| author | wenzelm | 
| Tue, 08 Nov 2011 12:03:51 +0100 | |
| changeset 45405 | 23e5af70af07 | 
| parent 39302 | d7728f65b353 | 
| child 46577 | e5438c5797ae | 
| permissions | -rw-r--r-- | 
| 7186 | 1 | (* Title: HOL/UNITY/Lift_prog.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1999 University of Cambridge | |
| 4 | ||
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 5 | lift_prog, etc: replication of components and arrays of processes. | 
| 7186 | 6 | *) | 
| 7 | ||
| 13798 | 8 | header{*Replication of Components*}
 | 
| 9 | ||
| 16417 | 10 | theory Lift_prog imports Rename begin | 
| 7186 | 11 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 12 | definition insert_map :: "[nat, 'b, nat=>'b] => (nat=>'b)" where | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8122diff
changeset | 13 | "insert_map i z f k == if k<i then f k | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8122diff
changeset | 14 | else if k=i then z | 
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
10834diff
changeset | 15 | else f(k - 1)" | 
| 7186 | 16 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 17 | definition delete_map :: "[nat, nat=>'b] => (nat=>'b)" where | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8122diff
changeset | 18 | "delete_map i g k == if k<i then g k else g (Suc k)" | 
| 7186 | 19 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 20 | definition lift_map :: "[nat, 'b * ((nat=>'b) * 'c)] => (nat=>'b) * 'c" where | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8122diff
changeset | 21 | "lift_map i == %(s,(f,uu)). (insert_map i s f, uu)" | 
| 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8122diff
changeset | 22 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 23 | definition drop_map :: "[nat, (nat=>'b) * 'c] => 'b * ((nat=>'b) * 'c)" where | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8122diff
changeset | 24 | "drop_map i == %(g, uu). (g i, (delete_map i g, uu))" | 
| 7186 | 25 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 26 | definition lift_set :: "[nat, ('b * ((nat=>'b) * 'c)) set] => ((nat=>'b) * 'c) set" where
 | 
| 10834 | 27 | "lift_set i A == lift_map i ` A" | 
| 7186 | 28 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 29 | definition lift :: "[nat, ('b * ((nat=>'b) * 'c)) program] => ((nat=>'b) * 'c) program" where
 | 
| 8251 
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
 paulson parents: 
