| author | wenzelm | 
| Tue, 28 Aug 2012 12:52:14 +0200 | |
| changeset 48960 | 17dbe95eaa2a | 
| parent 46822 | 95f1e700b712 | 
| child 58871 | c399ae4b836f | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/IMP/Equiv.thy | 
| 40945 | 2 | Author: Heiko Loetzbeyer and Robert Sandner, TU München | 
| 482 | 3 | *) | 
| 4 | ||
| 12610 | 5 | header {* Equivalence *}
 | 
| 6 | ||
| 16417 | 7 | theory Equiv imports Denotation Com begin | 
| 12606 | 8 | |
| 9 | lemma aexp_iff [rule_format]: | |
| 12610 | 10 | "[| a \<in> aexp; sigma: loc -> nat |] | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
40945diff
changeset | 11 | ==> \<forall>n. <a,sigma> -a-> n \<longleftrightarrow> A(a,sigma) = n" | 
| 12610 | 12 | apply (erule aexp.induct) | 
| 13 | apply (force intro!: evala.intros)+ | |
| 14 | done | |
| 12606 | 15 | |
| 16 | declare aexp_iff [THEN iffD1, simp] | |
| 17 | aexp_iff [THEN iffD2, intro!] | |
| 18 | ||
| 12610 | 19 | inductive_cases [elim!]: | 
| 20 | "<true,sigma> -b-> x" | |
| 21 | "<false,sigma> -b-> x" | |
| 22 | "<ROp(f,a0,a1),sigma> -b-> x" | |
| 23 | "<noti(b),sigma> -b-> x" | |
| 24 | "<b0 andi b1,sigma> -b-> x" | |
| 25 | "<b0 ori b1,sigma> -b-> x" | |
| 12606 | 26 | |
| 27 | ||
| 28 | lemma bexp_iff [rule_format]: | |
| 12610 | 29 | "[| b \<in> bexp; sigma: loc -> nat |] | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
40945diff
changeset | 30 | ==> \<forall>w. <b,sigma> -b-> w \<longleftrightarrow> B(b,sigma) = w" | 
| 12610 | 31 | apply (erule bexp.induct) | 
| 32 | apply (auto intro!: evalb.intros) | |
| 33 | done | |
| 12606 | 34 | |
| 35 | declare bexp_iff [THEN iffD1, simp] | |
| 36 | bexp_iff [THEN iffD2, intro!] | |
| 37 | ||
| 38 | lemma com1: "<c,sigma> -c-> sigma' ==> <sigma,sigma'> \<in> C(c)" | |
| 12610 | 39 | apply (erule evalc.induct) | 
| 13612 | 40 | apply (simp_all (no_asm_simp)) | 
| 12610 | 41 |      txt {* @{text assign} *}
 | 
| 42 | apply (simp add: update_type) | |
| 43 |     txt {* @{text comp} *}
 | |
| 44 | apply fast | |
| 45 |    txt {* @{text while} *}
 | |
| 46 | apply (erule Gamma_bnd_mono [THEN lfp_unfold, THEN ssubst, OF C_subset]) | |
| 47 | apply (simp add: Gamma_def) | |
| 48 |   txt {* recursive case of @{text while} *}
 | |
| 49 | apply (erule Gamma_bnd_mono [THEN lfp_unfold, THEN ssubst, OF C_subset]) | |
| 50 | apply (auto simp add: Gamma_def) | |
| 51 | done | |
| 12606 | 52 | |
| 53 | declare B_type [intro!] A_type [intro!] | |
| 54 | declare evalc.intros [intro] | |
| 55 | ||
| 56 | ||
| 57 | lemma com2 [rule_format]: "c \<in> com ==> \<forall>x \<in> C(c). <c,fst(x)> -c-> snd(x)" | |
| 12610 | 58 | apply (erule com.induct) | 
| 59 |       txt {* @{text skip} *}
 | |
| 60 | apply force | |
| 61 |      txt {* @{text assign} *}
 | |
| 62 | apply force | |
| 63 |     txt {* @{text comp} *}
 | |
| 64 | apply force | |
| 65 |    txt {* @{text while} *}
 | |
| 66 | apply safe | |
| 67 | apply simp_all | |
| 68 | apply (frule Gamma_bnd_mono [OF C_subset], erule Fixedpt.induct, assumption) | |
| 69 | apply (unfold Gamma_def) | |
| 70 | apply force | |
| 19796 | 71 |   txt {* @{text "if"} *}
 | 
| 12610 | 72 | apply auto | 
| 73 | done | |
| 12606 | 74 | |
| 75 | ||
| 12610 | 76 | subsection {* Main theorem *}
 | 
| 12606 | 77 | |
| 12610 | 78 | theorem com_equivalence: | 
| 79 |     "c \<in> com ==> C(c) = {io \<in> (loc->nat) \<times> (loc->nat). <c,fst(io)> -c-> snd(io)}"
 | |
| 80 | by (force intro: C_subset [THEN subsetD] elim: com2 dest: com1) | |
| 12606 | 81 | |
| 82 | end | |
| 83 |