(* Author: Gerwin Klein, Tobias Nipkow *)
subsection "Big-Step Semantics of Commands"
theory Big_Step imports Com begin
text \<open>
The big-step semantics is a straight-forward inductive definition
with concrete syntax. Note that the first parameter is a tuple,
so the syntax becomes \<open>(c,s) \<Rightarrow> s'\<close>.
\<close>
text_raw\<open>\snip{BigStepdef}{0}{1}{%\<close>
inductive
big_step :: "com \ state \ state \ bool" (infix "\" 55)
where
Skip: "(SKIP,s) \ s" |
Assign: "(x ::= a,s) \ s(x := aval a s)" |
Seq: "\ (c\<^sub>1,s\<^sub>1) \ s\<^sub>2; (c\<^sub>2,s\<^sub>2) \ s\<^sub>3 \ \ (c\<^sub>1;;c\<^sub>2, s\<^sub>1) \ s\<^sub>3" |
IfTrue: "\ bval b s; (c\<^sub>1,s) \ t \ \ (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \ t" |
IfFalse: "\ \bval b s; (c\<^sub>2,s) \ t \ \ (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \ t" |
WhileFalse: "\bval b s \ (WHILE b DO c,s) \ s" |
WhileTrue:
"\ bval b s\<^sub>1; (c,s\<^sub>1) \ s\<^sub>2; (WHILE b DO c, s\<^sub>2) \ s\<^sub>3 \
\<Longrightarrow> (WHILE b DO c, s\<^sub>1) \<Rightarrow> s\<^sub>3"
text_raw\<open>}%endsnip\<close>
text_raw\<open>\snip{BigStepEx}{1}{2}{%\<close>
schematic_goal ex: "(''x'' ::= N 5;; ''y'' ::= V ''x'', s) \ ?t"
apply(rule Seq)
apply(rule Assign)
apply simp
apply(rule Assign)
done
text_raw\<open>}%endsnip\<close>
thm ex[simplified]
text\<open>We want to execute the big-step rules:\<close>
code_pred big_step .
text\<open>For inductive definitions we need command
\texttt{values} instead of \texttt{value}.\<close>
values "{t. (SKIP, \_. 0) \ t}"
text\<open>We need to translate the result state into a list
to display it.\<close>
values "{map t [''x''] |t. (SKIP, <''x'' := 42>) \ t}"
values "{map t [''x''] |t. (''x'' ::= N 2, <''x'' := 42>) \ t}"
values "{map t [''x'',''y''] |t.
(WHILE Less (V ''x'') (V ''y'') DO (''x'' ::= Plus (V ''x'') (N 5)),
<''x'' := 0, ''y'' := 13>) \<Rightarrow> t}"
text\<open>Proof automation:\<close>
text \<open>The introduction rules are good for automatically
construction small program executions. The recursive cases
may require backtracking, so we declare the set as unsafe
intro rules.\<close>
declare big_step.intros [intro]
text\<open>The standard induction rule
@{thm [display] big_step.induct [no_vars]}\<close>
thm big_step.induct
text\<open>
This induction schema is almost perfect for our purposes, but
our trick for reusing the tuple syntax means that the induction
schema has two parameters instead of the \<open>c\<close>, \<open>s\<close>,
and \<open>s'\<close> that we are likely to encounter. Splitting
the tuple parameter fixes this:
\<close>
lemmas big_step_induct = big_step.induct[split_format(complete)]
thm big_step_induct
text \<open>
@{thm [display] big_step_induct [no_vars]}
\<close>
subsection "Rule inversion"
text\<open>What can we deduce from \<^prop>\<open>(SKIP,s) \<Rightarrow> t\<close> ?
That \<^prop>\<open>s = t\<close>. This is how we can automatically prove it:\<close>
inductive_cases SkipE[elim!]: "(SKIP,s) \ t"
thm SkipE
text\<open>This is an \emph{elimination rule}. The [elim] attribute tells auto,
blast and friends (but not simp!) to use it automatically; [elim!] means that
it is applied eagerly.
Similarly for the other commands:\<close>
inductive_cases AssignE[elim!]: "(x ::= a,s) \ t"
thm AssignE
inductive_cases SeqE[elim!]: "(c1;;c2,s1) \ s3"
thm SeqE
inductive_cases IfE[elim!]: "(IF b THEN c1 ELSE c2,s) \ t"
thm IfE
inductive_cases WhileE[elim]: "(WHILE b DO c,s) \ t"
thm WhileE
text\<open>Only [elim]: [elim!] would not terminate.\<close>
text\<open>An automatic example:\<close>
lemma "(IF b THEN SKIP ELSE SKIP, s) \ t \ t = s"
by blast
text\<open>Rule inversion by hand via the ``cases'' method:\<close>
lemma assumes "(IF b THEN SKIP ELSE SKIP, s) \ t"
shows "t = s"
proof-
from assms show ?thesis
proof cases \<comment> \<open>inverting assms\<close>
case IfTrue thm IfTrue
thus ?thesis by blast
next
case IfFalse thus ?thesis by blast
qed
qed
(* Using rule inversion to prove simplification rules: *)
lemma assign_simp:
"(x ::= a,s) \ s' \ (s' = s(x := aval a s))"
by auto
text \<open>An example combining rule inversion and derivations\<close>
lemma Seq_assoc:
"(c1;; c2;; c3, s) \ s' \ (c1;; (c2;; c3), s) \ s'"
proof
assume "(c1;; c2;; c3, s) \ s'"
then obtain s1 s2 where
c1: "(c1, s) \ s1" and
c2: "(c2, s1) \ s2" and
c3: "(c3, s2) \ s'" by auto
from c2 c3
have "(c2;; c3, s1) \ s'" by (rule Seq)
with c1
show "(c1;; (c2;; c3), s) \ s'" by (rule Seq)
next
\<comment> \<open>The other direction is analogous\<close>
assume "(c1;; (c2;; c3), s) \ s'"
thus "(c1;; c2;; c3, s) \ s'" by auto
qed
subsection "Command Equivalence"
text \<open>
We call two statements \<open>c\<close> and \<open>c'\<close> equivalent wrt.\ the
big-step semantics when \emph{\<open>c\<close> started in \<open>s\<close> terminates
in \<open>s'\<close> iff \<open>c'\<close> started in the same \<open>s\<close> also terminates
in the same \<open>s'\<close>}. Formally:
\<close>
text_raw\<open>\snip{BigStepEquiv}{0}{1}{%\<close>
abbreviation
equiv_c :: "com \ com \ bool" (infix "\" 50) where
"c \ c' \ (\s t. (c,s) \ t = (c',s) \ t)"
text_raw\<open>}%endsnip\<close>
text \<open>
Warning: \<open>\<sim>\<close> is the symbol written \verb!\ < s i m >! (without spaces).
As an example, we show that loop unfolding is an equivalence
transformation on programs:
\<close>
lemma unfold_while:
"(WHILE b DO c) \ (IF b THEN c;; WHILE b DO c ELSE SKIP)" (is "?w \ ?iw")
proof -
\<comment> \<open>to show the equivalence, we look at the derivation tree for\<close>
\<comment> \<open>each side and from that construct a derivation tree for the other side\<close>
have "(?iw, s) \ t" if assm: "(?w, s) \ t" for s t
proof -
from assm show ?thesis
proof cases \<comment> \<open>rule inversion on \<open>(?w, s) \<Rightarrow> t\<close>\<close>
case WhileFalse
thus ?thesis by blast
next
case WhileTrue
from \<open>bval b s\<close> \<open>(?w, s) \<Rightarrow> t\<close> obtain s' where
"(c, s) \ s'" and "(?w, s') \ t" by auto
\<comment> \<open>now we can build a derivation tree for the \<^text>\<open>IF\<close>\<close>
\<comment> \<open>first, the body of the True-branch:\<close>
hence "(c;; ?w, s) \ t" by (rule Seq)
\<comment> \<open>then the whole \<^text>\<open>IF\<close>\<close>
with \<open>bval b s\<close> show ?thesis by (rule IfTrue)
qed
qed
moreover
\<comment> \<open>now the other direction:\<close>
have "(?w, s) \ t" if assm: "(?iw, s) \ t" for s t
proof -
from assm show ?thesis
proof cases \<comment> \<open>rule inversion on \<open>(?iw, s) \<Rightarrow> t\<close>\<close>
case IfFalse
hence "s = t" using \<open>(?iw, s) \<Rightarrow> t\<close> by blast
thus ?thesis using \<open>\<not>bval b s\<close> by blast
next
case IfTrue
\<comment> \<open>and for this, only the Seq-rule is applicable:\<close>
from \<open>(c;; ?w, s) \<Rightarrow> t\<close> obtain s' where
"(c, s) \ s'" and "(?w, s') \ t" by auto
\<comment> \<open>with this information, we can build a derivation tree for \<^text>\<open>WHILE\<close>\<close>
with \<open>bval b s\<close> show ?thesis by (rule WhileTrue)
qed
qed
ultimately
show ?thesis by blast
qed
text \<open>Luckily, such lengthy proofs are seldom necessary. Isabelle can
prove many such facts automatically.\<close>
lemma while_unfold:
"(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)
\<sim>
(IF b2 THEN (IF b1 THEN c11 ELSE c2) ELSE (IF b1 THEN c12 ELSE c2))"
by blast
lemma sim_while_cong_aux:
"(WHILE b DO c,s) \ t \ c \ c' \ (WHILE b DO c',s) \ t"
apply(induction "WHILE b DO c" s t arbitrary: b c rule: big_step_induct)
apply blast
apply blast
done
lemma sim_while_cong: "c \ c' \ WHILE b DO c \ WHILE b DO c'"
by (metis sim_while_cong_aux)
text \<open>Command equivalence is an equivalence relation, i.e.\ it is
reflexive, symmetric, and transitive. Because we used an abbreviation
above, Isabelle derives this automatically.\<close>
lemma sim_refl: "c \ c" by simp
lemma sim_sym: "(c \ c') = (c' \ c)" by auto
lemma sim_trans: "c \ c' \ c' \ c'' \ c \ c''" by auto
subsection "Execution is deterministic"
text \<open>This proof is automatic.\<close>
theorem big_step_determ: "\ (c,s) \ t; (c,s) \ u \ \ u = t"
by (induction arbitrary: u rule: big_step.induct) blast+
text \<open>
This is the proof as you might present it in a lecture. The remaining
cases are simple enough to be proved automatically:
\<close>
text_raw\<open>\snip{BigStepDetLong}{0}{2}{%\<close>
theorem
"(c,s) \ t \ (c,s) \ t' \ t' = t"
proof (induction arbitrary: t' rule: big_step.induct)
\<comment> \<open>the only interesting case, \<^text>\<open>WhileTrue\<close>:\<close>
fix b c s s\<^sub>1 t t'
\<comment> \<open>The assumptions of the rule:\<close>
assume "bval b s" and "(c,s) \ s\<^sub>1" and "(WHILE b DO c,s\<^sub>1) \ t"
\<comment> \<open>Ind.Hyp; note the \<^text>\<open>\<And>\<close> because of arbitrary:\<close>
assume IHc: "\t'. (c,s) \ t' \ t' = s\<^sub>1"
assume IHw: "\t'. (WHILE b DO c,s\<^sub>1) \ t' \ t' = t"
\<comment> \<open>Premise of implication:\<close>
assume "(WHILE b DO c,s) \ t'"
with \<open>bval b s\<close> obtain s\<^sub>1' where
c: "(c,s) \ s\<^sub>1'" and
w: "(WHILE b DO c,s\<^sub>1') \ t'"
by auto
from c IHc have "s\<^sub>1' = s\<^sub>1" by blast
with w IHw show "t' = t" by blast
qed blast+ \<comment> \<open>prove the rest automatically\<close>
text_raw\<open>}%endsnip\<close>
end
¤ Dauer der Verarbeitung: 0.2 Sekunden
(vorverarbeitet)
¤
|
Haftungshinweis
Die Informationen auf dieser Webseite wurden
nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit,
noch Qualität der bereit gestellten Informationen zugesichert.
Bemerkung:
Die farbliche Syntaxdarstellung ist noch experimentell.
|