header "Transition Semantics of Commands"
theory Transition imports Natural begin
subsection "The transition relation"
text {*
We formalize the transition semantics as in \cite{Nielson}. This
makes some of the rules a bit more intuitive, but also requires
some more (internal) formal overhead.
Since configurations that have terminated are written without
a statement, the transition relation is not
@{typ "((com × state) × (com × state)) set"}
but instead:
@{typ "((com option × state) × (com option × state)) set"}
Some syntactic sugar that we will use to hide the
@{text option} part in configurations:
*}
abbreviation
angle :: "[com, state] => com option × state" ("<_,_>") where
"<c,s> == (Some c, s)"
abbreviation
angle2 :: "state => com option × state" ("<_>") where
"<s> == (None, s)"
notation (xsymbols)
angle ("〈_,_〉") and
angle2 ("〈_〉")
notation (HTML output)
angle ("〈_,_〉") and
angle2 ("〈_〉")
text {*
Now, finally, we are set to write down the rules for our small step semantics:
*}
inductive_set
evalc1 :: "((com option × state) × (com option × state)) set"
and evalc1' :: "[(com option×state),(com option×state)] => bool"
("_ -->\<^sub>1 _" [60,60] 61)
where
"cs -->\<^sub>1 cs' == (cs,cs') ∈ evalc1"
| Skip: "〈\<SKIP>, s〉 -->\<^sub>1 〈s〉"
| Assign: "〈x :== a, s〉 -->\<^sub>1 〈s[x \<mapsto> a s]〉"
| Semi1: "〈c0,s〉 -->\<^sub>1 〈s'〉 ==> 〈c0;c1,s〉 -->\<^sub>1 〈c1,s'〉"
| Semi2: "〈c0,s〉 -->\<^sub>1 〈c0',s'〉 ==> 〈c0;c1,s〉 -->\<^sub>1 〈c0';c1,s'〉"
| IfTrue: "b s ==> 〈\<IF> b \<THEN> c1 \<ELSE> c2,s〉 -->\<^sub>1 〈c1,s〉"
| IfFalse: "¬b s ==> 〈\<IF> b \<THEN> c1 \<ELSE> c2,s〉 -->\<^sub>1 〈c2,s〉"
| While: "〈\<WHILE> b \<DO> c,s〉 -->\<^sub>1 〈\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>,s〉"
lemmas [intro] = evalc1.intros -- "again, use these rules in automatic proofs"
text {*
More syntactic sugar for the transition relation, and its
iteration.
*}
abbreviation
evalcn :: "[(com option×state),nat,(com option×state)] => bool"
("_ -_->\<^sub>1 _" [60,60,60] 60) where
"cs -n->\<^sub>1 cs' == (cs,cs') ∈ evalc1^^n"
abbreviation
evalc' :: "[(com option×state),(com option×state)] => bool"
("_ -->\<^sub>1\<^sup>* _" [60,60] 60) where
"cs -->\<^sub>1\<^sup>* cs' == (cs,cs') ∈ evalc1^*"
declare rel_pow_Suc_E2 [elim!]
text {*
As for the big step semantics you can read these rules in a
syntax directed way:
*}
lemma SKIP_1: "〈\<SKIP>, s〉 -->\<^sub>1 y = (y = 〈s〉)"
by (induct y, rule, cases set: evalc1, auto)
lemma Assign_1: "〈x :== a, s〉 -->\<^sub>1 y = (y = 〈s[x \<mapsto> a s]〉)"
by (induct y, rule, cases set: evalc1, auto)
lemma Cond_1:
"〈\<IF> b \<THEN> c1 \<ELSE> c2, s〉 -->\<^sub>1 y = ((b s --> y = 〈c1, s〉) ∧ (¬b s --> y = 〈c2, s〉))"
by (induct y, rule, cases set: evalc1, auto)
lemma While_1:
"〈\<WHILE> b \<DO> c, s〉 -->\<^sub>1 y = (y = 〈\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>, s〉)"
by (induct y, rule, cases set: evalc1, auto)
lemmas [simp] = SKIP_1 Assign_1 Cond_1 While_1
subsection "Examples"
lemma
"s x = 0 ==> 〈\<WHILE> λs. s x ≠ 1 \<DO> (x:== λs. s x+1), s〉 -->\<^sub>1\<^sup>* 〈s[x \<mapsto> 1]〉"
(is "_ ==> 〈?w, _〉 -->\<^sub>1\<^sup>* _")
proof -
let ?c = "x:== λs. s x+1"
let ?if = "\<IF> λs. s x ≠ 1 \<THEN> ?c; ?w \<ELSE> \<SKIP>"
assume [simp]: "s x = 0"
have "〈?w, s〉 -->\<^sub>1 〈?if, s〉" ..
also have "〈?if, s〉 -->\<^sub>1 〈?c; ?w, s〉" by simp
also have "〈?c; ?w, s〉 -->\<^sub>1 〈?w, s[x \<mapsto> 1]〉" by (rule Semi1) simp
also have "〈?w, s[x \<mapsto> 1]〉 -->\<^sub>1 〈?if, s[x \<mapsto> 1]〉" ..
also have "〈?if, s[x \<mapsto> 1]〉 -->\<^sub>1 〈\<SKIP>, s[x \<mapsto> 1]〉" by (simp add: update_def)
also have "〈\<SKIP>, s[x \<mapsto> 1]〉 -->\<^sub>1 〈s[x \<mapsto> 1]〉" ..
finally show ?thesis ..
qed
lemma
"s x = 2 ==> 〈\<WHILE> λs. s x ≠ 1 \<DO> (x:== λs. s x+1), s〉 -->\<^sub>1\<^sup>* s'"
(is "_ ==> 〈?w, _〉 -->\<^sub>1\<^sup>* s'")
proof -
let ?c = "x:== λs. s x+1"
let ?if = "\<IF> λs. s x ≠ 1 \<THEN> ?c; ?w \<ELSE> \<SKIP>"
assume [simp]: "s x = 2"
note update_def [simp]
have "〈?w, s〉 -->\<^sub>1 〈?if, s〉" ..
also have "〈?if, s〉 -->\<^sub>1 〈?c; ?w, s〉" by simp
also have "〈?c; ?w, s〉 -->\<^sub>1 〈?w, s[x \<mapsto> 3]〉" by (rule Semi1) simp
also have "〈?w, s[x \<mapsto> 3]〉 -->\<^sub>1 〈?if, s[x \<mapsto> 3]〉" ..
also have "〈?if, s[x \<mapsto> 3]〉 -->\<^sub>1 〈?c; ?w, s[x \<mapsto> 3]〉" by simp
also have "〈?c; ?w, s[x \<mapsto> 3]〉 -->\<^sub>1 〈?w, s[x \<mapsto> 4]〉" by (rule Semi1) simp
also have "〈?w, s[x \<mapsto> 4]〉 -->\<^sub>1 〈?if, s[x \<mapsto> 4]〉" ..
also have "〈?if, s[x \<mapsto> 4]〉 -->\<^sub>1 〈?c; ?w, s[x \<mapsto> 4]〉" by simp
also have "〈?c; ?w, s[x \<mapsto> 4]〉 -->\<^sub>1 〈?w, s[x \<mapsto> 5]〉" by (rule Semi1) simp
oops
subsection "Basic properties"
text {*
There are no \emph{stuck} programs:
*}
lemma no_stuck: "∃y. 〈c,s〉 -->\<^sub>1 y"
proof (induct c)
-- "case Semi:"
fix c1 c2 assume "∃y. 〈c1,s〉 -->\<^sub>1 y"
then obtain y where "〈c1,s〉 -->\<^sub>1 y" ..
then obtain c1' s' where "〈c1,s〉 -->\<^sub>1 〈s'〉 ∨ 〈c1,s〉 -->\<^sub>1 〈c1',s'〉"
by (cases y, cases "fst y") auto
thus "∃s'. 〈c1;c2,s〉 -->\<^sub>1 s'" by auto
next
-- "case If:"
fix b c1 c2 assume "∃y. 〈c1,s〉 -->\<^sub>1 y" and "∃y. 〈c2,s〉 -->\<^sub>1 y"
thus "∃y. 〈\<IF> b \<THEN> c1 \<ELSE> c2, s〉 -->\<^sub>1 y" by (cases "b s") auto
qed auto -- "the rest is trivial"
text {*
If a configuration does not contain a statement, the program
has terminated and there is no next configuration:
*}
lemma stuck [elim!]: "〈s〉 -->\<^sub>1 y ==> P"
by (induct y, auto elim: evalc1.cases)
lemma evalc_None_retrancl [simp, dest!]: "〈s〉 -->\<^sub>1\<^sup>* s' ==> s' = 〈s〉"
by (induct set: rtrancl) auto
lemmas [simp del] = relpow.simps
lemma rel_pow_0 [simp]: "!!R::('a*'a) set. R ^^ 0 = Id" by (simp add: relpow.simps)
lemma rel_pow_Suc_0 [simp]: "!!R::('a*'a) set. R ^^ Suc 0 = R" by (simp add: relpow.simps)
lemma evalc1_None_0 [simp]: "〈s〉 -n->\<^sub>1 y = (n = 0 ∧ y = 〈s〉)"
by (cases n) auto
lemma SKIP_n: "〈\<SKIP>, s〉 -n->\<^sub>1 〈s'〉 ==> s' = s ∧ n=1"
by (cases n) auto
subsection "Equivalence to natural semantics (after Nielson and Nielson)"
text {*
We first need two lemmas about semicolon statements:
decomposition and composition.
*}
lemma semiD:
"〈c1; c2, s〉 -n->\<^sub>1 〈s''〉 ==>
∃i j s'. 〈c1, s〉 -i->\<^sub>1 〈s'〉 ∧ 〈c2, s'〉 -j->\<^sub>1 〈s''〉 ∧ n = i+j"
proof (induct n arbitrary: c1 c2 s s'')
case 0
then show ?case by simp
next
case (Suc n)
from `〈c1; c2, s〉 -Suc n->\<^sub>1 〈s''〉`
obtain co s''' where
1: "〈c1; c2, s〉 -->\<^sub>1 (co, s''')" and
n: "(co, s''') -n->\<^sub>1 〈s''〉"
by auto
from 1
show "∃i j s'. 〈c1, s〉 -i->\<^sub>1 〈s'〉 ∧ 〈c2, s'〉 -j->\<^sub>1 〈s''〉 ∧ Suc n = i+j"
(is "∃i j s'. ?Q i j s'")
proof (cases set: evalc1)
case Semi1
from `co = Some c2` and `〈c1, s〉 -->\<^sub>1 〈s'''〉` and 1 n
have "?Q 1 n s'''" by simp
thus ?thesis by blast
next
case (Semi2 c1')
note c1 = `〈c1, s〉 -->\<^sub>1 〈c1', s'''〉`
with `co = Some (c1'; c2)` and n
have "〈c1'; c2, s'''〉 -n->\<^sub>1 〈s''〉" by simp
with Suc.hyps obtain i j s0 where
c1': "〈c1',s'''〉 -i->\<^sub>1 〈s0〉" and
c2: "〈c2,s0〉 -j->\<^sub>1 〈s''〉" and
i: "n = i+j"
by fast
from c1 c1'
have "〈c1,s〉 -(i+1)->\<^sub>1 〈s0〉" by (auto intro: rel_pow_Suc_I2)
with c2 i
have "?Q (i+1) j s0" by simp
thus ?thesis by blast
qed
qed
lemma semiI:
"〈c0,s〉 -n->\<^sub>1 〈s''〉 ==> 〈c1,s''〉 -->\<^sub>1\<^sup>* 〈s'〉 ==> 〈c0; c1, s〉 -->\<^sub>1\<^sup>* 〈s'〉"
proof (induct n arbitrary: c0 s s'')
case 0
from `〈c0,s〉 -(0::nat)->\<^sub>1 〈s''〉`
have False by simp
thus ?case ..
next
case (Suc n)
note c0 = `〈c0,s〉 -Suc n->\<^sub>1 〈s''〉`
note c1 = `〈c1,s''〉 -->\<^sub>1\<^sup>* 〈s'〉`
note IH = `!!c0 s s''.
〈c0,s〉 -n->\<^sub>1 〈s''〉 ==> 〈c1,s''〉 -->\<^sub>1\<^sup>* 〈s'〉 ==> 〈c0; c1,s〉 -->\<^sub>1\<^sup>* 〈s'〉`
from c0 obtain y where
1: "〈c0,s〉 -->\<^sub>1 y" and n: "y -n->\<^sub>1 〈s''〉" by blast
from 1 obtain c0' s0' where
"y = 〈s0'〉 ∨ y = 〈c0', s0'〉"
by (cases y, cases "fst y") auto
moreover
{ assume y: "y = 〈s0'〉"
with n have "s'' = s0'" by simp
with y 1 have "〈c0; c1,s〉 -->\<^sub>1 〈c1, s''〉" by blast
with c1 have "〈c0; c1,s〉 -->\<^sub>1\<^sup>* 〈s'〉" by (blast intro: rtrancl_trans)
}
moreover
{ assume y: "y = 〈c0', s0'〉"
with n have "〈c0', s0'〉 -n->\<^sub>1 〈s''〉" by blast
with IH c1 have "〈c0'; c1,s0'〉 -->\<^sub>1\<^sup>* 〈s'〉" by blast
moreover
from y 1 have "〈c0; c1,s〉 -->\<^sub>1 〈c0'; c1,s0'〉" by blast
hence "〈c0; c1,s〉 -->\<^sub>1\<^sup>* 〈c0'; c1,s0'〉" by blast
ultimately
have "〈c0; c1,s〉 -->\<^sub>1\<^sup>* 〈s'〉" by (blast intro: rtrancl_trans)
}
ultimately
show "〈c0; c1,s〉 -->\<^sub>1\<^sup>* 〈s'〉" by blast
qed
text {*
The easy direction of the equivalence proof:
*}
lemma evalc_imp_evalc1:
assumes "〈c,s〉 -->\<^sub>c s'"
shows "〈c, s〉 -->\<^sub>1\<^sup>* 〈s'〉"
using prems
proof induct
fix s show "〈\<SKIP>,s〉 -->\<^sub>1\<^sup>* 〈s〉" by auto
next
fix x a s show "〈x :== a ,s〉 -->\<^sub>1\<^sup>* 〈s[x\<mapsto>a s]〉" by auto
next
fix c0 c1 s s'' s'
assume "〈c0,s〉 -->\<^sub>1\<^sup>* 〈s''〉"
then obtain n where "〈c0,s〉 -n->\<^sub>1 〈s''〉" by (blast dest: rtrancl_imp_rel_pow)
moreover
assume "〈c1,s''〉 -->\<^sub>1\<^sup>* 〈s'〉"
ultimately
show "〈c0; c1,s〉 -->\<^sub>1\<^sup>* 〈s'〉" by (rule semiI)
next
fix s::state and b c0 c1 s'
assume "b s"
hence "〈\<IF> b \<THEN> c0 \<ELSE> c1,s〉 -->\<^sub>1 〈c0,s〉" by simp
also assume "〈c0,s〉 -->\<^sub>1\<^sup>* 〈s'〉"
finally show "〈\<IF> b \<THEN> c0 \<ELSE> c1,s〉 -->\<^sub>1\<^sup>* 〈s'〉" .
next
fix s::state and b c0 c1 s'
assume "¬b s"
hence "〈\<IF> b \<THEN> c0 \<ELSE> c1,s〉 -->\<^sub>1 〈c1,s〉" by simp
also assume "〈c1,s〉 -->\<^sub>1\<^sup>* 〈s'〉"
finally show "〈\<IF> b \<THEN> c0 \<ELSE> c1,s〉 -->\<^sub>1\<^sup>* 〈s'〉" .
next
fix b c and s::state
assume b: "¬b s"
let ?if = "\<IF> b \<THEN> c; \<WHILE> b \<DO> c \<ELSE> \<SKIP>"
have "〈\<WHILE> b \<DO> c,s〉 -->\<^sub>1 〈?if, s〉" by blast
also have "〈?if,s〉 -->\<^sub>1 〈\<SKIP>, s〉" by (simp add: b)
also have "〈\<SKIP>, s〉 -->\<^sub>1 〈s〉" by blast
finally show "〈\<WHILE> b \<DO> c,s〉 -->\<^sub>1\<^sup>* 〈s〉" ..
next
fix b c s s'' s'
let ?w = "\<WHILE> b \<DO> c"
let ?if = "\<IF> b \<THEN> c; ?w \<ELSE> \<SKIP>"
assume w: "〈?w,s''〉 -->\<^sub>1\<^sup>* 〈s'〉"
assume c: "〈c,s〉 -->\<^sub>1\<^sup>* 〈s''〉"
assume b: "b s"
have "〈?w,s〉 -->\<^sub>1 〈?if, s〉" by blast
also have "〈?if, s〉 -->\<^sub>1 〈c; ?w, s〉" by (simp add: b)
also
from c obtain n where "〈c,s〉 -n->\<^sub>1 〈s''〉" by (blast dest: rtrancl_imp_rel_pow)
with w have "〈c; ?w,s〉 -->\<^sub>1\<^sup>* 〈s'〉" by - (rule semiI)
finally show "〈\<WHILE> b \<DO> c,s〉 -->\<^sub>1\<^sup>* 〈s'〉" ..
qed
text {*
Finally, the equivalence theorem:
*}
theorem evalc_equiv_evalc1:
"〈c, s〉 -->\<^sub>c s' = 〈c,s〉 -->\<^sub>1\<^sup>* 〈s'〉"
proof
assume "〈c,s〉 -->\<^sub>c s'"
then show "〈c, s〉 -->\<^sub>1\<^sup>* 〈s'〉" by (rule evalc_imp_evalc1)
next
assume "〈c, s〉 -->\<^sub>1\<^sup>* 〈s'〉"
then obtain n where "〈c, s〉 -n->\<^sub>1 〈s'〉" by (blast dest: rtrancl_imp_rel_pow)
moreover
have "〈c, s〉 -n->\<^sub>1 〈s'〉 ==> 〈c,s〉 -->\<^sub>c s'"
proof (induct arbitrary: c s s' rule: less_induct)
fix n
assume IH: "!!m c s s'. m < n ==> 〈c,s〉 -m->\<^sub>1 〈s'〉 ==> 〈c,s〉 -->\<^sub>c s'"
fix c s s'
assume c: "〈c, s〉 -n->\<^sub>1 〈s'〉"
then obtain m where n: "n = Suc m" by (cases n) auto
with c obtain y where
c': "〈c, s〉 -->\<^sub>1 y" and m: "y -m->\<^sub>1 〈s'〉" by blast
show "〈c,s〉 -->\<^sub>c s'"
proof (cases c)
case SKIP
with c n show ?thesis by auto
next
case Assign
with c n show ?thesis by auto
next
fix c1 c2 assume semi: "c = (c1; c2)"
with c obtain i j s'' where
c1: "〈c1, s〉 -i->\<^sub>1 〈s''〉" and
c2: "〈c2, s''〉 -j->\<^sub>1 〈s'〉" and
ij: "n = i+j"
by (blast dest: semiD)
from c1 c2 obtain
"0 < i" and "0 < j" by (cases i, auto, cases j, auto)
with ij obtain
i: "i < n" and j: "j < n" by simp
from IH i c1
have "〈c1,s〉 -->\<^sub>c s''" .
moreover
from IH j c2
have "〈c2,s''〉 -->\<^sub>c s'" .
moreover
note semi
ultimately
show "〈c,s〉 -->\<^sub>c s'" by blast
next
fix b c1 c2 assume If: "c = \<IF> b \<THEN> c1 \<ELSE> c2"
{ assume True: "b s = True"
with If c n
have "〈c1,s〉 -m->\<^sub>1 〈s'〉" by auto
with n IH
have "〈c1,s〉 -->\<^sub>c s'" by blast
with If True
have "〈c,s〉 -->\<^sub>c s'" by blast
}
moreover
{ assume False: "b s = False"
with If c n
have "〈c2,s〉 -m->\<^sub>1 〈s'〉" by auto
with n IH
have "〈c2,s〉 -->\<^sub>c s'" by blast
with If False
have "〈c,s〉 -->\<^sub>c s'" by blast
}
ultimately
show "〈c,s〉 -->\<^sub>c s'" by (cases "b s") auto
next
fix b c' assume w: "c = \<WHILE> b \<DO> c'"
with c n
have "〈\<IF> b \<THEN> c'; \<WHILE> b \<DO> c' \<ELSE> \<SKIP>,s〉 -m->\<^sub>1 〈s'〉"
(is "〈?if,_〉 -m->\<^sub>1 _") by auto
with n IH
have "〈\<IF> b \<THEN> c'; \<WHILE> b \<DO> c' \<ELSE> \<SKIP>,s〉 -->\<^sub>c s'" by blast
moreover note unfold_while [of b c']
-- {* @{thm unfold_while [of b c']} *}
ultimately
have "〈\<WHILE> b \<DO> c',s〉 -->\<^sub>c s'" by (blast dest: equivD2)
with w show "〈c,s〉 -->\<^sub>c s'" by simp
qed
qed
ultimately
show "〈c,s〉 -->\<^sub>c s'" by blast
qed
subsection "Winskel's Proof"
declare rel_pow_0_E [elim!]
text {*
Winskel's small step rules are a bit different \cite{Winskel};
we introduce their equivalents as derived rules:
*}
lemma whileFalse1 [intro]:
"¬ b s ==> 〈\<WHILE> b \<DO> c,s〉 -->\<^sub>1\<^sup>* 〈s〉" (is "_ ==> 〈?w, s〉 -->\<^sub>1\<^sup>* 〈s〉")
proof -
assume "¬b s"
have "〈?w, s〉 -->\<^sub>1 〈\<IF> b \<THEN> c;?w \<ELSE> \<SKIP>, s〉" ..
also from `¬b s` have "〈\<IF> b \<THEN> c;?w \<ELSE> \<SKIP>, s〉 -->\<^sub>1 〈\<SKIP>, s〉" ..
also have "〈\<SKIP>, s〉 -->\<^sub>1 〈s〉" ..
finally show "〈?w, s〉 -->\<^sub>1\<^sup>* 〈s〉" ..
qed
lemma whileTrue1 [intro]:
"b s ==> 〈\<WHILE> b \<DO> c,s〉 -->\<^sub>1\<^sup>* 〈c;\<WHILE> b \<DO> c, s〉"
(is "_ ==> 〈?w, s〉 -->\<^sub>1\<^sup>* 〈c;?w,s〉")
proof -
assume "b s"
have "〈?w, s〉 -->\<^sub>1 〈\<IF> b \<THEN> c;?w \<ELSE> \<SKIP>, s〉" ..
also from `b s` have "〈\<IF> b \<THEN> c;?w \<ELSE> \<SKIP>, s〉 -->\<^sub>1 〈c;?w, s〉" ..
finally show "〈?w, s〉 -->\<^sub>1\<^sup>* 〈c;?w,s〉" ..
qed
inductive_cases evalc1_SEs:
"〈\<SKIP>,s〉 -->\<^sub>1 (co, s')"
"〈x:==a,s〉 -->\<^sub>1 (co, s')"
"〈c1;c2, s〉 -->\<^sub>1 (co, s')"
"〈\<IF> b \<THEN> c1 \<ELSE> c2, s〉 -->\<^sub>1 (co, s')"
"〈\<WHILE> b \<DO> c, s〉 -->\<^sub>1 (co, s')"
inductive_cases evalc1_E: "〈\<WHILE> b \<DO> c, s〉 -->\<^sub>1 (co, s')"
declare evalc1_SEs [elim!]
lemma evalc_impl_evalc1: "〈c,s〉 -->\<^sub>c s1 ==> 〈c,s〉 -->\<^sub>1\<^sup>* 〈s1〉"
apply (induct set: evalc)
-- SKIP
apply blast
-- ASSIGN
apply fast
-- SEMI
apply (fast dest: rtrancl_imp_UN_rel_pow intro: semiI)
-- IF
apply (fast intro: converse_rtrancl_into_rtrancl)
apply (fast intro: converse_rtrancl_into_rtrancl)
-- WHILE
apply blast
apply (blast dest: rtrancl_imp_UN_rel_pow intro: converse_rtrancl_into_rtrancl semiI)
done
lemma lemma2:
"〈c;d,s〉 -n->\<^sub>1 〈u〉 ==> ∃t m. 〈c,s〉 -->\<^sub>1\<^sup>* 〈t〉 ∧ 〈d,t〉 -m->\<^sub>1 〈u〉 ∧ m ≤ n"
apply (induct n arbitrary: c d s u)
-- "case n = 0"
apply fastsimp
-- "induction step"
apply (fast intro!: le_SucI le_refl dest!: rel_pow_Suc_D2
elim!: rel_pow_imp_rtrancl converse_rtrancl_into_rtrancl)
done
lemma evalc1_impl_evalc:
"〈c,s〉 -->\<^sub>1\<^sup>* 〈t〉 ==> 〈c,s〉 -->\<^sub>c t"
apply (induct c arbitrary: s t)
apply (safe dest!: rtrancl_imp_UN_rel_pow)
-- SKIP
apply (simp add: SKIP_n)
-- ASSIGN
apply (fastsimp elim: rel_pow_E2)
-- SEMI
apply (fast dest!: rel_pow_imp_rtrancl lemma2)
-- IF
apply (erule rel_pow_E2)
apply simp
apply (fast dest!: rel_pow_imp_rtrancl)
-- "WHILE, induction on the length of the computation"
apply (rename_tac b c s t n)
apply (erule_tac P = "?X -n->\<^sub>1 ?Y" in rev_mp)
apply (rule_tac x = "s" in spec)
apply (induct_tac n rule: nat_less_induct)
apply (intro strip)
apply (erule rel_pow_E2)
apply simp
apply (simp only: split_paired_all)
apply (erule evalc1_E)
apply simp
apply (case_tac "b x")
-- WhileTrue
apply (erule rel_pow_E2)
apply simp
apply (clarify dest!: lemma2)
apply atomize
apply (erule allE, erule allE, erule impE, assumption)
apply (erule_tac x=mb in allE, erule impE, fastsimp)
apply blast
-- WhileFalse
apply (erule rel_pow_E2)
apply simp
apply (simp add: SKIP_n)
done
text {* proof of the equivalence of evalc and evalc1 *}
lemma evalc1_eq_evalc: "(〈c, s〉 -->\<^sub>1\<^sup>* 〈t〉) = (〈c,s〉 -->\<^sub>c t)"
by (fast elim!: evalc1_impl_evalc evalc_impl_evalc1)
subsection "A proof without n"
text {*
The inductions are a bit awkward to write in this section,
because @{text None} as result statement in the small step
semantics doesn't have a direct counterpart in the big step
semantics.
Winskel's small step rule set (using the @{text "\<SKIP>"} statement
to indicate termination) is better suited for this proof.
*}
lemma my_lemma1:
assumes "〈c1,s1〉 -->\<^sub>1\<^sup>* 〈s2〉"
and "〈c2,s2〉 -->\<^sub>1\<^sup>* cs3"
shows "〈c1;c2,s1〉 -->\<^sub>1\<^sup>* cs3"
proof -
-- {* The induction rule needs @{text P} to be a function of @{term "Some c1"} *}
from prems
have "〈(λc. if c = None then c2 else the c; c2) (Some c1),s1〉 -->\<^sub>1\<^sup>* cs3"
apply (induct rule: converse_rtrancl_induct2)
apply simp
apply (rename_tac c s')
apply simp
apply (rule conjI)
apply fast
apply clarify
apply (case_tac c)
apply (auto intro: converse_rtrancl_into_rtrancl)
done
then show ?thesis by simp
qed
lemma evalc_impl_evalc1': "〈c,s〉 -->\<^sub>c s1 ==> 〈c,s〉 -->\<^sub>1\<^sup>* 〈s1〉"
apply (induct set: evalc)
-- SKIP
apply fast
-- ASSIGN
apply fast
-- SEMI
apply (fast intro: my_lemma1)
-- IF
apply (fast intro: converse_rtrancl_into_rtrancl)
apply (fast intro: converse_rtrancl_into_rtrancl)
-- WHILE
apply fast
apply (fast intro: converse_rtrancl_into_rtrancl my_lemma1)
done
text {*
The opposite direction is based on a Coq proof done by Ranan Fraer and
Yves Bertot. The following sketch is from an email by Ranan Fraer.
\begin{verbatim}
First we've broke it into 2 lemmas:
Lemma 1
((c,s) --> (SKIP,t)) => (<c,s> -c-> t)
This is a quick one, dealing with the cases skip, assignment
and while_false.
Lemma 2
((c,s) -*-> (c',s')) /\ <c',s'> -c'-> t
=>
<c,s> -c-> t
This is proved by rule induction on the -*-> relation
and the induction step makes use of a third lemma:
Lemma 3
((c,s) --> (c',s')) /\ <c',s'> -c'-> t
=>
<c,s> -c-> t
This captures the essence of the proof, as it shows that <c',s'>
behaves as the continuation of <c,s> with respect to the natural
semantics.
The proof of Lemma 3 goes by rule induction on the --> relation,
dealing with the cases sequence1, sequence2, if_true, if_false and
while_true. In particular in the case (sequence1) we make use again
of Lemma 1.
\end{verbatim}
*}
inductive_cases evalc1_term_cases: "〈c,s〉 -->\<^sub>1 〈s'〉"
lemma FB_lemma3:
"(c,s) -->\<^sub>1 (c',s') ==> c ≠ None ==>
〈if c'=None then \<SKIP> else the c',s'〉 -->\<^sub>c t ==> 〈the c,s〉 -->\<^sub>c t"
by (induct arbitrary: t set: evalc1)
(auto elim!: evalc1_term_cases equivD2 [OF unfold_while])
lemma FB_lemma2:
"(c,s) -->\<^sub>1\<^sup>* (c',s') ==> c ≠ None ==>
〈if c' = None then \<SKIP> else the c',s'〉 -->\<^sub>c t ==> 〈the c,s〉 -->\<^sub>c t"
apply (induct rule: converse_rtrancl_induct2, force)
apply (fastsimp elim!: evalc1_term_cases intro: FB_lemma3)
done
lemma evalc1_impl_evalc': "〈c,s〉 -->\<^sub>1\<^sup>* 〈t〉 ==> 〈c,s〉 -->\<^sub>c t"
by (fastsimp dest: FB_lemma2)
end