8122diff
changeset | 30 | "lift i == rename (lift_map i)" | 
| 7186 | 31 | |
| 32 | (*simplifies the expression of specifications*) | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 33 | definition sub :: "['a, 'a=>'b] => 'b" where | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 34 | "sub == %i f. f i" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 35 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 36 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 37 | declare insert_map_def [simp] delete_map_def [simp] | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 38 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 39 | lemma insert_map_inverse: "delete_map i (insert_map i x f) = f" | 
| 13798 | 40 | by (rule ext, simp) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 41 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 42 | lemma insert_map_delete_map_eq: "(insert_map i x (delete_map i g)) = g(i:=x)" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 43 | apply (rule ext) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 44 | apply (auto split add: nat_diff_split) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 45 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 46 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 47 | subsection{*Injectiveness proof*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 48 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 49 | lemma insert_map_inject1: "(insert_map i x f) = (insert_map i y g) ==> x=y" | 
| 13798 | 50 | by (drule_tac x = i in fun_cong, simp) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 51 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 52 | lemma insert_map_inject2: "(insert_map i x f) = (insert_map i y g) ==> f=g" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 53 | apply (drule_tac f = "delete_map i" in arg_cong) | 
| 13798 | 54 | apply (simp add: insert_map_inverse) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 55 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 56 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 57 | lemma insert_map_inject': | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 58 | "(insert_map i x f) = (insert_map i y g) ==> x=y & f=g" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 59 | by (blast dest: insert_map_inject1 insert_map_inject2) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 60 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 61 | lemmas insert_map_inject = insert_map_inject' [THEN conjE, elim!] | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 62 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 63 | (*The general case: we don't assume i=i'*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 64 | lemma lift_map_eq_iff [iff]: | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 65 | "(lift_map i (s,(f,uu)) = lift_map i' (s',(f',uu'))) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 66 | = (uu = uu' & insert_map i s f = insert_map i' s' f')" | 
| 13798 | 67 | by (unfold lift_map_def, auto) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 68 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 69 | (*The !!s allows the automatic splitting of the bound variable*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 70 | lemma drop_map_lift_map_eq [simp]: "!!s. drop_map i (lift_map i s) = s" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 71 | apply (unfold lift_map_def drop_map_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 72 | apply (force intro: insert_map_inverse) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 73 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 74 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 75 | lemma inj_lift_map: "inj (lift_map i)" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 76 | apply (unfold lift_map_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 77 | apply (rule inj_onI, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 78 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 79 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 80 | subsection{*Surjectiveness proof*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 81 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 82 | lemma lift_map_drop_map_eq [simp]: "!!s. lift_map i (drop_map i s) = s" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 83 | apply (unfold lift_map_def drop_map_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 84 | apply (force simp add: insert_map_delete_map_eq) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 85 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 86 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 87 | lemma drop_map_inject [dest!]: "(drop_map i s) = (drop_map i s') ==> s=s'" | 
| 13798 | 88 | by (drule_tac f = "lift_map i" in arg_cong, simp) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 89 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 90 | lemma surj_lift_map: "surj (lift_map i)" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 91 | apply (rule surjI) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 92 | apply (rule lift_map_drop_map_eq) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 93 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 94 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 95 | lemma bij_lift_map [iff]: "bij (lift_map i)" | 
| 13798 | 96 | by (simp add: bij_def inj_lift_map surj_lift_map) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 97 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 98 | lemma inv_lift_map_eq [simp]: "inv (lift_map i) = drop_map i" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 99 | by (rule inv_equality, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 100 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 101 | lemma inv_drop_map_eq [simp]: "inv (drop_map i) = lift_map i" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 102 | by (rule inv_equality, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 103 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 104 | lemma bij_drop_map [iff]: "bij (drop_map i)" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 105 | by (simp del: inv_lift_map_eq add: inv_lift_map_eq [symmetric] bij_imp_bij_inv) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 106 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 107 | (*sub's main property!*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 108 | lemma sub_apply [simp]: "sub i f = f i" | 
| 13798 | 109 | by (simp add: sub_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 110 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 111 | lemma all_total_lift: "all_total F ==> all_total (lift i F)" | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 112 | by (simp add: lift_def rename_def Extend.all_total_extend) | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 113 | |
| 13836 | 114 | lemma insert_map_upd_same: "(insert_map i t f)(i := s) = insert_map i s f" | 
| 115 | by (rule ext, auto) | |
| 116 | ||
| 117 | lemma insert_map_upd: | |
| 118 | "(insert_map j t f)(i := s) = | |
| 119 | (if i=j then insert_map i s f | |
| 120 | else if i<j then insert_map j t (f(i:=s)) | |
| 121 | else insert_map j t (f(i - Suc 0 := s)))" | |
| 122 | apply (rule ext) | |
| 123 | apply (simp split add: nat_diff_split) | |
| 124 |  txt{*This simplification is VERY slow*}
 | |
| 125 | done | |
| 126 | ||
| 127 | lemma insert_map_eq_diff: | |
| 128 | "[| insert_map i s f = insert_map j t g; i\<noteq>j |] | |
| 129 | ==> \<exists>g'. insert_map i s' f = insert_map j t g'" | |
| 130 | apply (subst insert_map_upd_same [symmetric]) | |
| 131 | apply (erule ssubst) | |
| 132 | apply (simp only: insert_map_upd if_False split: split_if, blast) | |
| 133 | done | |
| 134 | ||
| 135 | lemma lift_map_eq_diff: | |
| 136 | "[| lift_map i (s,(f,uu)) = lift_map j (t,(g,vv)); i\<noteq>j |] | |
| 137 | ==> \<exists>g'. lift_map i (s',(f,uu)) = lift_map j (t,(g',vv))" | |
| 138 | apply (unfold lift_map_def, auto) | |
| 139 | apply (blast dest: insert_map_eq_diff) | |
| 140 | done | |
| 141 | ||
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 142 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 143 | subsection{*The Operator @{term lift_set}*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 144 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 145 | lemma lift_set_empty [simp]: "lift_set i {} = {}"
 | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 146 | by (unfold lift_set_def, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 147 | |
| 13805 | 148 | lemma lift_set_iff: "(lift_map i x \<in> lift_set i A) = (x \<in> A)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 149 | apply (unfold lift_set_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 150 | apply (rule inj_lift_map [THEN inj_image_mem_iff]) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 151 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 152 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 153 | (*Do we really need both this one and its predecessor?*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 154 | lemma lift_set_iff2 [iff]: | 
| 13805 | 155 | "((f,uu) \<in> lift_set i A) = ((f i, (delete_map i f, uu)) \<in> A)" | 
| 13798 | 156 | by (simp add: lift_set_def mem_rename_set_iff drop_map_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 157 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 158 | |
| 13805 | 159 | lemma lift_set_mono: "A \<subseteq> B ==> lift_set i A \<subseteq> lift_set i B" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 160 | apply (unfold lift_set_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 161 | apply (erule image_mono) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 162 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 163 | |
| 13805 | 164 | lemma lift_set_Un_distrib: "lift_set i (A \<union> B) = lift_set i A \<union> lift_set i B" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 165 | by (simp add: lift_set_def image_Un) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 166 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 167 | lemma lift_set_Diff_distrib: "lift_set i (A-B) = lift_set i A - lift_set i B" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 168 | apply (unfold lift_set_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 169 | apply (rule inj_lift_map [THEN image_set_diff]) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 170 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 171 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 172 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 173 | subsection{*The Lattice Operations*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 174 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 175 | lemma bij_lift [iff]: "bij (lift i)" | 
| 13798 | 176 | by (simp add: lift_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 177 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 178 | lemma lift_SKIP [simp]: "lift i SKIP = SKIP" | 
| 13798 | 179 | by (simp add: lift_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 180 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 181 | lemma lift_Join [simp]: "lift i (F Join G) = lift i F Join lift i G" | 
| 13798 | 182 | by (simp add: lift_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 183 | |
| 13805 | 184 | lemma lift_JN [simp]: "lift j (JOIN I F) = (\<Squnion>i \<in> I. lift j (F i))" | 
| 13798 | 185 | by (simp add: lift_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 186 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 187 | subsection{*Safety: constrains, stable, invariant*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 188 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 189 | lemma lift_constrains: | 
| 13805 | 190 | "(lift i F \<in> (lift_set i A) co (lift_set i B)) = (F \<in> A co B)" | 
| 13798 | 191 | by (simp add: lift_def lift_set_def rename_constrains) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 192 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 193 | lemma lift_stable: | 
| 13805 | 194 | "(lift i F \<in> stable (lift_set i A)) = (F \<in> stable A)" | 
| 13798 | 195 | by (simp add: lift_def lift_set_def rename_stable) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 196 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 197 | lemma lift_invariant: | 
| 13805 | 198 | "(lift i F \<in> invariant (lift_set i A)) = (F \<in> invariant A)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 199 | by (simp add: lift_def lift_set_def rename_invariant) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 200 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 201 | lemma lift_Constrains: | 
| 13805 | 202 | "(lift i F \<in> (lift_set i A) Co (lift_set i B)) = (F \<in> A Co B)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 203 | by (simp add: lift_def lift_set_def rename_Constrains) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 204 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 205 | lemma lift_Stable: | 
| 13805 | 206 | "(lift i F \<in> Stable (lift_set i A)) = (F \<in> Stable A)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 207 | by (simp add: lift_def lift_set_def rename_Stable) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 208 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 209 | lemma lift_Always: | 
| 13805 | 210 | "(lift i F \<in> Always (lift_set i A)) = (F \<in> Always A)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 211 | by (simp add: lift_def lift_set_def rename_Always) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 212 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 213 | subsection{*Progress: transient, ensures*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 214 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 215 | lemma lift_transient: | 
| 13805 | 216 | "(lift i F \<in> transient (lift_set i A)) = (F \<in> transient A)" | 
| 13798 | 217 | by (simp add: lift_def lift_set_def rename_transient) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 218 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 219 | lemma lift_ensures: | 
| 13805 | 220 | "(lift i F \<in> (lift_set i A) ensures (lift_set i B)) = | 
| 221 | (F \<in> A ensures B)" | |
| 13798 | 222 | by (simp add: lift_def lift_set_def rename_ensures) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 223 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 224 | lemma lift_leadsTo: | 
| 13805 | 225 | "(lift i F \<in> (lift_set i A) leadsTo (lift_set i B)) = | 
| 226 | (F \<in> A leadsTo B)" | |
| 13798 | 227 | by (simp add: lift_def lift_set_def rename_leadsTo) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 228 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 229 | lemma lift_LeadsTo: | 
| 13805 | 230 | "(lift i F \<in> (lift_set i A) LeadsTo (lift_set i B)) = | 
| 231 | (F \<in> A LeadsTo B)" | |
| 13798 | 232 | by (simp add: lift_def lift_set_def rename_LeadsTo) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 233 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 234 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 235 | (** guarantees **) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 236 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 237 | lemma lift_lift_guarantees_eq: | 
| 13805 | 238 | "(lift i F \<in> (lift i ` X) guarantees (lift i ` Y)) = | 
| 239 | (F \<in> X guarantees Y)" | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 240 | apply (unfold lift_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 241 | apply (subst bij_lift_map [THEN rename_rename_guarantees_eq, symmetric]) | 
| 13798 | 242 | apply (simp add: o_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 243 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 244 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 245 | lemma lift_guarantees_eq_lift_inv: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 246 | "(lift i F \<in> X guarantees Y) = | 
| 13805 | 247 | (F \<in> (rename (drop_map i) ` X) guarantees (rename (drop_map i) ` Y))" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 248 | by (simp add: bij_lift_map [THEN rename_guarantees_eq_rename_inv] lift_def) | 
| 7186 | 249 | |
| 250 | ||
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 251 | (*To preserve snd means that the second component is there just to allow | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 252 | guarantees properties to be stated. Converse fails, for lift i F can | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 253 | change function components other than i*) | 
| 13805 | 254 | lemma lift_preserves_snd_I: "F \<in> preserves snd ==> lift i F \<in> preserves snd" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 255 | apply (drule_tac w1=snd in subset_preserves_o [THEN subsetD]) | 
| 13798 | 256 | apply (simp add: lift_def rename_preserves) | 
| 14101 | 257 | apply (simp add: lift_map_def o_def split_def del: split_comp_eq) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 258 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 259 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 260 | lemma delete_map_eqE': | 
| 13805 | 261 | "(delete_map i g) = (delete_map i g') ==> \<exists>x. g = g'(i:=x)" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 262 | apply (drule_tac f = "insert_map i (g i) " in arg_cong) | 
| 13798 | 263 | apply (simp add: insert_map_delete_map_eq) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 264 | apply (erule exI) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 265 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 266 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 267 | lemmas delete_map_eqE = delete_map_eqE' [THEN exE, elim!] | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 268 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 269 | lemma delete_map_neq_apply: | 
| 13805 | 270 | "[| delete_map j g = delete_map j g'; i\<noteq>j |] ==> g i = g' i" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 271 | by force | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 272 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 273 | (*A set of the form (A <*> UNIV) ignores the second (dummy) state component*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 274 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 275 | lemma vimage_o_fst_eq [simp]: "(f o fst) -` A = (f-`A) <*> UNIV" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 276 | by auto | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 277 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 278 | lemma vimage_sub_eq_lift_set [simp]: | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 279 | "(sub i -`A) <*> UNIV = lift_set i (A <*> UNIV)" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 280 | by auto | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 281 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 282 | lemma mem_lift_act_iff [iff]: | 
| 13805 | 283 | "((s,s') \<in> extend_act (%(x,u::unit). lift_map i x) act) = | 
| 284 | ((drop_map i s, drop_map i s') \<in> act)" | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 285 | apply (unfold extend_act_def, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 286 | apply (rule bexI, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 287 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 288 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 289 | lemma preserves_snd_lift_stable: | 
| 13805 | 290 | "[| F \<in> preserves snd; i\<noteq>j |] | 
| 291 | ==> lift j F \<in> stable (lift_set i (A <*> UNIV))" | |
| 13798 | 292 | apply (auto simp add: lift_def lift_set_def stable_def constrains_def | 
| 293 | rename_def extend_def mem_rename_set_iff) | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 294 | apply (auto dest!: preserves_imp_eq simp add: lift_map_def drop_map_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 295 | apply (drule_tac x = i in fun_cong, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 296 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 297 | |
| 13805 | 298 | (*If i\<noteq>j then lift j F does nothing to lift_set i, and the | 
| 299 | premise ensures A \<subseteq> B.*) | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 300 | lemma constrains_imp_lift_constrains: | 
| 13805 | 301 | "[| F i \<in> (A <*> UNIV) co (B <*> UNIV); | 
| 302 | F j \<in> preserves snd |] | |
| 303 | ==> lift j (F j) \<in> (lift_set i (A <*> UNIV)) co (lift_set i (B <*> UNIV))" | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 304 | apply (case_tac "i=j") | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 305 | apply (simp add: lift_def lift_set_def rename_constrains) | 
| 13798 | 306 | apply (erule preserves_snd_lift_stable[THEN stableD, THEN constrains_weaken_R], | 
| 307 | assumption) | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 308 | apply (erule constrains_imp_subset [THEN lift_set_mono]) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 309 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 310 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 311 | (*USELESS??*) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 312 | lemma lift_map_image_Times: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 313 | "lift_map i ` (A <*> UNIV) = | 
| 13805 | 314 |       (\<Union>s \<in> A. \<Union>f. {insert_map i s f}) <*> UNIV"
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 315 | apply (auto intro!: bexI image_eqI simp add: lift_map_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 316 | apply (rule split_conv [symmetric]) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 317 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 318 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 319 | lemma lift_preserves_eq: | 
| 13805 | 320 | "(lift i F \<in> preserves v) = (F \<in> preserves (v o lift_map i))" | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 321 | by (simp add: lift_def rename_preserves) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 322 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 323 | (*A useful rewrite. If o, sub have been rewritten out already then can also | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 324 | use it as rewrite_rule [sub_def, o_def] lift_preserves_sub*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 325 | lemma lift_preserves_sub: | 
| 13805 | 326 | "F \<in> preserves snd | 
| 327 | ==> lift i F \<in> preserves (v o sub j o fst) = | |
| 328 | (if i=j then F \<in> preserves (v o fst) else True)" | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 329 | apply (drule subset_preserves_o [THEN subsetD]) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 330 | apply (simp add: lift_preserves_eq o_def drop_map_lift_map_eq) | 
| 13798 | 331 | apply (auto cong del: if_weak_cong | 
| 14101 | 332 | simp add: lift_map_def eq_commute split_def o_def simp del:split_comp_eq) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 333 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 334 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 335 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 336 | subsection{*Lemmas to Handle Function Composition (o) More Consistently*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 337 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 338 | (*Lets us prove one version of a theorem and store others*) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 339 | lemma o_equiv_assoc: "f o g = h ==> f' o f o g = f' o h" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 340 | by (simp add: fun_eq_iff o_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 341 | |
| 13805 | 342 | lemma o_equiv_apply: "f o g = h ==> \<forall>x. f(g x) = h x" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 343 | by (simp add: fun_eq_iff o_def) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 344 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 345 | lemma fst_o_lift_map: "sub i o fst o lift_map i = fst" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 346 | apply (rule ext) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 347 | apply (auto simp add: o_def lift_map_def sub_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 348 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 349 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 350 | lemma snd_o_lift_map: "snd o lift_map i = snd o snd" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 351 | apply (rule ext) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 352 | apply (auto simp add: o_def lift_map_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 353 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 354 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 355 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 356 | subsection{*More lemmas about extend and project*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 357 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 358 | text{*They could be moved to theory Extend or Project*}
 | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 359 | |
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 360 | lemma extend_act_extend_act: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 361 | "extend_act h' (extend_act h act) = | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 362 | extend_act (%(x,(y,y')). h'(h(x,y),y')) act" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 363 | apply (auto elim!: rev_bexI simp add: extend_act_def, blast) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 364 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 365 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 366 | lemma project_act_project_act: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 367 | "project_act h (project_act h' act) = | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 368 | project_act (%(x,(y,y')). h'(h(x,y),y')) act" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 369 | by (auto elim!: rev_bexI simp add: project_act_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 370 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 371 | lemma project_act_extend_act: | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 372 | "project_act h (extend_act h' act) = | 
| 13805 | 373 |         {(x,x'). \<exists>s s' y y' z. (s,s') \<in> act &  
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 374 | h(x,y) = h'(s,z) & h(x',y') = h'(s',z)}" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 375 | by (simp add: extend_act_def project_act_def, blast) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 376 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 377 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 378 | subsection{*OK and "lift"*}
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 379 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 380 | lemma act_in_UNION_preserves_fst: | 
| 13805 | 381 |      "act \<subseteq> {(x,x'). fst x = fst x'} ==> act \<in> UNION (preserves fst) Acts"
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 382 | apply (rule_tac a = "mk_program (UNIV,{act},UNIV) " in UN_I)
 | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 383 | apply (auto simp add: preserves_def stable_def constrains_def) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 384 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 385 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 386 | lemma UNION_OK_lift_I: | 
| 13805 | 387 | "[| \<forall>i \<in> I. F i \<in> preserves snd; | 
| 388 | \<forall>i \<in> I. UNION (preserves fst) Acts \<subseteq> AllowedActs (F i) |] | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 389 | ==> OK I (%i. lift i (F i))" | 
| 13790 | 390 | apply (auto simp add: OK_def lift_def rename_def Extend.Acts_extend) | 
| 13798 | 391 | apply (simp add: Extend.AllowedActs_extend project_act_extend_act) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 392 | apply (rename_tac "act") | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 393 | apply (subgoal_tac | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 394 |        "{(x, x'). \<exists>s f u s' f' u'. 
 | 
| 13805 | 395 | ((s, f, u), s', f', u') \<in> act & | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 396 | lift_map j x = lift_map i (s, f, u) & | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 397 | lift_map j x' = lift_map i (s', f', u') } | 
| 13805 | 398 |                 \<subseteq> { (x,x') . fst x = fst x'}")
 | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 399 | apply (blast intro: act_in_UNION_preserves_fst, clarify) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 400 | apply (drule_tac x = j in fun_cong)+ | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 401 | apply (drule_tac x = i in bspec, assumption) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 402 | apply (frule preserves_imp_eq, auto) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 403 | done | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 404 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 405 | lemma OK_lift_I: | 
| 13805 | 406 | "[| \<forall>i \<in> I. F i \<in> preserves snd; | 
| 407 | \<forall>i \<in> I. preserves fst \<subseteq> Allowed (F i) |] | |
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 408 | ==> OK I (%i. lift i (F i))" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 409 | by (simp add: safety_prop_AllowedActs_iff_Allowed UNION_OK_lift_I) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 410 | |
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 411 | lemma Allowed_lift [simp]: "Allowed (lift i F) = lift i ` (Allowed F)" | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 412 | by (simp add: lift_def Allowed_rename) | 
| 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 413 | |
| 13790 | 414 | lemma lift_image_preserves: | 
| 415 | "lift i ` preserves v = preserves (v o drop_map i)" | |
| 13798 | 416 | by (simp add: rename_image_preserves lift_def inv_lift_map_eq) | 
| 13786 
ab8f39f48a6f
More conversion of UNITY to Isar new-style theories
 paulson parents: 
11701diff
changeset | 417 | |
| 7186 | 418 | end |