author | wenzelm |
Tue, 13 Sep 2005 22:19:40 +0200 | |
changeset 17356 | 09afdf37cdb3 |
parent 16417 | 9bc16273c2d4 |
child 24327 | a207114007c6 |
permissions | -rw-r--r-- |
11479 | 1 |
(* Title: ZF/UNITY/Comp.thy |
2 |
ID: $Id$ |
|
3 |
Author: Sidi O Ehmety, Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
From Chandy and Sanders, "Reasoning About Program Composition", |
|
7 |
Technical Report 2000-003, University of Florida, 2000. |
|
8 |
||
9 |
Revised by Sidi Ehmety on January 2001 |
|
10 |
||
11 |
Added: a strong form of the order relation over component and localize |
|
12 |
||
13 |
Theory ported from HOL. |
|
14 |
||
15 |
*) |
|
16 |
||
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
17 |
header{*Composition*} |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
18 |
|
16417 | 19 |
theory Comp imports Union Increasing begin |
11479 | 20 |
|
21 |
constdefs |
|
22 |
||
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
23 |
component :: "[i,i]=>o" (infixl "component" 65) |
11479 | 24 |
"F component H == (EX G. F Join G = H)" |
25 |
||
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
26 |
strict_component :: "[i,i]=>o" (infixl "strict'_component" 65) |
11479 | 27 |
"F strict_component H == F component H & F~=H" |
28 |
||
29 |
(* A stronger form of the component relation *) |
|
30 |
component_of :: "[i,i]=>o" (infixl "component'_of" 65) |
|
31 |
"F component_of H == (EX G. F ok G & F Join G = H)" |
|
32 |
||
33 |
strict_component_of :: "[i,i]=>o" (infixl "strict'_component'_of" 65) |
|
34 |
"F strict_component_of H == F component_of H & F~=H" |
|
35 |
||
12195 | 36 |
(* Preserves a state function f, in particular a variable *) |
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
37 |
preserves :: "(i=>i)=>i" |
11479 | 38 |
"preserves(f) == {F:program. ALL z. F: stable({s:state. f(s) = z})}" |
39 |
||
14046 | 40 |
fun_pair :: "[i=>i, i =>i] =>(i=>i)" |
41 |
"fun_pair(f, g) == %x. <f(x), g(x)>" |
|
42 |
||
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
43 |
localize :: "[i=>i, i] => i" |
11479 | 44 |
"localize(f,F) == mk_program(Init(F), Acts(F), |
14046 | 45 |
AllowedActs(F) Int (UN G:preserves(f). Acts(G)))" |
11479 | 46 |
|
47 |
||
14093
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
48 |
(*** component and strict_component relations ***) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
49 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
50 |
lemma componentI: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
51 |
"H component F | H component G ==> H component (F Join G)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
52 |
apply (unfold component_def, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
53 |
apply (rule_tac x = "Ga Join G" in exI) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
54 |
apply (rule_tac [2] x = "G Join F" in exI) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
55 |
apply (auto simp add: Join_ac) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
56 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
57 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
58 |
lemma component_eq_subset: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
59 |
"G \<in> program ==> (F component G) <-> |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
60 |
(Init(G) <= Init(F) & Acts(F) <= Acts(G) & AllowedActs(G) <= AllowedActs(F))" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
61 |
apply (unfold component_def, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
62 |
apply (rule exI) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
63 |
apply (rule program_equalityI, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
64 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
65 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
66 |
lemma component_SKIP [simp]: "F \<in> program ==> SKIP component F" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
67 |
apply (unfold component_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
68 |
apply (rule_tac x = F in exI) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
69 |
apply (force intro: Join_SKIP_left) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
70 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
71 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
72 |
lemma component_refl [simp]: "F \<in> program ==> F component F" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
73 |
apply (unfold component_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
74 |
apply (rule_tac x = F in exI) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
75 |
apply (force intro: Join_SKIP_right) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
76 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
77 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
78 |
lemma SKIP_minimal: "F component SKIP ==> programify(F) = SKIP" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
79 |
apply (rule program_equalityI) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
80 |
apply (simp_all add: component_eq_subset, blast) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
81 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
82 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
83 |
lemma component_Join1: "F component (F Join G)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
84 |
by (unfold component_def, blast) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
85 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
86 |
lemma component_Join2: "G component (F Join G)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
87 |
apply (unfold component_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
88 |
apply (simp (no_asm) add: Join_commute) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
89 |
apply blast |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
90 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
91 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
92 |
lemma Join_absorb1: "F component G ==> F Join G = G" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
93 |
by (auto simp add: component_def Join_left_absorb) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
94 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
95 |
lemma Join_absorb2: "G component F ==> F Join G = F" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
96 |
by (auto simp add: Join_ac component_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
97 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
98 |
lemma JN_component_iff: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
99 |
"H \<in> program==>(JOIN(I,F) component H) <-> (\<forall>i \<in> I. F(i) component H)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
100 |
apply (case_tac "I=0", force) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
101 |
apply (simp (no_asm_simp) add: component_eq_subset) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
102 |
apply auto |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
103 |
apply blast |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
104 |
apply (rename_tac "y") |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
105 |
apply (drule_tac c = y and A = "AllowedActs (H)" in subsetD) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
106 |
apply (blast elim!: not_emptyE)+ |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
107 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
108 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
109 |
lemma component_JN: "i \<in> I ==> F(i) component (\<Squnion>i \<in> I. (F(i)))" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
110 |
apply (unfold component_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
111 |
apply (blast intro: JN_absorb) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
112 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
113 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
114 |
lemma component_trans: "[| F component G; G component H |] ==> F component H" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
115 |
apply (unfold component_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
116 |
apply (blast intro: Join_assoc [symmetric]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
117 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
118 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
119 |
lemma component_antisym: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
120 |
"[| F \<in> program; G \<in> program |] ==>(F component G & G component F) --> F = G" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
121 |
apply (simp (no_asm_simp) add: component_eq_subset) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
122 |
apply clarify |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
123 |
apply (rule program_equalityI, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
124 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
125 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
126 |
lemma Join_component_iff: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
127 |
"H \<in> program ==> |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
128 |
((F Join G) component H) <-> (F component H & G component H)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
129 |
apply (simp (no_asm_simp) add: component_eq_subset) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
130 |
apply blast |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
131 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
132 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
133 |
lemma component_constrains: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
134 |
"[| F component G; G \<in> A co B; F \<in> program |] ==> F \<in> A co B" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
135 |
apply (frule constrainsD2) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
136 |
apply (auto simp add: constrains_def component_eq_subset) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
137 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
138 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
139 |
(* Used in Guar.thy to show that programs are partially ordered*) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
140 |
(* bind_thm ("program_less_le", strict_component_def RS meta_eq_to_obj_eq);*) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
141 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
142 |
(*** preserves ***) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
143 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
144 |
lemma preserves_is_safety_prop [simp]: "safety_prop(preserves(f))" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
145 |
apply (unfold preserves_def safety_prop_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
146 |
apply (auto dest: ActsD simp add: stable_def constrains_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
147 |
apply (drule_tac c = act in subsetD, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
148 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
149 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
150 |
lemma preservesI [rule_format]: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
151 |
"\<forall>z. F \<in> stable({s \<in> state. f(s) = z}) ==> F \<in> preserves(f)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
152 |
apply (auto simp add: preserves_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
153 |
apply (blast dest: stableD2) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
154 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
155 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
156 |
lemma preserves_imp_eq: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
157 |
"[| F \<in> preserves(f); act \<in> Acts(F); <s,t> \<in> act |] ==> f(s) = f(t)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
158 |
apply (unfold preserves_def stable_def constrains_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
159 |
apply (subgoal_tac "s \<in> state & t \<in> state") |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
160 |
prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
161 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
162 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
163 |
lemma Join_preserves [iff]: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
164 |
"(F Join G \<in> preserves(v)) <-> |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
165 |
(programify(F) \<in> preserves(v) & programify(G) \<in> preserves(v))" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
166 |
by (auto simp add: preserves_def INT_iff) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
167 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
168 |
lemma JN_preserves [iff]: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
169 |
"(JOIN(I,F): preserves(v)) <-> (\<forall>i \<in> I. programify(F(i)):preserves(v))" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
170 |
by (auto simp add: JN_stable preserves_def INT_iff) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
171 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
172 |
lemma SKIP_preserves [iff]: "SKIP \<in> preserves(v)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
173 |
by (auto simp add: preserves_def INT_iff) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
174 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
175 |
lemma fun_pair_apply [simp]: "fun_pair(f,g,x) = <f(x), g(x)>" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
176 |
apply (unfold fun_pair_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
177 |
apply (simp (no_asm)) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
178 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
179 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
180 |
lemma preserves_fun_pair: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
181 |
"preserves(fun_pair(v,w)) = preserves(v) Int preserves(w)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
182 |
apply (rule equalityI) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
183 |
apply (auto simp add: preserves_def stable_def constrains_def, blast+) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
184 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
185 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
186 |
lemma preserves_fun_pair_iff [iff]: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
187 |
"F \<in> preserves(fun_pair(v, w)) <-> F \<in> preserves(v) Int preserves(w)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
188 |
by (simp add: preserves_fun_pair) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
189 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
190 |
lemma fun_pair_comp_distrib: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
191 |
"(fun_pair(f, g) comp h)(x) = fun_pair(f comp h, g comp h, x)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
192 |
by (simp add: fun_pair_def metacomp_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
193 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
194 |
lemma comp_apply [simp]: "(f comp g)(x) = f(g(x))" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
195 |
by (simp add: metacomp_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
196 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
197 |
lemma preserves_type: "preserves(v)<=program" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
198 |
by (unfold preserves_def, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
199 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
200 |
lemma preserves_into_program [TC]: "F \<in> preserves(f) ==> F \<in> program" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
201 |
by (blast intro: preserves_type [THEN subsetD]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
202 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
203 |
lemma subset_preserves_comp: "preserves(f) <= preserves(g comp f)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
204 |
apply (auto simp add: preserves_def stable_def constrains_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
205 |
apply (drule_tac x = "f (xb)" in spec) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
206 |
apply (drule_tac x = act in bspec, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
207 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
208 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
209 |
lemma imp_preserves_comp: "F \<in> preserves(f) ==> F \<in> preserves(g comp f)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
210 |
by (blast intro: subset_preserves_comp [THEN subsetD]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
211 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
212 |
lemma preserves_subset_stable: "preserves(f) <= stable({s \<in> state. P(f(s))})" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
213 |
apply (auto simp add: preserves_def stable_def constrains_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
214 |
apply (rename_tac s' s) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
215 |
apply (subgoal_tac "f (s) = f (s') ") |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
216 |
apply (force+) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
217 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
218 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
219 |
lemma preserves_imp_stable: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
220 |
"F \<in> preserves(f) ==> F \<in> stable({s \<in> state. P(f(s))})" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
221 |
by (blast intro: preserves_subset_stable [THEN subsetD]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
222 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
223 |
lemma preserves_imp_increasing: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
224 |
"[| F \<in> preserves(f); \<forall>x \<in> state. f(x):A |] ==> F \<in> Increasing.increasing(A, r, f)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
225 |
apply (unfold Increasing.increasing_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
226 |
apply (auto intro: preserves_into_program) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
227 |
apply (rule_tac P = "%x. <k, x>:r" in preserves_imp_stable, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
228 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
229 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
230 |
lemma preserves_id_subset_stable: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
231 |
"st_set(A) ==> preserves(%x. x) <= stable(A)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
232 |
apply (unfold preserves_def stable_def constrains_def, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
233 |
apply (drule_tac x = xb in spec) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
234 |
apply (drule_tac x = act in bspec) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
235 |
apply (auto dest: ActsD) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
236 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
237 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
238 |
lemma preserves_id_imp_stable: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
239 |
"[| F \<in> preserves(%x. x); st_set(A) |] ==> F \<in> stable(A)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
240 |
by (blast intro: preserves_id_subset_stable [THEN subsetD]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
241 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
242 |
(** Added by Sidi **) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
243 |
(** component_of **) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
244 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
245 |
(* component_of is stronger than component *) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
246 |
lemma component_of_imp_component: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
247 |
"F component_of H ==> F component H" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
248 |
apply (unfold component_def component_of_def, blast) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
249 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
250 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
251 |
(* component_of satisfies many of component's properties *) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
252 |
lemma component_of_refl [simp]: "F \<in> program ==> F component_of F" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
253 |
apply (unfold component_of_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
254 |
apply (rule_tac x = SKIP in exI, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
255 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
256 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
257 |
lemma component_of_SKIP [simp]: "F \<in> program ==>SKIP component_of F" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
258 |
apply (unfold component_of_def, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
259 |
apply (rule_tac x = F in exI, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
260 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
261 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
262 |
lemma component_of_trans: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
263 |
"[| F component_of G; G component_of H |] ==> F component_of H" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
264 |
apply (unfold component_of_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
265 |
apply (blast intro: Join_assoc [symmetric]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
266 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
267 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
268 |
(** localize **) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
269 |
lemma localize_Init_eq [simp]: "Init(localize(v,F)) = Init(F)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
270 |
by (unfold localize_def, simp) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
271 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
272 |
lemma localize_Acts_eq [simp]: "Acts(localize(v,F)) = Acts(F)" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
273 |
by (unfold localize_def, simp) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
274 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
275 |
lemma localize_AllowedActs_eq [simp]: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
276 |
"AllowedActs(localize(v,F)) = AllowedActs(F) Int (\<Union>G \<in> preserves(v). Acts(G))" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
277 |
apply (unfold localize_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
278 |
apply (rule equalityI) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
279 |
apply (auto dest: Acts_type [THEN subsetD]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
280 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
281 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
282 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
283 |
(** Theorems used in ClientImpl **) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
284 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
285 |
lemma stable_localTo_stable2: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
286 |
"[| F \<in> stable({s \<in> state. P(f(s), g(s))}); G \<in> preserves(f); G \<in> preserves(g) |] |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
287 |
==> F Join G \<in> stable({s \<in> state. P(f(s), g(s))})" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
288 |
apply (auto dest: ActsD preserves_into_program simp add: stable_def constrains_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
289 |
apply (case_tac "act \<in> Acts (F) ") |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
290 |
apply auto |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
291 |
apply (drule preserves_imp_eq) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
292 |
apply (drule_tac [3] preserves_imp_eq, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
293 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
294 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
295 |
lemma Increasing_preserves_Stable: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
296 |
"[| F \<in> stable({s \<in> state. <f(s), g(s)>:r}); G \<in> preserves(f); |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
297 |
F Join G \<in> Increasing(A, r, g); |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
298 |
\<forall>x \<in> state. f(x):A & g(x):A |] |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
299 |
==> F Join G \<in> Stable({s \<in> state. <f(s), g(s)>:r})" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
300 |
apply (auto simp add: stable_def Stable_def Increasing_def Constrains_def all_conj_distrib) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
301 |
apply (simp_all add: constrains_type [THEN subsetD] preserves_type [THEN subsetD]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
302 |
apply (blast intro: constrains_weaken) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
303 |
(*The G case remains*) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
304 |
apply (auto dest: ActsD simp add: preserves_def stable_def constrains_def ball_conj_distrib all_conj_distrib) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
305 |
(*We have a G-action, so delete assumptions about F-actions*) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
306 |
apply (erule_tac V = "\<forall>act \<in> Acts (F) . ?P (act)" in thin_rl) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
307 |
apply (erule_tac V = "\<forall>k\<in>A. \<forall>act \<in> Acts (F) . ?P (k,act)" in thin_rl) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
308 |
apply (subgoal_tac "f (x) = f (xa) ") |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
309 |
apply (auto dest!: bspec) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
310 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
311 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
312 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
313 |
(** Lemma used in AllocImpl **) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
314 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
315 |
lemma Constrains_UN_left: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
316 |
"[| \<forall>x \<in> I. F \<in> A(x) Co B; F \<in> program |] ==> F:(\<Union>x \<in> I. A(x)) Co B" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
317 |
by (unfold Constrains_def constrains_def, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
318 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
319 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
320 |
lemma stable_Join_Stable: |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
321 |
"[| F \<in> stable({s \<in> state. P(f(s), g(s))}); |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
322 |
\<forall>k \<in> A. F Join G \<in> Stable({s \<in> state. P(k, g(s))}); |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
323 |
G \<in> preserves(f); \<forall>s \<in> state. f(s):A|] |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
324 |
==> F Join G \<in> Stable({s \<in> state. P(f(s), g(s))})" |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
325 |
apply (unfold stable_def Stable_def preserves_def) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
326 |
apply (rule_tac A = "(\<Union>k \<in> A. {s \<in> state. f(s)=k} Int {s \<in> state. P (f (s), g (s))})" in Constrains_weaken_L) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
327 |
prefer 2 apply blast |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
328 |
apply (rule Constrains_UN_left, auto) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
329 |
apply (rule_tac A = "{s \<in> state. f(s)=k} Int {s \<in> state. P (f (s), g (s))} Int {s \<in> state. P (k, g (s))}" and A' = "({s \<in> state. f(s)=k} Un {s \<in> state. P (f (s), g (s))}) Int {s \<in> state. P (k, g (s))}" in Constrains_weaken) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
330 |
prefer 2 apply blast |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
331 |
prefer 2 apply blast |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
332 |
apply (rule Constrains_Int) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
333 |
apply (rule constrains_imp_Constrains) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
334 |
apply (auto simp add: constrains_type [THEN subsetD]) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
335 |
apply (blast intro: constrains_weaken) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
336 |
apply (drule_tac x = k in spec) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
337 |
apply (blast intro: constrains_weaken) |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
338 |
done |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
339 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
340 |
ML |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
341 |
{* |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
342 |
val component_of_def = thm "component_of_def"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
343 |
val component_def = thm "component_def"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
344 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
345 |
val componentI = thm "componentI"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
346 |
val component_eq_subset = thm "component_eq_subset"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
347 |
val component_SKIP = thm "component_SKIP"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
348 |
val component_refl = thm "component_refl"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
349 |
val SKIP_minimal = thm "SKIP_minimal"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
350 |
val component_Join1 = thm "component_Join1"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
351 |
val component_Join2 = thm "component_Join2"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
352 |
val Join_absorb1 = thm "Join_absorb1"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
353 |
val Join_absorb2 = thm "Join_absorb2"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
354 |
val JN_component_iff = thm "JN_component_iff"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
355 |
val component_JN = thm "component_JN"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
356 |
val component_trans = thm "component_trans"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
357 |
val component_antisym = thm "component_antisym"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
358 |
val Join_component_iff = thm "Join_component_iff"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
359 |
val component_constrains = thm "component_constrains"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
360 |
val preserves_is_safety_prop = thm "preserves_is_safety_prop"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
361 |
val preservesI = thm "preservesI"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
362 |
val preserves_imp_eq = thm "preserves_imp_eq"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
363 |
val Join_preserves = thm "Join_preserves"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
364 |
val JN_preserves = thm "JN_preserves"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
365 |
val SKIP_preserves = thm "SKIP_preserves"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
366 |
val fun_pair_apply = thm "fun_pair_apply"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
367 |
val preserves_fun_pair = thm "preserves_fun_pair"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
368 |
val preserves_fun_pair_iff = thm "preserves_fun_pair_iff"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
369 |
val fun_pair_comp_distrib = thm "fun_pair_comp_distrib"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
370 |
val comp_apply = thm "comp_apply"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
371 |
val preserves_type = thm "preserves_type"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
372 |
val preserves_into_program = thm "preserves_into_program"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
373 |
val subset_preserves_comp = thm "subset_preserves_comp"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
374 |
val imp_preserves_comp = thm "imp_preserves_comp"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
375 |
val preserves_subset_stable = thm "preserves_subset_stable"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
376 |
val preserves_imp_stable = thm "preserves_imp_stable"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
377 |
val preserves_imp_increasing = thm "preserves_imp_increasing"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
378 |
val preserves_id_subset_stable = thm "preserves_id_subset_stable"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
379 |
val preserves_id_imp_stable = thm "preserves_id_imp_stable"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
380 |
val component_of_imp_component = thm "component_of_imp_component"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
381 |
val component_of_refl = thm "component_of_refl"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
382 |
val component_of_SKIP = thm "component_of_SKIP"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
383 |
val component_of_trans = thm "component_of_trans"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
384 |
val localize_Init_eq = thm "localize_Init_eq"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
385 |
val localize_Acts_eq = thm "localize_Acts_eq"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
386 |
val localize_AllowedActs_eq = thm "localize_AllowedActs_eq"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
387 |
val stable_localTo_stable2 = thm "stable_localTo_stable2"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
388 |
val Increasing_preserves_Stable = thm "Increasing_preserves_Stable"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
389 |
val Constrains_UN_left = thm "Constrains_UN_left"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
390 |
val stable_Join_Stable = thm "stable_Join_Stable"; |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
391 |
*} |
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
392 |
|
24382760fd89
converting more theories to Isar scripts, and tidying
paulson
parents:
14046
diff
changeset
|
393 |
end |