Theory Natural

Up to index of Isabelle/HOL/IMP

theory Natural
imports Com

(*  Title:        HOL/IMP/Natural.thy
ID: $Id$
Author: Tobias Nipkow & Robert Sandner, TUM
Isar Version: Gerwin Klein, 2001; additional proofs by Lawrence Paulson
Copyright 1996 TUM
*)


header "Natural Semantics of Commands"

theory Natural imports Com begin

subsection "Execution of commands"

text {*
We write @{text "⟨c,s⟩ -->c s'"} for \emph{Statement @{text c}, started
in state @{text s}, terminates in state @{text s'}}. Formally,
@{text "⟨c,s⟩ -->c s'"} is just another form of saying \emph{the tuple
@{text "(c,s,s')"} is part of the relation @{text evalc}}:
*}


definition
update :: "('a => 'b) => 'a => 'b => ('a => 'b)" ("_/[_ ::= /_]" [900,0,0] 900) where
"update = fun_upd"


notation (xsymbols)
update ("_/[_ \<mapsto> /_]" [900,0,0] 900)


text {* Disable conflicting syntax from HOL Map theory. *}

no_syntax
"_maplet" :: "['a, 'a] => maplet" ("_ /|->/ _")
"_maplets" :: "['a, 'a] => maplet" ("_ /[|->]/ _")
"" :: "maplet => maplets" ("_")
"_Maplets" :: "[maplet, maplets] => maplets" ("_,/ _")
"_MapUpd" :: "['a ~=> 'b, maplets] => 'a ~=> 'b" ("_/'(_')" [900,0]900)
"_Map" :: "maplets => 'a ~=> 'b" ("(1[_])")


text {*
The big-step execution relation @{text evalc} is defined inductively:
*}

inductive
evalc :: "[com,state,state] => bool" ("⟨_,_⟩/ -->c _" [0,0,60] 60)
where
Skip: "⟨\<SKIP>,s⟩ -->c s"
| Assign: "⟨x :== a,s⟩ -->c s[x\<mapsto>a s]"

| Semi: "⟨c0,s⟩ -->c s'' ==> ⟨c1,s''⟩ -->c s' ==> ⟨c0; c1, s⟩ -->c s'"

| IfTrue: "b s ==> ⟨c0,s⟩ -->c s' ==> ⟨\<IF> b \<THEN> c0 \<ELSE> c1, s⟩ -->c s'"
| IfFalse: "¬b s ==> ⟨c1,s⟩ -->c s' ==> ⟨\<IF> b \<THEN> c0 \<ELSE> c1, s⟩ -->c s'"

| WhileFalse: "¬b s ==> ⟨\<WHILE> b \<DO> c,s⟩ -->c s"
| WhileTrue: "b s ==> ⟨c,s⟩ -->c s'' ==> ⟨\<WHILE> b \<DO> c, s''⟩ -->c s'
==> ⟨\<WHILE> b \<DO> c, s⟩ -->c s'"


lemmas evalc.intros [intro] -- "use those rules in automatic proofs"

text {*
The induction principle induced by this definition looks like this:

@{thm [display] evalc.induct [no_vars]}


(@{text "!!"} and @{text "==>"} are Isabelle's
meta symbols for @{text "∀"} and @{text "-->"})
*}


text {*
The rules of @{text evalc} are syntax directed, i.e.~for each
syntactic category there is always only one rule applicable. That
means we can use the rules in both directions. This property is called rule inversion.
*}

inductive_cases skipE [elim!]: "⟨\<SKIP>,s⟩ -->c s'"
inductive_cases semiE [elim!]: "⟨c0; c1, s⟩ -->c s'"
inductive_cases assignE [elim!]: "⟨x :== a,s⟩ -->c s'"
inductive_cases ifE [elim!]: "⟨\<IF> b \<THEN> c0 \<ELSE> c1, s⟩ -->c s'"
inductive_cases whileE [elim]: "⟨\<WHILE> b \<DO> c,s⟩ -->c s'"

text {* The next proofs are all trivial by rule inversion.
*}


lemma skip:
"⟨\<SKIP>,s⟩ -->c s' = (s' = s)"

by auto

lemma assign:
"⟨x :== a,s⟩ -->c s' = (s' = s[x\<mapsto>a s])"

by auto

lemma semi:
"⟨c0; c1, s⟩ -->c s' = (∃s''. ⟨c0,s⟩ -->c s'' ∧ ⟨c1,s''⟩ -->c s')"

by auto

lemma ifTrue:
"b s ==> ⟨\<IF> b \<THEN> c0 \<ELSE> c1, s⟩ -->c s' = ⟨c0,s⟩ -->c s'"

by auto

lemma ifFalse:
"¬b s ==> ⟨\<IF> b \<THEN> c0 \<ELSE> c1, s⟩ -->c s' = ⟨c1,s⟩ -->c s'"

by auto

lemma whileFalse:
"¬ b s ==> ⟨\<WHILE> b \<DO> c,s⟩ -->c s' = (s' = s)"

by auto

lemma whileTrue:
"b s ==>
⟨\<WHILE> b \<DO> c, s⟩ -->c s' =
(∃s''. ⟨c,s⟩ -->c s'' ∧ ⟨\<WHILE> b \<DO> c, s''⟩ -->c s')"

by auto

text "Again, Isabelle may use these rules in automatic proofs:"
lemmas evalc_cases [simp] = skip assign ifTrue ifFalse whileFalse semi whileTrue

subsection "Equivalence of statements"

text {*
We call two statements @{text c} and @{text c'} equivalent wrt.~the
big-step semantics when \emph{@{text c} started in @{text s} terminates
in @{text s'} iff @{text c'} started in the same @{text s} also terminates
in the same @{text s'}}. Formally:
*}

definition
equiv_c :: "com => com => bool" ("_ ∼ _" [56, 56] 55) where
"c ∼ c' = (∀s s'. ⟨c, s⟩ -->c s' = ⟨c', s⟩ -->c s')"


text {*
Proof rules telling Isabelle to unfold the definition
if there is something to be proved about equivalent
statements: *}

lemma equivI [intro!]:
"(!!s s'. ⟨c, s⟩ -->c s' = ⟨c', s⟩ -->c s') ==> c ∼ c'"

by (unfold equiv_c_def) blast

lemma equivD1:
"c ∼ c' ==> ⟨c, s⟩ -->c s' ==> ⟨c', s⟩ -->c s'"

by (unfold equiv_c_def) blast

lemma equivD2:
"c ∼ c' ==> ⟨c', s⟩ -->c s' ==> ⟨c, s⟩ -->c s'"

by (unfold equiv_c_def) blast

text {*
As an example, we show that loop unfolding is an equivalence
transformation on programs:
*}

lemma unfold_while:
"(\<WHILE> b \<DO> c) ∼ (\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>)" (is "?w ∼ ?if")

proof -
-- "to show the equivalence, we look at the derivation tree for"
-- "each side and from that construct a derivation tree for the other side"

{ fix s s' assume w: "⟨?w, s⟩ -->c s'"
-- "as a first thing we note that, if @{text b} is @{text False} in state @{text s},"
-- "then both statements do nothing:"

hence "¬b s ==> s = s'" by blast
hence "¬b s ==> ⟨?if, s⟩ -->c s'" by blast
moreover
-- "on the other hand, if @{text b} is @{text True} in state @{text s},"
-- {* then only the @{text WhileTrue} rule can have been used to derive @{text "⟨?w, s⟩ -->c s'"} *}

{ assume b: "b s"
with w obtain s'' where
"⟨c, s⟩ -->c s''" and "⟨?w, s''⟩ -->c s'"
by (cases set: evalc) auto
-- "now we can build a derivation tree for the @{text \<IF>}"
-- "first, the body of the True-branch:"

hence "⟨c; ?w, s⟩ -->c s'" by (rule Semi)
-- "then the whole @{text \<IF>}"

with b have "⟨?if, s⟩ -->c s'" by (rule IfTrue)
}
ultimately
-- "both cases together give us what we want:"

have "⟨?if, s⟩ -->c s'" by blast
}
moreover
-- "now the other direction:"

{ fix s s' assume "if": "⟨?if, s⟩ -->c s'"
-- "again, if @{text b} is @{text False} in state @{text s}, then the False-branch"
-- "of the @{text \<IF>} is executed, and both statements do nothing:"

hence "¬b s ==> s = s'" by blast
hence "¬b s ==> ⟨?w, s⟩ -->c s'" by blast
moreover
-- "on the other hand, if @{text b} is @{text True} in state @{text s},"
-- {* then this time only the @{text IfTrue} rule can have be used *}

{ assume b: "b s"
with "if" have "⟨c; ?w, s⟩ -->c s'" by (cases set: evalc) auto
-- "and for this, only the Semi-rule is applicable:"

then obtain s'' where
"⟨c, s⟩ -->c s''" and "⟨?w, s''⟩ -->c s'"
by (cases set: evalc) auto
-- "with this information, we can build a derivation tree for the @{text \<WHILE>}"

with b
have "⟨?w, s⟩ -->c s'" by (rule WhileTrue)
}
ultimately
-- "both cases together again give us what we want:"

have "⟨?w, s⟩ -->c s'" by blast
}
ultimately
show ?thesis by blast
qed

text {*
Happily, such lengthy proofs are seldom necessary. Isabelle can prove many such facts automatically.
*}


lemma
"(\<WHILE> b \<DO> c) ∼ (\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>)"

by blast

lemma triv_if:
"(\<IF> b \<THEN> c \<ELSE> c) ∼ c"

by blast

lemma commute_if:
"(\<IF> b1 \<THEN> (\<IF> b2 \<THEN> c11 \<ELSE> c12) \<ELSE> c2)

(\<IF> b2 \<THEN> (\<IF> b1 \<THEN> c11 \<ELSE> c2) \<ELSE> (\<IF> b1 \<THEN> c12 \<ELSE> c2))"

by blast

lemma while_equiv:
"⟨c0, s⟩ -->c u ==> c ∼ c' ==> (c0 = \<WHILE> b \<DO> c) ==> ⟨\<WHILE> b \<DO> c', s⟩ -->c u"

by (induct rule: evalc.induct) (auto simp add: equiv_c_def)

lemma equiv_while:
"c ∼ c' ==> (\<WHILE> b \<DO> c) ∼ (\<WHILE> b \<DO> c')"

by (simp add: equiv_c_def) (metis equiv_c_def while_equiv)


text {*
Program equivalence is an equivalence relation.
*}


lemma equiv_refl:
"c ∼ c"

by blast

lemma equiv_sym:
"c1 ∼ c2 ==> c2 ∼ c1"

by (auto simp add: equiv_c_def)

lemma equiv_trans:
"c1 ∼ c2 ==> c2 ∼ c3 ==> c1 ∼ c3"

by (auto simp add: equiv_c_def)

text {*
Program constructions preserve equivalence.
*}


lemma equiv_semi:
"c1 ∼ c1' ==> c2 ∼ c2' ==> (c1; c2) ∼ (c1'; c2')"

by (force simp add: equiv_c_def)

lemma equiv_if:
"c1 ∼ c1' ==> c2 ∼ c2' ==> (\<IF> b \<THEN> c1 \<ELSE> c2) ∼ (\<IF> b \<THEN> c1' \<ELSE> c2')"

by (force simp add: equiv_c_def)

lemma while_never: "⟨c, s⟩ -->c u ==> c ≠ \<WHILE> (λs. True) \<DO> c1"
apply (induct rule: evalc.induct)
apply auto
done

lemma equiv_while_True:
"(\<WHILE> (λs. True) \<DO> c1) ∼ (\<WHILE> (λs. True) \<DO> c2)"

by (blast dest: while_never)


subsection "Execution is deterministic"

text {*
This proof is automatic.
*}

theorem "⟨c,s⟩ -->c t ==> ⟨c,s⟩ -->c u ==> u = t"
by (induct arbitrary: u rule: evalc.induct) blast+


text {*
The following proof presents all the details:
*}

theorem com_det:
assumes "⟨c,s⟩ -->c t" and "⟨c,s⟩ -->c u"
shows "u = t"

using prems
proof (induct arbitrary: u set: evalc)
fix s u assume "⟨\<SKIP>,s⟩ -->c u"
thus "u = s" by blast
next
fix a s x u assume "⟨x :== a,s⟩ -->c u"
thus "u = s[x \<mapsto> a s]" by blast
next
fix c0 c1 s s1 s2 u
assume IH0: "!!u. ⟨c0,s⟩ -->c u ==> u = s2"
assume IH1: "!!u. ⟨c1,s2⟩ -->c u ==> u = s1"

assume "⟨c0;c1, s⟩ -->c u"
then obtain s' where
c0: "⟨c0,s⟩ -->c s'" and
c1: "⟨c1,s'⟩ -->c u"

by auto

from c0 IH0 have "s'=s2" by blast
with c1 IH1 show "u=s1" by blast
next
fix b c0 c1 s s1 u
assume IH: "!!u. ⟨c0,s⟩ -->c u ==> u = s1"

assume "b s" and "⟨\<IF> b \<THEN> c0 \<ELSE> c1,s⟩ -->c u"
hence "⟨c0, s⟩ -->c u" by blast
with IH show "u = s1" by blast
next
fix b c0 c1 s s1 u
assume IH: "!!u. ⟨c1,s⟩ -->c u ==> u = s1"

assume "¬b s" and "⟨\<IF> b \<THEN> c0 \<ELSE> c1,s⟩ -->c u"
hence "⟨c1, s⟩ -->c u" by blast
with IH show "u = s1" by blast
next
fix b c s u
assume "¬b s" and "⟨\<WHILE> b \<DO> c,s⟩ -->c u"
thus "u = s" by blast
next
fix b c s s1 s2 u
assume "IHc": "!!u. ⟨c,s⟩ -->c u ==> u = s2"
assume "IHw": "!!u. ⟨\<WHILE> b \<DO> c,s2⟩ -->c u ==> u = s1"

assume "b s" and "⟨\<WHILE> b \<DO> c,s⟩ -->c u"
then obtain s' where
c: "⟨c,s⟩ -->c s'" and
w: "⟨\<WHILE> b \<DO> c,s'⟩ -->c u"

by auto

from c "IHc" have "s' = s2" by blast
with w "IHw" show "u = s1" by blast
qed


text {*
This is the proof as you might present it in a lecture. The remaining
cases are simple enough to be proved automatically:
*}

theorem
assumes "⟨c,s⟩ -->c t" and "⟨c,s⟩ -->c u"
shows "u = t"

using prems
proof (induct arbitrary: u)
-- "the simple @{text \<SKIP>} case for demonstration:"

fix s u assume "⟨\<SKIP>,s⟩ -->c u"
thus "u = s" by blast
next
-- "and the only really interesting case, @{text \<WHILE>}:"

fix b c s s1 s2 u
assume "IHc": "!!u. ⟨c,s⟩ -->c u ==> u = s2"
assume "IHw": "!!u. ⟨\<WHILE> b \<DO> c,s2⟩ -->c u ==> u = s1"

assume "b s" and "⟨\<WHILE> b \<DO> c,s⟩ -->c u"
then obtain s' where
c: "⟨c,s⟩ -->c s'" and
w: "⟨\<WHILE> b \<DO> c,s'⟩ -->c u"

by auto

from c "IHc" have "s' = s2" by blast
with w "IHw" show "u = s1" by blast
qed blast+ -- "prove the rest automatically"

end