# HG changeset patch # User nipkow # Date 1306957849 -7200 # Node ID 2a05c1f7c08c18ce2c383405c021f8c70b410b76 # Parent 11fce8564415cb6955fe852f2c9ea15e2a90639d Removed old IMP files diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Compiler0.thy --- a/src/HOL/IMP/Compiler0.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,262 +0,0 @@ -(* Title: HOL/IMP/Compiler0.thy - Author: Tobias Nipkow, TUM - Copyright 1996 TUM - -This is an early version of the compiler, where the abstract machine -has an explicit pc. This turned out to be awkward, and a second -development was started. See Machines.thy and Compiler.thy. -*) - -header "A Simple Compiler" - -theory Compiler0 imports Natural begin - -subsection "An abstract, simplistic machine" - -text {* There are only three instructions: *} -datatype instr = ASIN loc aexp | JMPF bexp nat | JMPB nat - -text {* We describe execution of programs in the machine by - an operational (small step) semantics: -*} - -inductive_set - stepa1 :: "instr list \ ((state\nat) \ (state\nat))set" - and stepa1' :: "[instr list,state,nat,state,nat] \ bool" - ("_ \ (3\_,_\/ -1\ \_,_\)" [50,0,0,0,0] 50) - for P :: "instr list" -where - "P \ \s,m\ -1\ \t,n\ == ((s,m),t,n) : stepa1 P" -| ASIN[simp]: - "\ n \ P \ \s,n\ -1\ \s[x\ a s],Suc n\" -| JMPFT[simp,intro]: - "\ n \ P \ \s,n\ -1\ \s,Suc n\" -| JMPFF[simp,intro]: - "\ n \ P \ \s,n\ -1\ \s,m\" -| JMPB[simp]: - "\ n \ P \ \s,n\ -1\ \s,j\" - -abbreviation - stepa :: "[instr list,state,nat,state,nat] \ bool" - ("_ \/ (3\_,_\/ -*\ \_,_\)" [50,0,0,0,0] 50) where - "P \ \s,m\ -*\ \t,n\ == ((s,m),t,n) : ((stepa1 P)^*)" - -abbreviation - stepan :: "[instr list,state,nat,nat,state,nat] \ bool" - ("_ \/ (3\_,_\/ -(_)\ \_,_\)" [50,0,0,0,0,0] 50) where - "P \ \s,m\ -(i)\ \t,n\ == ((s,m),t,n) : (stepa1 P ^^ i)" - -subsection "The compiler" - -primrec compile :: "com \ instr list" where -"compile \ = []" | -"compile (x:==a) = [ASIN x a]" | -"compile (c1;c2) = compile c1 @ compile c2" | -"compile (\ b \ c1 \ c2) = - [JMPF b (length(compile c1) + 2)] @ compile c1 @ - [JMPF (%x. False) (length(compile c2)+1)] @ compile c2" | -"compile (\ b \ c) = [JMPF b (length(compile c) + 2)] @ compile c @ - [JMPB (length(compile c)+1)]" - -declare nth_append[simp] - -subsection "Context lifting lemmas" - -text {* - Some lemmas for lifting an execution into a prefix and suffix - of instructions; only needed for the first proof. -*} -lemma app_right_1: - assumes "is1 \ \s1,i1\ -1\ \s2,i2\" - shows "is1 @ is2 \ \s1,i1\ -1\ \s2,i2\" - using assms - by induct auto - -lemma app_left_1: - assumes "is2 \ \s1,i1\ -1\ \s2,i2\" - shows "is1 @ is2 \ \s1,size is1+i1\ -1\ \s2,size is1+i2\" - using assms - by induct auto - -declare rtrancl_induct2 [induct set: rtrancl] - -lemma app_right: - assumes "is1 \ \s1,i1\ -*\ \s2,i2\" - shows "is1 @ is2 \ \s1,i1\ -*\ \s2,i2\" - using assms -proof induct - show "is1 @ is2 \ \s1,i1\ -*\ \s1,i1\" by simp -next - fix s1' i1' s2 i2 - assume "is1 @ is2 \ \s1,i1\ -*\ \s1',i1'\" - and "is1 \ \s1',i1'\ -1\ \s2,i2\" - thus "is1 @ is2 \ \s1,i1\ -*\ \s2,i2\" - by (blast intro: app_right_1 rtrancl_trans) -qed - -lemma app_left: - assumes "is2 \ \s1,i1\ -*\ \s2,i2\" - shows "is1 @ is2 \ \s1,size is1+i1\ -*\ \s2,size is1+i2\" -using assms -proof induct - show "is1 @ is2 \ \s1,length is1 + i1\ -*\ \s1,length is1 + i1\" by simp -next - fix s1' i1' s2 i2 - assume "is1 @ is2 \ \s1,length is1 + i1\ -*\ \s1',length is1 + i1'\" - and "is2 \ \s1',i1'\ -1\ \s2,i2\" - thus "is1 @ is2 \ \s1,length is1 + i1\ -*\ \s2,length is1 + i2\" - by (blast intro: app_left_1 rtrancl_trans) -qed - -lemma app_left2: - "\ is2 \ \s1,i1\ -*\ \s2,i2\; j1 = size is1+i1; j2 = size is1+i2 \ \ - is1 @ is2 \ \s1,j1\ -*\ \s2,j2\" - by (simp add: app_left) - -lemma app1_left: - assumes "is \ \s1,i1\ -*\ \s2,i2\" - shows "instr # is \ \s1,Suc i1\ -*\ \s2,Suc i2\" -proof - - from app_left [OF assms, of "[instr]"] - show ?thesis by simp -qed - -subsection "Compiler correctness" - -declare rtrancl_into_rtrancl[trans] - converse_rtrancl_into_rtrancl[trans] - rtrancl_trans[trans] - -text {* - The first proof; The statement is very intuitive, - but application of induction hypothesis requires the above lifting lemmas -*} -theorem - assumes "\c,s\ \\<^sub>c t" - shows "compile c \ \s,0\ -*\ \t,length(compile c)\" (is "?P c s t") - using assms -proof induct - show "\s. ?P \ s s" by simp -next - show "\a s x. ?P (x :== a) s (s[x\ a s])" by force -next - fix c0 c1 s0 s1 s2 - assume "?P c0 s0 s1" - hence "compile c0 @ compile c1 \ \s0,0\ -*\ \s1,length(compile c0)\" - by (rule app_right) - moreover assume "?P c1 s1 s2" - hence "compile c0 @ compile c1 \ \s1,length(compile c0)\ -*\ - \s2,length(compile c0)+length(compile c1)\" - proof - - show "\is1 is2 s1 s2 i2. - is2 \ \s1,0\ -*\ \s2,i2\ \ - is1 @ is2 \ \s1,size is1\ -*\ \s2,size is1+i2\" - using app_left[of _ 0] by simp - qed - ultimately have "compile c0 @ compile c1 \ \s0,0\ -*\ - \s2,length(compile c0)+length(compile c1)\" - by (rule rtrancl_trans) - thus "?P (c0; c1) s0 s2" by simp -next - fix b c0 c1 s0 s1 - let ?comp = "compile(\ b \ c0 \ c1)" - assume "b s0" and IH: "?P c0 s0 s1" - hence "?comp \ \s0,0\ -1\ \s0,1\" by auto - also from IH - have "?comp \ \s0,1\ -*\ \s1,length(compile c0)+1\" - by(auto intro:app1_left app_right) - also have "?comp \ \s1,length(compile c0)+1\ -1\ \s1,length ?comp\" - by(auto) - finally show "?P (\ b \ c0 \ c1) s0 s1" . -next - fix b c0 c1 s0 s1 - let ?comp = "compile(\ b \ c0 \ c1)" - assume "\b s0" and IH: "?P c1 s0 s1" - hence "?comp \ \s0,0\ -1\ \s0,length(compile c0) + 2\" by auto - also from IH - have "?comp \ \s0,length(compile c0)+2\ -*\ \s1,length ?comp\" - by (force intro!: app_left2 app1_left) - finally show "?P (\ b \ c0 \ c1) s0 s1" . -next - fix b c and s::state - assume "\b s" - thus "?P (\ b \ c) s s" by force -next - fix b c and s0::state and s1 s2 - let ?comp = "compile(\ b \ c)" - assume "b s0" and - IHc: "?P c s0 s1" and IHw: "?P (\ b \ c) s1 s2" - hence "?comp \ \s0,0\ -1\ \s0,1\" by auto - also from IHc - have "?comp \ \s0,1\ -*\ \s1,length(compile c)+1\" - by (auto intro: app1_left app_right) - also have "?comp \ \s1,length(compile c)+1\ -1\ \s1,0\" by simp - also note IHw - finally show "?P (\ b \ c) s0 s2". -qed - -text {* - Second proof; statement is generalized to cater for prefixes and suffixes; - needs none of the lifting lemmas, but instantiations of pre/suffix. - *} -(* -theorem assumes A: "\c,s\ \\<^sub>c t" -shows "\a z. a@compile c@z \ \s,size a\ -*\ \t,size a + size(compile c)\" - (is "\a z. ?P c s t a z") -proof - - from A show "\a z. ?thesis a z" - proof induct - case Skip thus ?case by simp - next - case Assign thus ?case by (force intro!: ASIN) - next - fix c1 c2 s s' s'' a z - assume IH1: "\a z. ?P c1 s s' a z" and IH2: "\a z. ?P c2 s' s'' a z" - from IH1 IH2[of "a@compile c1"] - show "?P (c1;c2) s s'' a z" - by(simp add:add_assoc[THEN sym])(blast intro:rtrancl_trans) - next -(* at this point I gave up converting to structured proofs *) -(* \ b \ c0 \ c1; case b is true *) - apply(intro strip) - (* instantiate assumption sufficiently for later: *) - apply(erule_tac x = "a@[?I]" in allE) - apply(simp) - (* execute JMPF: *) - apply(rule converse_rtrancl_into_rtrancl) - apply(force intro!: JMPFT) - (* execute compile c0: *) - apply(rule rtrancl_trans) - apply(erule allE) - apply assumption - (* execute JMPF: *) - apply(rule r_into_rtrancl) - apply(force intro!: JMPFF) -(* end of case b is true *) - apply(intro strip) - apply(erule_tac x = "a@[?I]@compile c0@[?J]" in allE) - apply(simp add:add_assoc) - apply(rule converse_rtrancl_into_rtrancl) - apply(force intro!: JMPFF) - apply(blast) - apply(force intro: JMPFF) -apply(intro strip) -apply(erule_tac x = "a@[?I]" in allE) -apply(erule_tac x = a in allE) -apply(simp) -apply(rule converse_rtrancl_into_rtrancl) - apply(force intro!: JMPFT) -apply(rule rtrancl_trans) - apply(erule allE) - apply assumption -apply(rule converse_rtrancl_into_rtrancl) - apply(force intro!: JMPB) -apply(simp) -done -*) -text {* Missing: the other direction! I did much of it, and although -the main lemma is very similar to the one in the new development, the -lemmas surrounding it seemed much more complicated. In the end I gave -up. *} - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Examples.thy --- a/src/HOL/IMP/Examples.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,81 +0,0 @@ -(* Title: HOL/IMP/Examples.thy - Author: David von Oheimb, TUM -*) - -header "Examples" - -theory Examples imports Natural begin - -definition - factorial :: "loc => loc => com" where - "factorial a b = (b :== (%s. 1); - \ (%s. s a ~= 0) \ - (b :== (%s. s b * s a); a :== (%s. s a - 1)))" - -declare update_def [simp] - -subsection "An example due to Tony Hoare" - -lemma lemma1: - assumes 1: "!x. P x \ Q x" - and 2: "\w,s\ \\<^sub>c t" - shows "w = While P c \ \While Q c,t\ \\<^sub>c u \ \While Q c,s\ \\<^sub>c u" - using 2 apply induct - using 1 apply auto - done - -lemma lemma2 [rule_format (no_asm)]: - "[| !x. P x \ Q x; \w,s\ \\<^sub>c u |] ==> - !c. w = While Q c \ \While P c; While Q c,s\ \\<^sub>c u" -apply (erule evalc.induct) -apply (simp_all (no_asm_simp)) -apply blast -apply (case_tac "P s") -apply auto -done - -lemma Hoare_example: "!x. P x \ Q x ==> - (\While P c; While Q c, s\ \\<^sub>c t) = (\While Q c, s\ \\<^sub>c t)" - by (blast intro: lemma1 lemma2 dest: semi [THEN iffD1]) - - -subsection "Factorial" - -lemma factorial_3: "a~=b ==> - \factorial a b, Mem(a:=3)\ \\<^sub>c Mem(b:=6, a:=0)" - by (simp add: factorial_def) - -text {* the same in single step mode: *} -lemmas [simp del] = evalc_cases -lemma "a~=b \ \factorial a b, Mem(a:=3)\ \\<^sub>c Mem(b:=6, a:=0)" -apply (unfold factorial_def) -apply (frule not_sym) -apply (rule evalc.intros) -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply simp -apply (rule evalc.intros) -apply simp -done - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Expr.thy --- a/src/HOL/IMP/Expr.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,106 +0,0 @@ -(* Title: HOL/IMP/Expr.thy - Author: Heiko Loetzbeyer & Robert Sandner & Tobias Nipkow, TUM -*) - -header "Expressions" - -theory Expr imports Main begin - -text {* - Arithmetic expressions and Boolean expressions. - Not used in the rest of the language, but included for completeness. -*} - -subsection "Arithmetic expressions" -typedecl loc - -type_synonym state = "loc => nat" - -datatype - aexp = N nat - | X loc - | Op1 "nat => nat" aexp - | Op2 "nat => nat => nat" aexp aexp - -subsection "Evaluation of arithmetic expressions" - -inductive - evala :: "[aexp*state,nat] => bool" (infixl "-a->" 50) -where - N: "(N(n),s) -a-> n" -| X: "(X(x),s) -a-> s(x)" -| Op1: "(e,s) -a-> n ==> (Op1 f e,s) -a-> f(n)" -| Op2: "[| (e0,s) -a-> n0; (e1,s) -a-> n1 |] - ==> (Op2 f e0 e1,s) -a-> f n0 n1" - -lemmas [intro] = N X Op1 Op2 - - -subsection "Boolean expressions" - -datatype - bexp = true - | false - | ROp "nat => nat => bool" aexp aexp - | noti bexp - | andi bexp bexp (infixl "andi" 60) - | ori bexp bexp (infixl "ori" 60) - -subsection "Evaluation of boolean expressions" - -inductive - evalb :: "[bexp*state,bool] => bool" (infixl "-b->" 50) - -- "avoid clash with ML constructors true, false" -where - tru: "(true,s) -b-> True" -| fls: "(false,s) -b-> False" -| ROp: "[| (a0,s) -a-> n0; (a1,s) -a-> n1 |] - ==> (ROp f a0 a1,s) -b-> f n0 n1" -| noti: "(b,s) -b-> w ==> (noti(b),s) -b-> (~w)" -| andi: "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |] - ==> (b0 andi b1,s) -b-> (w0 & w1)" -| ori: "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |] - ==> (b0 ori b1,s) -b-> (w0 | w1)" - -lemmas [intro] = tru fls ROp noti andi ori - -subsection "Denotational semantics of arithmetic and boolean expressions" - -primrec A :: "aexp => state => nat" -where - "A(N(n)) = (%s. n)" -| "A(X(x)) = (%s. s(x))" -| "A(Op1 f a) = (%s. f(A a s))" -| "A(Op2 f a0 a1) = (%s. f (A a0 s) (A a1 s))" - -primrec B :: "bexp => state => bool" -where - "B(true) = (%s. True)" -| "B(false) = (%s. False)" -| "B(ROp f a0 a1) = (%s. f (A a0 s) (A a1 s))" -| "B(noti(b)) = (%s. ~(B b s))" -| "B(b0 andi b1) = (%s. (B b0 s) & (B b1 s))" -| "B(b0 ori b1) = (%s. (B b0 s) | (B b1 s))" - -inductive_simps - evala_simps [simp]: - "(N(n),s) -a-> n'" - "(X(x),sigma) -a-> i" - "(Op1 f e,sigma) -a-> i" - "(Op2 f a1 a2,sigma) -a-> i" - "((true,sigma) -b-> w)" - "((false,sigma) -b-> w)" - "((ROp f a0 a1,sigma) -b-> w)" - "((noti(b),sigma) -b-> w)" - "((b0 andi b1,sigma) -b-> w)" - "((b0 ori b1,sigma) -b-> w)" - - -lemma aexp_iff: "((a,s) -a-> n) = (A a s = n)" - by (induct a arbitrary: n) auto - -lemma bexp_iff: - "((b,s) -b-> w) = (B b s = w)" - by (induct b arbitrary: w) (auto simp add: aexp_iff) - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Hoare.thy --- a/src/HOL/IMP/Hoare.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,40 +0,0 @@ -(* Title: HOL/IMP/Hoare.thy - Author: Tobias Nipkow -*) - -header "Inductive Definition of Hoare Logic" - -theory Hoare imports Natural begin - -type_synonym assn = "state => bool" - -inductive - hoare :: "assn => com => assn => bool" ("|- ({(1_)}/ (_)/ {(1_)})" 50) -where - skip: "|- {P}\{P}" -| ass: "|- {%s. P(s[x\a s])} x:==a {P}" -| semi: "[| |- {P}c{Q}; |- {Q}d{R} |] ==> |- {P} c;d {R}" -| If: "[| |- {%s. P s & b s}c{Q}; |- {%s. P s & ~b s}d{Q} |] ==> - |- {P} \ b \ c \ d {Q}" -| While: "|- {%s. P s & b s} c {P} ==> - |- {P} \ b \ c {%s. P s & ~b s}" -| conseq: "[| !s. P' s --> P s; |- {P}c{Q}; !s. Q s --> Q' s |] ==> - |- {P'}c{Q'}" - -lemma strengthen_pre: "[| !s. P' s --> P s; |- {P}c{Q} |] ==> |- {P'}c{Q}" -by (blast intro: conseq) - -lemma weaken_post: "[| |- {P}c{Q}; !s. Q s --> Q' s |] ==> |- {P}c{Q'}" -by (blast intro: conseq) - -lemma While': -assumes "|- {%s. P s & b s} c {P}" and "ALL s. P s & \ b s \ Q s" -shows "|- {P} \ b \ c {Q}" -by(rule weaken_post[OF While[OF assms(1)] assms(2)]) - - -lemmas [simp] = skip ass semi If - -lemmas [intro!] = hoare.skip hoare.ass hoare.semi hoare.If - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Hoare_Den.thy --- a/src/HOL/IMP/Hoare_Den.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,133 +0,0 @@ -(* Title: HOL/IMP/Hoare_Den.thy - Author: Tobias Nipkow -*) - -header "Soundness and Completeness wrt Denotational Semantics" - -theory Hoare_Den imports Hoare Denotation begin - -definition - hoare_valid :: "[assn,com,assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" 50) where - "|= {P}c{Q} = (!s t. (s,t) : C(c) --> P s --> Q t)" - - -lemma hoare_sound: "|- {P}c{Q} ==> |= {P}c{Q}" -proof(induct rule: hoare.induct) - case (While P b c) - { fix s t - let ?G = "Gamma b (C c)" - assume "(s,t) \ lfp ?G" - hence "P s \ P t \ \ b t" - proof(rule lfp_induct2) - show "mono ?G" by(rule Gamma_mono) - next - fix s t assume "(s,t) \ ?G (lfp ?G \ {(s,t). P s \ P t \ \ b t})" - thus "P s \ P t \ \ b t" using While.hyps - by(auto simp: hoare_valid_def Gamma_def) - qed - } - thus ?case by(simp add:hoare_valid_def) -qed (auto simp: hoare_valid_def) - - -definition - wp :: "com => assn => assn" where - "wp c Q = (%s. !t. (s,t) : C(c) --> Q t)" - -lemma wp_SKIP: "wp \ Q = Q" -by (simp add: wp_def) - -lemma wp_Ass: "wp (x:==a) Q = (%s. Q(s[x\a s]))" -by (simp add: wp_def) - -lemma wp_Semi: "wp (c;d) Q = wp c (wp d Q)" -by (rule ext) (auto simp: wp_def) - -lemma wp_If: - "wp (\ b \ c \ d) Q = (%s. (b s --> wp c Q s) & (~b s --> wp d Q s))" -by (rule ext) (auto simp: wp_def) - -lemma wp_While_If: - "wp (\ b \ c) Q s = - wp (IF b THEN c;\ b \ c ELSE SKIP) Q s" -by(simp only: wp_def C_While_If) - -(*Not suitable for rewriting: LOOPS!*) -lemma wp_While_if: - "wp (\ b \ c) Q s = (if b s then wp (c;\ b \ c) Q s else Q s)" -by(simp add:wp_While_If wp_If wp_SKIP) - -lemma wp_While_True: "b s ==> - wp (\ b \ c) Q s = wp (c;\ b \ c) Q s" -by(simp add: wp_While_if) - -lemma wp_While_False: "~b s ==> wp (\ b \ c) Q s = Q s" -by(simp add: wp_While_if) - -lemmas [simp] = wp_SKIP wp_Ass wp_Semi wp_If wp_While_True wp_While_False - -lemma wp_While: "wp (\ b \ c) Q s = - (s : gfp(%S.{s. if b s then wp c (%s. s:S) s else Q s}))" -apply (simp (no_asm)) -apply (rule iffI) - apply (rule weak_coinduct) - apply (erule CollectI) - apply safe - apply simp - apply simp -apply (simp add: wp_def Gamma_def) -apply (intro strip) -apply (rule mp) - prefer 2 apply (assumption) -apply (erule lfp_induct2) -apply (fast intro!: monoI) -apply (subst gfp_unfold) - apply (fast intro!: monoI) -apply fast -done - -declare C_while [simp del] - -lemma wp_is_pre: "|- {wp c Q} c {Q}" -proof(induct c arbitrary: Q) - case SKIP show ?case by auto -next - case Assign show ?case by auto -next - case Semi thus ?case by auto -next - case (Cond b c1 c2) - let ?If = "IF b THEN c1 ELSE c2" - show ?case - proof(rule If) - show "|- {\s. wp ?If Q s \ b s} c1 {Q}" - proof(rule strengthen_pre[OF _ Cond(1)]) - show "\s. wp ?If Q s \ b s \ wp c1 Q s" by auto - qed - show "|- {\s. wp ?If Q s \ \ b s} c2 {Q}" - proof(rule strengthen_pre[OF _ Cond(2)]) - show "\s. wp ?If Q s \ \ b s \ wp c2 Q s" by auto - qed - qed -next - case (While b c) - let ?w = "WHILE b DO c" - show ?case - proof(rule While') - show "|- {\s. wp ?w Q s \ b s} c {wp ?w Q}" - proof(rule strengthen_pre[OF _ While(1)]) - show "\s. wp ?w Q s \ b s \ wp c (wp ?w Q) s" by auto - qed - show "\s. wp ?w Q s \ \ b s \ Q s" by auto - qed -qed - -lemma hoare_relative_complete: assumes "|= {P}c{Q}" shows "|- {P}c{Q}" -proof(rule conseq) - show "\s. P s \ wp c Q s" using assms - by (auto simp: hoare_valid_def wp_def) - show "|- {wp c Q} c {Q}" by(rule wp_is_pre) - show "\s. Q s \ Q s" by auto -qed - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Hoare_Op.thy --- a/src/HOL/IMP/Hoare_Op.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,105 +0,0 @@ -(* Title: HOL/IMP/Hoare_Op.thy - Author: Tobias Nipkow -*) - -header "Soundness and Completeness wrt Operational Semantics" - -theory Hoare_Op imports Hoare begin - -definition - hoare_valid :: "[assn,com,assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" 50) where - "|= {P}c{Q} = (!s t. \c,s\ \\<^sub>c t --> P s --> Q t)" - -lemma hoare_sound: "|- {P}c{Q} ==> |= {P}c{Q}" -proof(induct rule: hoare.induct) - case (While P b c) - { fix s t - assume "\WHILE b DO c,s\ \\<^sub>c t" - hence "P s \ P t \ \ b t" - proof(induct "WHILE b DO c" s t) - case WhileFalse thus ?case by blast - next - case WhileTrue thus ?case - using While(2) unfolding hoare_valid_def by blast - qed - - } - thus ?case unfolding hoare_valid_def by blast -qed (auto simp: hoare_valid_def) - -definition - wp :: "com => assn => assn" where - "wp c Q = (%s. !t. \c,s\ \\<^sub>c t --> Q t)" - -lemma wp_SKIP: "wp \ Q = Q" -by (simp add: wp_def) - -lemma wp_Ass: "wp (x:==a) Q = (%s. Q(s[x\a s]))" -by (simp add: wp_def) - -lemma wp_Semi: "wp (c;d) Q = wp c (wp d Q)" -by (rule ext) (auto simp: wp_def) - -lemma wp_If: - "wp (\ b \ c \ d) Q = (%s. (b s --> wp c Q s) & (~b s --> wp d Q s))" -by (rule ext) (auto simp: wp_def) - -lemma wp_While_If: - "wp (\ b \ c) Q s = - wp (IF b THEN c;\ b \ c ELSE SKIP) Q s" -unfolding wp_def by (metis equivD1 equivD2 unfold_while) - -lemma wp_While_True: "b s ==> - wp (\ b \ c) Q s = wp (c;\ b \ c) Q s" -by(simp add: wp_While_If wp_If wp_SKIP) - -lemma wp_While_False: "~b s ==> wp (\ b \ c) Q s = Q s" -by(simp add: wp_While_If wp_If wp_SKIP) - -lemmas [simp] = wp_SKIP wp_Ass wp_Semi wp_If wp_While_True wp_While_False - -lemma wp_is_pre: "|- {wp c Q} c {Q}" -proof(induct c arbitrary: Q) - case SKIP show ?case by auto -next - case Assign show ?case by auto -next - case Semi thus ?case by(auto intro: semi) -next - case (Cond b c1 c2) - let ?If = "IF b THEN c1 ELSE c2" - show ?case - proof(rule If) - show "|- {\s. wp ?If Q s \ b s} c1 {Q}" - proof(rule strengthen_pre[OF _ Cond(1)]) - show "\s. wp ?If Q s \ b s \ wp c1 Q s" by auto - qed - show "|- {\s. wp ?If Q s \ \ b s} c2 {Q}" - proof(rule strengthen_pre[OF _ Cond(2)]) - show "\s. wp ?If Q s \ \ b s \ wp c2 Q s" by auto - qed - qed -next - case (While b c) - let ?w = "WHILE b DO c" - have "|- {wp ?w Q} ?w {\s. wp ?w Q s \ \ b s}" - proof(rule hoare.While) - show "|- {\s. wp ?w Q s \ b s} c {wp ?w Q}" - proof(rule strengthen_pre[OF _ While(1)]) - show "\s. wp ?w Q s \ b s \ wp c (wp ?w Q) s" by auto - qed - qed - thus ?case - proof(rule weaken_post) - show "\s. wp ?w Q s \ \ b s \ Q s" by auto - qed -qed - -lemma hoare_relative_complete: assumes "|= {P}c{Q}" shows "|- {P}c{Q}" -proof(rule strengthen_pre) - show "\s. P s \ wp c Q s" using assms - by (auto simp: hoare_valid_def wp_def) - show "|- {wp c Q} c {Q}" by(rule wp_is_pre) -qed - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Live.thy --- a/src/HOL/IMP/Live.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,212 +0,0 @@ -theory Live imports Natural -begin - -text{* Which variables/locations does an expression depend on? -Any set of variables that completely determine the value of the expression, -in the worst case all locations: *} - -consts Dep :: "((loc \ 'a) \ 'b) \ loc set" -specification (Dep) -dep_on: "(\x\Dep e. s x = t x) \ e s = e t" -by(rule_tac x="%x. UNIV" in exI)(simp add: fun_eq_iff[symmetric]) - -text{* The following definition of @{const Dep} looks very tempting -@{prop"Dep e = {a. EX s t. (ALL x. x\a \ s x = t x) \ e s \ e t}"} -but does not work in case @{text e} depends on an infinite set of variables. -For example, if @{term"e s"} tests if @{text s} is 0 at infinitely many locations. Then @{term"Dep e"} incorrectly yields the empty set! - -If we had a concrete representation of expressions, we would simply write -a recursive free-variables function. -*} - -primrec L :: "com \ loc set \ loc set" where -"L SKIP A = A" | -"L (x :== e) A = A-{x} \ Dep e" | -"L (c1; c2) A = (L c1 \ L c2) A" | -"L (IF b THEN c1 ELSE c2) A = Dep b \ L c1 A \ L c2 A" | -"L (WHILE b DO c) A = Dep b \ A \ L c A" - -primrec "kill" :: "com \ loc set" where -"kill SKIP = {}" | -"kill (x :== e) = {x}" | -"kill (c1; c2) = kill c1 \ kill c2" | -"kill (IF b THEN c1 ELSE c2) = Dep b \ kill c1 \ kill c2" | -"kill (WHILE b DO c) = {}" - -primrec gen :: "com \ loc set" where -"gen SKIP = {}" | -"gen (x :== e) = Dep e" | -"gen (c1; c2) = gen c1 \ (gen c2-kill c1)" | -"gen (IF b THEN c1 ELSE c2) = Dep b \ gen c1 \ gen c2" | -"gen (WHILE b DO c) = Dep b \ gen c" - -lemma L_gen_kill: "L c A = gen c \ (A - kill c)" -by(induct c arbitrary:A) auto - -lemma L_idemp: "L c (L c A) \ L c A" -by(fastsimp simp add:L_gen_kill) - -theorem L_sound: "\ x \ L c A. s x = t x \ \c,s\ \\<^sub>c s' \ \c,t\ \\<^sub>c t' \ - \x\A. s' x = t' x" -proof (induct c arbitrary: A s t s' t') - case SKIP then show ?case by auto -next - case (Assign x e) then show ?case - by (auto simp:update_def ball_Un dest!: dep_on) -next - case (Semi c1 c2) - from Semi(4) obtain s'' where s1: "\c1,s\ \\<^sub>c s''" and s2: "\c2,s''\ \\<^sub>c s'" - by auto - from Semi(5) obtain t'' where t1: "\c1,t\ \\<^sub>c t''" and t2: "\c2,t''\ \\<^sub>c t'" - by auto - show ?case using Semi(1)[OF _ s1 t1] Semi(2)[OF _ s2 t2] Semi(3) by fastsimp -next - case (Cond b c1 c2) - show ?case - proof cases - assume "b s" - hence s: "\c1,s\ \\<^sub>c s'" using Cond(4) by simp - have "b t" using `b s` Cond(3) by (simp add: ball_Un)(blast dest: dep_on) - hence t: "\c1,t\ \\<^sub>c t'" using Cond(5) by auto - show ?thesis using Cond(1)[OF _ s t] Cond(3) by fastsimp - next - assume "\ b s" - hence s: "\c2,s\ \\<^sub>c s'" using Cond(4) by auto - have "\ b t" using `\ b s` Cond(3) by (simp add: ball_Un)(blast dest: dep_on) - hence t: "\c2,t\ \\<^sub>c t'" using Cond(5) by auto - show ?thesis using Cond(2)[OF _ s t] Cond(3) by fastsimp - qed -next - case (While b c) note IH = this - { fix cw - have "\cw,s\ \\<^sub>c s' \ cw = (While b c) \ \cw,t\ \\<^sub>c t' \ - \ x \ L cw A. s x = t x \ \x\A. s' x = t' x" - proof (induct arbitrary: t A pred:evalc) - case WhileFalse - have "\ b t" using WhileFalse by (simp add: ball_Un)(blast dest:dep_on) - then have "t' = t" using WhileFalse by auto - then show ?case using WhileFalse by auto - next - case (WhileTrue _ s _ s'' s') - have "\c,s\ \\<^sub>c s''" using WhileTrue(2,6) by simp - have "b t" using WhileTrue by (simp add: ball_Un)(blast dest:dep_on) - then obtain t'' where "\c,t\ \\<^sub>c t''" and "\While b c,t''\ \\<^sub>c t'" - using WhileTrue(6,7) by auto - have "\x\Dep b \ A \ L c A. s'' x = t'' x" - using IH(1)[OF _ `\c,s\ \\<^sub>c s''` `\c,t\ \\<^sub>c t''`] WhileTrue(6,8) - by (auto simp:L_gen_kill) - then have "\x\L (While b c) A. s'' x = t'' x" by auto - then show ?case using WhileTrue(5,6) `\While b c,t''\ \\<^sub>c t'` by metis - qed auto } --- "a terser version" - { let ?w = "While b c" - have "\?w,s\ \\<^sub>c s' \ \?w,t\ \\<^sub>c t' \ - \ x \ L ?w A. s x = t x \ \x\A. s' x = t' x" - proof (induct ?w s s' arbitrary: t A pred:evalc) - case WhileFalse - have "\ b t" using WhileFalse by (simp add: ball_Un)(blast dest:dep_on) - then have "t' = t" using WhileFalse by auto - then show ?case using WhileFalse by simp - next - case (WhileTrue s s'' s') - have "b t" using WhileTrue by (simp add: ball_Un)(blast dest:dep_on) - then obtain t'' where "\c,t\ \\<^sub>c t''" and "\While b c,t''\ \\<^sub>c t'" - using WhileTrue(6,7) by auto - have "\x\Dep b \ A \ L c A. s'' x = t'' x" - using IH(1)[OF _ `\c,s\ \\<^sub>c s''` `\c,t\ \\<^sub>c t''`] WhileTrue(7) - by (auto simp:L_gen_kill) - then have "\x\L (While b c) A. s'' x = t'' x" by auto - then show ?case using WhileTrue(5) `\While b c,t''\ \\<^sub>c t'` by metis - qed } - from this[OF IH(3) IH(4,2)] show ?case by metis -qed - - -primrec bury :: "com \ loc set \ com" where -"bury SKIP _ = SKIP" | -"bury (x :== e) A = (if x:A then x:== e else SKIP)" | -"bury (c1; c2) A = (bury c1 (L c2 A); bury c2 A)" | -"bury (IF b THEN c1 ELSE c2) A = (IF b THEN bury c1 A ELSE bury c2 A)" | -"bury (WHILE b DO c) A = (WHILE b DO bury c (Dep b \ A \ L c A))" - -theorem bury_sound: - "\ x \ L c A. s x = t x \ \c,s\ \\<^sub>c s' \ \bury c A,t\ \\<^sub>c t' \ - \x\A. s' x = t' x" -proof (induct c arbitrary: A s t s' t') - case SKIP then show ?case by auto -next - case (Assign x e) then show ?case - by (auto simp:update_def ball_Un split:split_if_asm dest!: dep_on) -next - case (Semi c1 c2) - from Semi(4) obtain s'' where s1: "\c1,s\ \\<^sub>c s''" and s2: "\c2,s''\ \\<^sub>c s'" - by auto - from Semi(5) obtain t'' where t1: "\bury c1 (L c2 A),t\ \\<^sub>c t''" and t2: "\bury c2 A,t''\ \\<^sub>c t'" - by auto - show ?case using Semi(1)[OF _ s1 t1] Semi(2)[OF _ s2 t2] Semi(3) by fastsimp -next - case (Cond b c1 c2) - show ?case - proof cases - assume "b s" - hence s: "\c1,s\ \\<^sub>c s'" using Cond(4) by simp - have "b t" using `b s` Cond(3) by (simp add: ball_Un)(blast dest: dep_on) - hence t: "\bury c1 A,t\ \\<^sub>c t'" using Cond(5) by auto - show ?thesis using Cond(1)[OF _ s t] Cond(3) by fastsimp - next - assume "\ b s" - hence s: "\c2,s\ \\<^sub>c s'" using Cond(4) by auto - have "\ b t" using `\ b s` Cond(3) by (simp add: ball_Un)(blast dest: dep_on) - hence t: "\bury c2 A,t\ \\<^sub>c t'" using Cond(5) by auto - show ?thesis using Cond(2)[OF _ s t] Cond(3) by fastsimp - qed -next - case (While b c) note IH = this - { fix cw - have "\cw,s\ \\<^sub>c s' \ cw = (While b c) \ \bury cw A,t\ \\<^sub>c t' \ - \ x \ L cw A. s x = t x \ \x\A. s' x = t' x" - proof (induct arbitrary: t A pred:evalc) - case WhileFalse - have "\ b t" using WhileFalse by (simp add: ball_Un)(blast dest:dep_on) - then have "t' = t" using WhileFalse by auto - then show ?case using WhileFalse by auto - next - case (WhileTrue _ s _ s'' s') - have "\c,s\ \\<^sub>c s''" using WhileTrue(2,6) by simp - have "b t" using WhileTrue by (simp add: ball_Un)(blast dest:dep_on) - then obtain t'' where tt'': "\bury c (Dep b \ A \ L c A),t\ \\<^sub>c t''" - and "\bury (While b c) A,t''\ \\<^sub>c t'" - using WhileTrue(6,7) by auto - have "\x\Dep b \ A \ L c A. s'' x = t'' x" - using IH(1)[OF _ `\c,s\ \\<^sub>c s''` tt''] WhileTrue(6,8) - by (auto simp:L_gen_kill) - moreover then have "\x\L (While b c) A. s'' x = t'' x" by auto - ultimately show ?case - using WhileTrue(5,6) `\bury (While b c) A,t''\ \\<^sub>c t'` by metis - qed auto } - { let ?w = "While b c" - have "\?w,s\ \\<^sub>c s' \ \bury ?w A,t\ \\<^sub>c t' \ - \ x \ L ?w A. s x = t x \ \x\A. s' x = t' x" - proof (induct ?w s s' arbitrary: t A pred:evalc) - case WhileFalse - have "\ b t" using WhileFalse by (simp add: ball_Un)(blast dest:dep_on) - then have "t' = t" using WhileFalse by auto - then show ?case using WhileFalse by simp - next - case (WhileTrue s s'' s') - have "b t" using WhileTrue by (simp add: ball_Un)(blast dest:dep_on) - then obtain t'' where tt'': "\bury c (Dep b \ A \ L c A),t\ \\<^sub>c t''" - and "\bury (While b c) A,t''\ \\<^sub>c t'" - using WhileTrue(6,7) by auto - have "\x\Dep b \ A \ L c A. s'' x = t'' x" - using IH(1)[OF _ `\c,s\ \\<^sub>c s''` tt''] WhileTrue(7) - by (auto simp:L_gen_kill) - then have "\x\L (While b c) A. s'' x = t'' x" by auto - then show ?case - using WhileTrue(5) `\bury (While b c) A,t''\ \\<^sub>c t'` by metis - qed } - from this[OF IH(3) IH(4,2)] show ?case by metis -qed - - -end \ No newline at end of file diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Machines.thy --- a/src/HOL/IMP/Machines.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,182 +0,0 @@ -theory Machines -imports Natural -begin - -lemma converse_in_rel_pow_eq: - "((x,z) \ R ^^ n) = (n=0 \ z=x \ (\m y. n = Suc m \ (x,y) \ R \ (y,z) \ R ^^ m))" -apply(rule iffI) - apply(blast elim:rel_pow_E2) -apply (auto simp: rel_pow_commute[symmetric]) -done - -subsection "Instructions" - -text {* There are only three instructions: *} -datatype instr = SET loc aexp | JMPF bexp nat | JMPB nat - -type_synonym instrs = "instr list" - -subsection "M0 with PC" - -inductive_set - exec01 :: "instr list \ ((nat\state) \ (nat\state))set" - and exec01' :: "[instrs, nat,state, nat,state] \ bool" - ("(_/ \ (1\_,/_\)/ -1\ (1\_,/_\))" [50,0,0,0,0] 50) - for P :: "instr list" -where - "p \ \i,s\ -1\ \j,t\ == ((i,s),j,t) : (exec01 p)" -| SET: "\ n \ P \ \n,s\ -1\ \Suc n,s[x\ a s]\" -| JMPFT: "\ n \ P \ \n,s\ -1\ \Suc n,s\" -| JMPFF: "\ nb s; m=n+i+1; m \ size P \ - \ P \ \n,s\ -1\ \m,s\" -| JMPB: "\ n n; j = n-i \ \ P \ \n,s\ -1\ \j,s\" - -abbreviation - exec0s :: "[instrs, nat,state, nat,state] \ bool" - ("(_/ \ (1\_,/_\)/ -*\ (1\_,/_\))" [50,0,0,0,0] 50) where - "p \ \i,s\ -*\ \j,t\ == ((i,s),j,t) : (exec01 p)^*" - -abbreviation - exec0n :: "[instrs, nat,state, nat, nat,state] \ bool" - ("(_/ \ (1\_,/_\)/ -_\ (1\_,/_\))" [50,0,0,0,0] 50) where - "p \ \i,s\ -n\ \j,t\ == ((i,s),j,t) : (exec01 p)^^n" - -subsection "M0 with lists" - -text {* We describe execution of programs in the machine by - an operational (small step) semantics: -*} - -type_synonym config = "instrs \ instrs \ state" - - -inductive_set - stepa1 :: "(config \ config)set" - and stepa1' :: "[instrs,instrs,state, instrs,instrs,state] \ bool" - ("((1\_,/_,/_\)/ -1\ (1\_,/_,/_\))" 50) -where - "\p,q,s\ -1\ \p',q',t\ == ((p,q,s),p',q',t) : stepa1" -| "\SET x a#p,q,s\ -1\ \p,SET x a#q,s[x\ a s]\" -| "b s \ \JMPF b i#p,q,s\ -1\ \p,JMPF b i#q,s\" -| "\ \ b s; i \ size p \ - \ \JMPF b i # p, q, s\ -1\ \drop i p, rev(take i p) @ JMPF b i # q, s\" -| "i \ size q - \ \JMPB i # p, q, s\ -1\ \rev(take i q) @ JMPB i # p, drop i q, s\" - -abbreviation - stepa :: "[instrs,instrs,state, instrs,instrs,state] \ bool" - ("((1\_,/_,/_\)/ -*\ (1\_,/_,/_\))" 50) where - "\p,q,s\ -*\ \p',q',t\ == ((p,q,s),p',q',t) : (stepa1^*)" - -abbreviation - stepan :: "[instrs,instrs,state, nat, instrs,instrs,state] \ bool" - ("((1\_,/_,/_\)/ -_\ (1\_,/_,/_\))" 50) where - "\p,q,s\ -i\ \p',q',t\ == ((p,q,s),p',q',t) : (stepa1^^i)" - -inductive_cases execE: "((i#is,p,s), (is',p',s')) : stepa1" - -lemma exec_simp[simp]: - "(\i#p,q,s\ -1\ \p',q',t\) = (case i of - SET x a \ t = s[x\ a s] \ p' = p \ q' = i#q | - JMPF b n \ t=s \ (if b s then p' = p \ q' = i#q - else n \ size p \ p' = drop n p \ q' = rev(take n p) @ i # q) | - JMPB n \ n \ size q \ t=s \ p' = rev(take n q) @ i # p \ q' = drop n q)" -apply(rule iffI) -defer -apply(clarsimp simp add: stepa1.intros split: instr.split_asm split_if_asm) -apply(erule execE) -apply(simp_all) -done - -lemma execn_simp[simp]: -"(\i#p,q,s\ -n\ \p'',q'',u\) = - (n=0 \ p'' = i#p \ q'' = q \ u = s \ - ((\m p' q' t. n = Suc m \ - \i#p,q,s\ -1\ \p',q',t\ \ \p',q',t\ -m\ \p'',q'',u\)))" -by(subst converse_in_rel_pow_eq, simp) - - -lemma exec_star_simp[simp]: "(\i#p,q,s\ -*\ \p'',q'',u\) = - (p'' = i#p & q''=q & u=s | - (\p' q' t. \i#p,q,s\ -1\ \p',q',t\ \ \p',q',t\ -*\ \p'',q'',u\))" -apply(simp add: rtrancl_is_UN_rel_pow del:exec_simp) -apply(blast) -done - -declare nth_append[simp] - -lemma rev_revD: "rev xs = rev ys \ xs = ys" -by simp - -lemma [simp]: "(rev xs @ rev ys = rev zs) = (ys @ xs = zs)" -apply(rule iffI) - apply(rule rev_revD, simp) -apply fastsimp -done - -lemma direction1: - "\q,p,s\ -1\ \q',p',t\ \ - rev p' @ q' = rev p @ q \ rev p @ q \ \size p,s\ -1\ \size p',t\" -apply(induct set: stepa1) - apply(simp add:exec01.SET) - apply(fastsimp intro:exec01.JMPFT) - apply simp - apply(rule exec01.JMPFF) - apply simp - apply fastsimp - apply simp - apply simp - apply simp -apply(fastsimp simp add:exec01.JMPB) -done - -(* -lemma rev_take: "\i. rev (take i xs) = drop (length xs - i) (rev xs)" -apply(induct xs) - apply simp_all -apply(case_tac i) -apply simp_all -done - -lemma rev_drop: "\i. rev (drop i xs) = take (length xs - i) (rev xs)" -apply(induct xs) - apply simp_all -apply(case_tac i) -apply simp_all -done -*) - -lemma direction2: - "rpq \ \sp,s\ -1\ \sp',t\ \ - rpq = rev p @ q & sp = size p & sp' = size p' \ - rev p' @ q' = rev p @ q \ \q,p,s\ -1\ \q',p',t\" -apply(induct arbitrary: p q p' q' set: exec01) - apply(clarsimp simp add: neq_Nil_conv append_eq_conv_conj) - apply(drule sym) - apply simp - apply(rule rev_revD) - apply simp - apply(clarsimp simp add: neq_Nil_conv append_eq_conv_conj) - apply(drule sym) - apply simp - apply(rule rev_revD) - apply simp - apply(simp (no_asm_use) add: neq_Nil_conv append_eq_conv_conj, clarify)+ - apply(drule sym) - apply simp - apply(rule rev_revD) - apply simp -apply(clarsimp simp add: neq_Nil_conv append_eq_conv_conj) -apply(drule sym) -apply(simp add:rev_take) -apply(rule rev_revD) -apply(simp add:rev_drop) -done - - -theorem M_eqiv: -"(\q,p,s\ -1\ \q',p',t\) = - (rev p' @ q' = rev p @ q \ rev p @ q \ \size p,s\ -1\ \size p',t\)" - by (blast dest: direction1 direction2) - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Natural.thy --- a/src/HOL/IMP/Natural.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,354 +0,0 @@ -(* Title: HOL/IMP/Natural.thy - 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\ \\<^sub>c s'"} for \emph{Statement @{text c}, started - in state @{text s}, terminates in state @{text s'}}. Formally, - @{text "\c,s\ \\<^sub>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 ("_/[_ \ /_]" [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" ("\_,_\/ \\<^sub>c _" [0,0,60] 60) -where - Skip: "\\,s\ \\<^sub>c s" -| Assign: "\x :== a,s\ \\<^sub>c s[x\a s]" - -| Semi: "\c0,s\ \\<^sub>c s'' \ \c1,s''\ \\<^sub>c s' \ \c0; c1, s\ \\<^sub>c s'" - -| IfTrue: "b s \ \c0,s\ \\<^sub>c s' \ \\ b \ c0 \ c1, s\ \\<^sub>c s'" -| IfFalse: "\b s \ \c1,s\ \\<^sub>c s' \ \\ b \ c0 \ c1, s\ \\<^sub>c s'" - -| WhileFalse: "\b s \ \\ b \ c,s\ \\<^sub>c s" -| WhileTrue: "b s \ \c,s\ \\<^sub>c s'' \ \\ b \ c, s''\ \\<^sub>c s' - \ \\ b \ c, s\ \\<^sub>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!]: "\\,s\ \\<^sub>c s'" -inductive_cases semiE [elim!]: "\c0; c1, s\ \\<^sub>c s'" -inductive_cases assignE [elim!]: "\x :== a,s\ \\<^sub>c s'" -inductive_cases ifE [elim!]: "\\ b \ c0 \ c1, s\ \\<^sub>c s'" -inductive_cases whileE [elim]: "\\ b \ c,s\ \\<^sub>c s'" - -text {* The next proofs are all trivial by rule inversion. -*} - -inductive_simps - skip: "\\,s\ \\<^sub>c s'" - and assign: "\x :== a,s\ \\<^sub>c s'" - and semi: "\c0; c1, s\ \\<^sub>c s'" - -lemma ifTrue: - "b s \ \\ b \ c0 \ c1, s\ \\<^sub>c s' = \c0,s\ \\<^sub>c s'" - by auto - -lemma ifFalse: - "\b s \ \\ b \ c0 \ c1, s\ \\<^sub>c s' = \c1,s\ \\<^sub>c s'" - by auto - -lemma whileFalse: - "\ b s \ \\ b \ c,s\ \\<^sub>c s' = (s' = s)" - by auto - -lemma whileTrue: - "b s \ - \\ b \ c, s\ \\<^sub>c s' = - (\s''. \c,s\ \\<^sub>c s'' \ \\ b \ c, s''\ \\<^sub>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\ \\<^sub>c s' = \c', s\ \\<^sub>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\ \\<^sub>c s' = \c', s\ \\<^sub>c s') \ c \ c'" - by (unfold equiv_c_def) blast - -lemma equivD1: - "c \ c' \ \c, s\ \\<^sub>c s' \ \c', s\ \\<^sub>c s'" - by (unfold equiv_c_def) blast - -lemma equivD2: - "c \ c' \ \c', s\ \\<^sub>c s' \ \c, s\ \\<^sub>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: - "(\ b \ c) \ (\ b \ c; \ b \ c \ \)" (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\ \\<^sub>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\ \\<^sub>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\ \\<^sub>c s'"} *} - { assume b: "b s" - with w obtain s'' where - "\c, s\ \\<^sub>c s''" and "\?w, s''\ \\<^sub>c s'" by (cases set: evalc) auto - -- "now we can build a derivation tree for the @{text \}" - -- "first, the body of the True-branch:" - hence "\c; ?w, s\ \\<^sub>c s'" by (rule Semi) - -- "then the whole @{text \}" - with b have "\?if, s\ \\<^sub>c s'" by (rule IfTrue) - } - ultimately - -- "both cases together give us what we want:" - have "\?if, s\ \\<^sub>c s'" by blast - } - moreover - -- "now the other direction:" - { fix s s' assume "if": "\?if, s\ \\<^sub>c s'" - -- "again, if @{text b} is @{text False} in state @{text s}, then the False-branch" - -- "of the @{text \} is executed, and both statements do nothing:" - hence "\b s \ s = s'" by blast - hence "\b s \ \?w, s\ \\<^sub>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\ \\<^sub>c s'" by (cases set: evalc) auto - -- "and for this, only the Semi-rule is applicable:" - then obtain s'' where - "\c, s\ \\<^sub>c s''" and "\?w, s''\ \\<^sub>c s'" by (cases set: evalc) auto - -- "with this information, we can build a derivation tree for the @{text \}" - with b - have "\?w, s\ \\<^sub>c s'" by (rule WhileTrue) - } - ultimately - -- "both cases together again give us what we want:" - have "\?w, s\ \\<^sub>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 - "(\ b \ c) \ (\ b \ c; \ b \ c \ \)" -by blast - -lemma triv_if: - "(\ b \ c \ c) \ c" -by blast - -lemma commute_if: - "(\ b1 \ (\ b2 \ c11 \ c12) \ c2) - \ - (\ b2 \ (\ b1 \ c11 \ c2) \ (\ b1 \ c12 \ c2))" -by blast - -lemma while_equiv: - "\c0, s\ \\<^sub>c u \ c \ c' \ (c0 = \ b \ c) \ \\ b \ c', s\ \\<^sub>c u" -by (induct rule: evalc.induct) (auto simp add: equiv_c_def) - -lemma equiv_while: - "c \ c' \ (\ b \ c) \ (\ b \ 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' \ (\ b \ c1 \ c2) \ (\ b \ c1' \ c2')" -by (force simp add: equiv_c_def) - -lemma while_never: "\c, s\ \\<^sub>c u \ c \ \ (\s. True) \ c1" -apply (induct rule: evalc.induct) -apply auto -done - -lemma equiv_while_True: - "(\ (\s. True) \ c1) \ (\ (\s. True) \ c2)" -by (blast dest: while_never) - - -subsection "Execution is deterministic" - -text {* -This proof is automatic. -*} -theorem "\c,s\ \\<^sub>c t \ \c,s\ \\<^sub>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\ \\<^sub>c t" and "\c,s\ \\<^sub>c u" - shows "u = t" - using assms -proof (induct arbitrary: u set: evalc) - fix s u assume "\\,s\ \\<^sub>c u" - thus "u = s" by blast -next - fix a s x u assume "\x :== a,s\ \\<^sub>c u" - thus "u = s[x \ a s]" by blast -next - fix c0 c1 s s1 s2 u - assume IH0: "\u. \c0,s\ \\<^sub>c u \ u = s2" - assume IH1: "\u. \c1,s2\ \\<^sub>c u \ u = s1" - - assume "\c0;c1, s\ \\<^sub>c u" - then obtain s' where - c0: "\c0,s\ \\<^sub>c s'" and - c1: "\c1,s'\ \\<^sub>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\ \\<^sub>c u \ u = s1" - - assume "b s" and "\\ b \ c0 \ c1,s\ \\<^sub>c u" - hence "\c0, s\ \\<^sub>c u" by blast - with IH show "u = s1" by blast -next - fix b c0 c1 s s1 u - assume IH: "\u. \c1,s\ \\<^sub>c u \ u = s1" - - assume "\b s" and "\\ b \ c0 \ c1,s\ \\<^sub>c u" - hence "\c1, s\ \\<^sub>c u" by blast - with IH show "u = s1" by blast -next - fix b c s u - assume "\b s" and "\\ b \ c,s\ \\<^sub>c u" - thus "u = s" by blast -next - fix b c s s1 s2 u - assume "IH\<^sub>c": "\u. \c,s\ \\<^sub>c u \ u = s2" - assume "IH\<^sub>w": "\u. \\ b \ c,s2\ \\<^sub>c u \ u = s1" - - assume "b s" and "\\ b \ c,s\ \\<^sub>c u" - then obtain s' where - c: "\c,s\ \\<^sub>c s'" and - w: "\\ b \ c,s'\ \\<^sub>c u" - by auto - - from c "IH\<^sub>c" have "s' = s2" by blast - with w "IH\<^sub>w" 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\ \\<^sub>c t" and "\c,s\ \\<^sub>c u" - shows "u = t" - using assms -proof (induct arbitrary: u) - -- "the simple @{text \} case for demonstration:" - fix s u assume "\\,s\ \\<^sub>c u" - thus "u = s" by blast -next - -- "and the only really interesting case, @{text \}:" - fix b c s s1 s2 u - assume "IH\<^sub>c": "\u. \c,s\ \\<^sub>c u \ u = s2" - assume "IH\<^sub>w": "\u. \\ b \ c,s2\ \\<^sub>c u \ u = s1" - - assume "b s" and "\\ b \ c,s\ \\<^sub>c u" - then obtain s' where - c: "\c,s\ \\<^sub>c s'" and - w: "\\ b \ c,s'\ \\<^sub>c u" - by auto - - from c "IH\<^sub>c" have "s' = s2" by blast - with w "IH\<^sub>w" show "u = s1" by blast -qed blast+ -- "prove the rest automatically" - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/Transition.thy --- a/src/HOL/IMP/Transition.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,644 +0,0 @@ -(* Title: HOL/IMP/Transition.thy - Author: Tobias Nipkow & Robert Sandner, TUM - Isar Version: Gerwin Klein, 2001 - Copyright 1996 TUM -*) - -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 - " == (Some c, s)" -abbreviation - angle2 :: "state \ com option \ state" ("<_>") where - " == (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: "\\, s\ \\<^sub>1 \s\" -| Assign: "\x :== a, s\ \\<^sub>1 \s[x \ 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 \ \\ b \ c1 \ c2,s\ \\<^sub>1 \c1,s\" -| IfFalse: "\b s \ \\ b \ c1 \ c2,s\ \\<^sub>1 \c2,s\" - -| While: "\\ b \ c,s\ \\<^sub>1 \\ b \ c; \ b \ c \ \,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: "\\, 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 \ a s]\)" - by (induct y, rule, cases set: evalc1, auto) - -lemma Cond_1: - "\\ b \ c1 \ 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: - "\\ b \ c, s\ \\<^sub>1 y = (y = \\ b \ c; \ b \ c \ \, 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 \ \\ \s. s x \ 1 \ (x:== \s. s x+1), s\ \\<^sub>1\<^sup>* \s[x \ 1]\" - (is "_ \ \?w, _\ \\<^sub>1\<^sup>* _") -proof - - let ?c = "x:== \s. s x+1" - let ?if = "\ \s. s x \ 1 \ ?c; ?w \ \" - 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 \ 1]\" by (rule Semi1) simp - also have "\?w, s[x \ 1]\ \\<^sub>1 \?if, s[x \ 1]\" .. - also have "\?if, s[x \ 1]\ \\<^sub>1 \\, s[x \ 1]\" by (simp add: update_def) - also have "\\, s[x \ 1]\ \\<^sub>1 \s[x \ 1]\" .. - finally show ?thesis .. -qed - -lemma - "s x = 2 \ \\ \s. s x \ 1 \ (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 = "\ \s. s x \ 1 \ ?c; ?w \ \" - 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 \ 3]\" by (rule Semi1) simp - also have "\?w, s[x \ 3]\ \\<^sub>1 \?if, s[x \ 3]\" .. - also have "\?if, s[x \ 3]\ \\<^sub>1 \?c; ?w, s[x \ 3]\" by simp - also have "\?c; ?w, s[x \ 3]\ \\<^sub>1 \?w, s[x \ 4]\" by (rule Semi1) simp - also have "\?w, s[x \ 4]\ \\<^sub>1 \?if, s[x \ 4]\" .. - also have "\?if, s[x \ 4]\ \\<^sub>1 \?c; ?w, s[x \ 4]\" by simp - also have "\?c; ?w, s[x \ 4]\ \\<^sub>1 \?w, s[x \ 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. \\ b \ c1 \ 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 - -(*<*) -(* FIXME: relpow.simps don't work *) -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: "\\, 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 assms -proof induct - fix s show "\\,s\ \\<^sub>1\<^sup>* \s\" by auto -next - fix x a s show "\x :== a ,s\ \\<^sub>1\<^sup>* \s[x\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 "\\ b \ c0 \ c1,s\ \\<^sub>1 \c0,s\" by simp - also assume "\c0,s\ \\<^sub>1\<^sup>* \s'\" - finally show "\\ b \ c0 \ c1,s\ \\<^sub>1\<^sup>* \s'\" . -next - fix s::state and b c0 c1 s' - assume "\b s" - hence "\\ b \ c0 \ c1,s\ \\<^sub>1 \c1,s\" by simp - also assume "\c1,s\ \\<^sub>1\<^sup>* \s'\" - finally show "\\ b \ c0 \ c1,s\ \\<^sub>1\<^sup>* \s'\" . -next - fix b c and s::state - assume b: "\b s" - let ?if = "\ b \ c; \ b \ c \ \" - have "\\ b \ c,s\ \\<^sub>1 \?if, s\" by blast - also have "\?if,s\ \\<^sub>1 \\, s\" by (simp add: b) - also have "\\, s\ \\<^sub>1 \s\" by blast - finally show "\\ b \ c,s\ \\<^sub>1\<^sup>* \s\" .. -next - fix b c s s'' s' - let ?w = "\ b \ c" - let ?if = "\ b \ c; ?w \ \" - 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 "\\ b \ 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 = \ b \ c1 \ 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 = \ b \ c'" - with c n - have "\\ b \ c'; \ b \ c' \ \,s\ -m\\<^sub>1 \s'\" - (is "\?if,_\ -m\\<^sub>1 _") by auto - with n IH - have "\\ b \ c'; \ b \ c' \ \,s\ \\<^sub>c s'" by blast - moreover note unfold_while [of b c'] - -- {* @{thm unfold_while [of b c']} *} - ultimately - have "\\ b \ 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 \ \\ b \ c,s\ \\<^sub>1\<^sup>* \s\" (is "_ \ \?w, s\ \\<^sub>1\<^sup>* \s\") -proof - - assume "\b s" - have "\?w, s\ \\<^sub>1 \\ b \ c;?w \ \, s\" .. - also from `\b s` have "\\ b \ c;?w \ \, s\ \\<^sub>1 \\, s\" .. - also have "\\, s\ \\<^sub>1 \s\" .. - finally show "\?w, s\ \\<^sub>1\<^sup>* \s\" .. -qed - -lemma whileTrue1 [intro]: - "b s \ \\ b \ c,s\ \\<^sub>1\<^sup>* \c;\ b \ c, s\" - (is "_ \ \?w, s\ \\<^sub>1\<^sup>* \c;?w,s\") -proof - - assume "b s" - have "\?w, s\ \\<^sub>1 \\ b \ c;?w \ \, s\" .. - also from `b s` have "\\ b \ c;?w \ \, s\ \\<^sub>1 \c;?w, s\" .. - finally show "\?w, s\ \\<^sub>1\<^sup>* \c;?w,s\" .. -qed - -inductive_cases evalc1_SEs: - "\\,s\ \\<^sub>1 (co, s')" - "\x:==a,s\ \\<^sub>1 (co, s')" - "\c1;c2, s\ \\<^sub>1 (co, s')" - "\\ b \ c1 \ c2, s\ \\<^sub>1 (co, s')" - "\\ b \ c, s\ \\<^sub>1 (co, s')" - -inductive_cases evalc1_E: "\\ b \ 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 "\"} 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 assms - 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-> t) - -This is a quick one, dealing with the cases skip, assignment -and while_false. - -Lemma 2 -((c,s) -*-> (c',s')) /\ -c'-> t - => - -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'-> t - => - -c-> t - -This captures the essence of the proof, as it shows that -behaves as the continuation of 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 \ 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 \ 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 diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IMP/VC.thy --- a/src/HOL/IMP/VC.thy Wed Jun 01 21:35:34 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,137 +0,0 @@ -(* Title: HOL/IMP/VC.thy - Author: Tobias Nipkow - -acom: annotated commands -vc: verification-conditions -awp: weakest (liberal) precondition -*) - -header "Verification Conditions" - -theory VC imports Hoare_Op begin - -datatype acom = Askip - | Aass loc aexp - | Asemi acom acom - | Aif bexp acom acom - | Awhile bexp assn acom - -primrec awp :: "acom => assn => assn" -where - "awp Askip Q = Q" -| "awp (Aass x a) Q = (\s. Q(s[x\a s]))" -| "awp (Asemi c d) Q = awp c (awp d Q)" -| "awp (Aif b c d) Q = (\s. (b s-->awp c Q s) & (~b s-->awp d Q s))" -| "awp (Awhile b I c) Q = I" - -primrec vc :: "acom => assn => assn" -where - "vc Askip Q = (\s. True)" -| "vc (Aass x a) Q = (\s. True)" -| "vc (Asemi c d) Q = (\s. vc c (awp d Q) s & vc d Q s)" -| "vc (Aif b c d) Q = (\s. vc c Q s & vc d Q s)" -| "vc (Awhile b I c) Q = (\s. (I s & ~b s --> Q s) & - (I s & b s --> awp c I s) & vc c I s)" - -primrec astrip :: "acom => com" -where - "astrip Askip = SKIP" -| "astrip (Aass x a) = (x:==a)" -| "astrip (Asemi c d) = (astrip c;astrip d)" -| "astrip (Aif b c d) = (\ b \ astrip c \ astrip d)" -| "astrip (Awhile b I c) = (\ b \ astrip c)" - -(* simultaneous computation of vc and awp: *) -primrec vcawp :: "acom => assn => assn \ assn" -where - "vcawp Askip Q = (\s. True, Q)" -| "vcawp (Aass x a) Q = (\s. True, \s. Q(s[x\a s]))" -| "vcawp (Asemi c d) Q = (let (vcd,wpd) = vcawp d Q; - (vcc,wpc) = vcawp c wpd - in (\s. vcc s & vcd s, wpc))" -| "vcawp (Aif b c d) Q = (let (vcd,wpd) = vcawp d Q; - (vcc,wpc) = vcawp c Q - in (\s. vcc s & vcd s, - \s.(b s --> wpc s) & (~b s --> wpd s)))" -| "vcawp (Awhile b I c) Q = (let (vcc,wpc) = vcawp c I - in (\s. (I s & ~b s --> Q s) & - (I s & b s --> wpc s) & vcc s, I))" - -(* -Soundness and completeness of vc -*) - -declare hoare.conseq [intro] - - -lemma vc_sound: "(ALL s. vc c Q s) \ |- {awp c Q} astrip c {Q}" -proof(induct c arbitrary: Q) - case (Awhile b I c) - show ?case - proof(simp, rule While') - from `\s. vc (Awhile b I c) Q s` - have vc: "ALL s. vc c I s" and IQ: "ALL s. I s \ \ b s \ Q s" and - awp: "ALL s. I s & b s --> awp c I s" by simp_all - from vc have "|- {awp c I} astrip c {I}" using Awhile.hyps by blast - with awp show "|- {\s. I s \ b s} astrip c {I}" - by(rule strengthen_pre) - show "\s. I s \ \ b s \ Q s" by(rule IQ) - qed -qed auto - - -lemma awp_mono: - "(!s. P s --> Q s) ==> awp c P s ==> awp c Q s" -proof (induct c arbitrary: P Q s) - case Asemi thus ?case by simp metis -qed simp_all - -lemma vc_mono: - "(!s. P s --> Q s) ==> vc c P s ==> vc c Q s" -proof(induct c arbitrary: P Q) - case Asemi thus ?case by simp (metis awp_mono) -qed simp_all - -lemma vc_complete: assumes der: "|- {P}c{Q}" - shows "(\ac. astrip ac = c & (\s. vc ac Q s) & (\s. P s --> awp ac Q s))" - (is "? ac. ?Eq P c Q ac") -using der -proof induct - case skip - show ?case (is "? ac. ?C ac") - proof show "?C Askip" by simp qed -next - case (ass P x a) - show ?case (is "? ac. ?C ac") - proof show "?C(Aass x a)" by simp qed -next - case (semi P c1 Q c2 R) - from semi.hyps obtain ac1 where ih1: "?Eq P c1 Q ac1" by fast - from semi.hyps obtain ac2 where ih2: "?Eq Q c2 R ac2" by fast - show ?case (is "? ac. ?C ac") - proof - show "?C(Asemi ac1 ac2)" - using ih1 ih2 by simp (fast elim!: awp_mono vc_mono) - qed -next - case (If P b c1 Q c2) - from If.hyps obtain ac1 where ih1: "?Eq (%s. P s & b s) c1 Q ac1" by fast - from If.hyps obtain ac2 where ih2: "?Eq (%s. P s & ~b s) c2 Q ac2" by fast - show ?case (is "? ac. ?C ac") - proof - show "?C(Aif b ac1 ac2)" - using ih1 ih2 by simp - qed -next - case (While P b c) - from While.hyps obtain ac where ih: "?Eq (%s. P s & b s) c P ac" by fast - show ?case (is "? ac. ?C ac") - proof show "?C(Awhile b P ac)" using ih by simp qed -next - case conseq thus ?case by(fast elim!: awp_mono vc_mono) -qed - -lemma vcawp_vc_awp: "vcawp c Q = (vc c Q, awp c Q)" - by (induct c arbitrary: Q) (simp_all add: Let_def) - -end diff -r 11fce8564415 -r 2a05c1f7c08c src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Jun 01 21:35:34 2011 +0200 +++ b/src/HOL/IsaMakefile Wed Jun 01 21:50:49 2011 +0200 @@ -525,9 +525,9 @@ HOL-IMP: HOL $(LOG)/HOL-IMP.gz -$(LOG)/HOL-IMP.gz: $(OUT)/HOL IMP/Com.thy IMP/Compiler0.thy \ - IMP/Compiler.thy IMP/Denotation.thy IMP/Expr.thy IMP/Hoare.thy \ - IMP/Natural.thy IMP/Examples.thy IMP/Transition.thy IMP/VC.thy \ +$(LOG)/HOL-IMP.gz: $(OUT)/HOL IMP/ASM.thy IMP/AExp.thy IMP/BExp.thy \ + IMP/Big_Step.thy IMP/Com.thy IMP/Compiler.thy IMP/Denotation.thy \ + IMP/Small_Step.thy IMP/Star.thy \ IMP/ROOT.ML IMP/document/root.tex IMP/document/root.bib @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL IMP