author | paulson |
Fri, 29 Oct 2004 15:16:02 +0200 | |
changeset 15270 | 8b3f707a78a7 |
parent 14981 | e73f8140af78 |
child 15531 | 08c8dad8e399 |
permissions | -rw-r--r-- |
12857 | 1 |
(* Title: HOL/Bali/Basis.thy |
12854 | 2 |
ID: $Id$ |
3 |
Author: David von Oheimb |
|
4 |
||
5 |
*) |
|
6 |
header {* Definitions extending HOL as logical basis of Bali *} |
|
7 |
||
8 |
theory Basis = Main: |
|
9 |
||
10 |
ML_setup {* |
|
11 |
Unify.search_bound := 40; |
|
12 |
Unify.trace_bound := 40; |
|
13 |
*} |
|
14 |
(*print_depth 100;*) |
|
15 |
(*Syntax.ambiguity_level := 1;*) |
|
16 |
||
17 |
section "misc" |
|
18 |
||
19 |
declare same_fstI [intro!] (*### TO HOL/Wellfounded_Relations *) |
|
20 |
||
21 |
ML {* |
|
13462 | 22 |
fun cond_simproc name pat pred thm = Simplifier.simproc (Thm.sign_of_thm thm) name [pat] |
23 |
(fn _ => fn _ => fn t => if pred t then None else Some (mk_meta_eq thm)); |
|
12854 | 24 |
*} |
25 |
||
26 |
declare split_if_asm [split] option.split [split] option.split_asm [split] |
|
27 |
ML {* |
|
28 |
simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac) |
|
29 |
*} |
|
30 |
declare if_weak_cong [cong del] option.weak_case_cong [cong del] |
|
31 |
declare length_Suc_conv [iff]; |
|
32 |
||
33 |
(*###to be phased out *) |
|
34 |
ML {* |
|
35 |
bind_thm ("make_imp", rearrange_prems [1,0] mp) |
|
36 |
*} |
|
37 |
||
38 |
lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}" |
|
39 |
apply auto |
|
40 |
done |
|
41 |
||
42 |
lemma subset_insertD: |
|
43 |
"A <= insert x B ==> A <= B & x ~: A | (EX B'. A = insert x B' & B' <= B)" |
|
44 |
apply (case_tac "x:A") |
|
45 |
apply (rule disjI2) |
|
46 |
apply (rule_tac x = "A-{x}" in exI) |
|
47 |
apply fast+ |
|
48 |
done |
|
49 |
||
50 |
syntax |
|
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12919
diff
changeset
|
51 |
"3" :: nat ("3") |
12854 | 52 |
"4" :: nat ("4") |
53 |
translations |
|
54 |
"3" == "Suc 2" |
|
55 |
"4" == "Suc 3" |
|
56 |
||
57 |
(*unused*) |
|
58 |
lemma range_bool_domain: "range f = {f True, f False}" |
|
59 |
apply auto |
|
60 |
apply (case_tac "xa") |
|
61 |
apply auto |
|
62 |
done |
|
63 |
||
13867 | 64 |
(* irrefl_tranclI in Transitive_Closure.thy is more general *) |
12854 | 65 |
lemma irrefl_tranclI': "r^-1 Int r^+ = {} ==> !x. (x, x) ~: r^+" |
13867 | 66 |
by(blast elim: tranclE dest: trancl_into_rtrancl) |
67 |
||
12854 | 68 |
|
69 |
lemma trancl_rtrancl_trancl: |
|
70 |
"\<lbrakk>(x,y)\<in>r^+; (y,z)\<in>r^*\<rbrakk> \<Longrightarrow> (x,z)\<in>r^+" |
|
71 |
by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2) |
|
72 |
||
73 |
lemma rtrancl_into_trancl3: |
|
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12919
diff
changeset
|
74 |
"\<lbrakk>(a,b)\<in>r^*; a\<noteq>b\<rbrakk> \<Longrightarrow> (a,b)\<in>r^+" |
12854 | 75 |
apply (drule rtranclD) |
76 |
apply auto |
|
77 |
done |
|
78 |
||
79 |
lemma rtrancl_into_rtrancl2: |
|
80 |
"\<lbrakk> (a, b) \<in> r; (b, c) \<in> r^* \<rbrakk> \<Longrightarrow> (a, c) \<in> r^*" |
|
81 |
by (auto intro: r_into_rtrancl rtrancl_trans) |
|
82 |
||
83 |
lemma triangle_lemma: |
|
84 |
"\<lbrakk> \<And> a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b=c; (a,x)\<in>r\<^sup>*; (a,y)\<in>r\<^sup>*\<rbrakk> |
|
85 |
\<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*" |
|
86 |
proof - |
|
87 |
note converse_rtrancl_induct = converse_rtrancl_induct [consumes 1] |
|
88 |
note converse_rtranclE = converse_rtranclE [consumes 1] |
|
89 |
assume unique: "\<And> a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b=c" |
|
90 |
assume "(a,x)\<in>r\<^sup>*" |
|
91 |
then show "(a,y)\<in>r\<^sup>* \<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*" |
|
92 |
proof (induct rule: converse_rtrancl_induct) |
|
93 |
assume "(x,y)\<in>r\<^sup>*" |
|
94 |
then show ?thesis |
|
95 |
by blast |
|
96 |
next |
|
97 |
fix a v |
|
98 |
assume a_v_r: "(a, v) \<in> r" and |
|
99 |
v_x_rt: "(v, x) \<in> r\<^sup>*" and |
|
100 |
a_y_rt: "(a, y) \<in> r\<^sup>*" and |
|
101 |
hyp: "(v, y) \<in> r\<^sup>* \<Longrightarrow> (x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*" |
|
102 |
from a_y_rt |
|
103 |
show "(x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*" |
|
104 |
proof (cases rule: converse_rtranclE) |
|
105 |
assume "a=y" |
|
106 |
with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*" |
|
107 |
by (auto intro: r_into_rtrancl rtrancl_trans) |
|
108 |
then show ?thesis |
|
109 |
by blast |
|
110 |
next |
|
111 |
fix w |
|
112 |
assume a_w_r: "(a, w) \<in> r" and |
|
113 |
w_y_rt: "(w, y) \<in> r\<^sup>*" |
|
114 |
from a_v_r a_w_r unique |
|
115 |
have "v=w" |
|
116 |
by auto |
|
117 |
with w_y_rt hyp |
|
118 |
show ?thesis |
|
119 |
by blast |
|
120 |
qed |
|
121 |
qed |
|
122 |
qed |
|
123 |
||
124 |
||
125 |
lemma rtrancl_cases [consumes 1, case_names Refl Trancl]: |
|
126 |
"\<lbrakk>(a,b)\<in>r\<^sup>*; a = b \<Longrightarrow> P; (a,b)\<in>r\<^sup>+ \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
|
127 |
apply (erule rtranclE) |
|
128 |
apply (auto dest: rtrancl_into_trancl1) |
|
129 |
done |
|
130 |
||
131 |
(* ### To Transitive_Closure *) |
|
132 |
theorems converse_rtrancl_induct |
|
133 |
= converse_rtrancl_induct [consumes 1,case_names Id Step] |
|
134 |
||
135 |
theorems converse_trancl_induct |
|
136 |
= converse_trancl_induct [consumes 1,case_names Single Step] |
|
137 |
||
138 |
(* context (theory "Set") *) |
|
139 |
lemma Ball_weaken:"\<lbrakk>Ball s P;\<And> x. P x\<longrightarrow>Q x\<rbrakk>\<Longrightarrow>Ball s Q" |
|
140 |
by auto |
|
141 |
||
142 |
(* context (theory "Finite") *) |
|
143 |
lemma finite_SetCompr2: "[| finite (Collect P); !y. P y --> finite (range (f y)) |] ==> |
|
144 |
finite {f y x |x y. P y}" |
|
145 |
apply (subgoal_tac "{f y x |x y. P y} = UNION (Collect P) (%y. range (f y))") |
|
146 |
prefer 2 apply fast |
|
147 |
apply (erule ssubst) |
|
148 |
apply (erule finite_UN_I) |
|
149 |
apply fast |
|
150 |
done |
|
151 |
||
152 |
||
153 |
(* ### TO theory "List" *) |
|
154 |
lemma list_all2_trans: "\<forall> a b c. P1 a b \<longrightarrow> P2 b c \<longrightarrow> P3 a c \<Longrightarrow> |
|
155 |
\<forall>xs2 xs3. list_all2 P1 xs1 xs2 \<longrightarrow> list_all2 P2 xs2 xs3 \<longrightarrow> list_all2 P3 xs1 xs3" |
|
156 |
apply (induct_tac "xs1") |
|
157 |
apply simp |
|
158 |
apply (rule allI) |
|
159 |
apply (induct_tac "xs2") |
|
160 |
apply simp |
|
161 |
apply (rule allI) |
|
162 |
apply (induct_tac "xs3") |
|
163 |
apply auto |
|
164 |
done |
|
165 |
||
166 |
||
167 |
section "pairs" |
|
168 |
||
169 |
lemma surjective_pairing5: "p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))), |
|
170 |
snd (snd (snd (snd p))))" |
|
171 |
apply auto |
|
172 |
done |
|
173 |
||
174 |
lemma fst_splitE [elim!]: |
|
175 |
"[| fst s' = x'; !!x s. [| s' = (x,s); x = x' |] ==> Q |] ==> Q" |
|
176 |
apply (cut_tac p = "s'" in surjective_pairing) |
|
177 |
apply auto |
|
178 |
done |
|
179 |
||
180 |
lemma fst_in_set_lemma [rule_format (no_asm)]: "(x, y) : set l --> x : fst ` set l" |
|
181 |
apply (induct_tac "l") |
|
182 |
apply auto |
|
183 |
done |
|
184 |
||
185 |
||
186 |
section "quantifiers" |
|
187 |
||
188 |
(*###to be phased out *) |
|
189 |
ML {* |
|
190 |
fun noAll_simpset () = simpset() setmksimps |
|
191 |
mksimps (filter (fn (x,_) => x<>"All") mksimps_pairs) |
|
192 |
*} |
|
193 |
||
194 |
lemma All_Ex_refl_eq2 [simp]: |
|
195 |
"(!x. (? b. x = f b & Q b) \<longrightarrow> P x) = (!b. Q b --> P (f b))" |
|
196 |
apply auto |
|
197 |
done |
|
198 |
||
199 |
lemma ex_ex_miniscope1 [simp]: |
|
200 |
"(EX w v. P w v & Q v) = (EX v. (EX w. P w v) & Q v)" |
|
201 |
apply auto |
|
202 |
done |
|
203 |
||
204 |
lemma ex_miniscope2 [simp]: |
|
205 |
"(EX v. P v & Q & R v) = (Q & (EX v. P v & R v))" |
|
206 |
apply auto |
|
207 |
done |
|
208 |
||
209 |
lemma ex_reorder31: "(\<exists>z x y. P x y z) = (\<exists>x y z. P x y z)" |
|
210 |
apply auto |
|
211 |
done |
|
212 |
||
213 |
lemma All_Ex_refl_eq1 [simp]: "(!x. (? b. x = f b) --> P x) = (!b. P (f b))" |
|
214 |
apply auto |
|
215 |
done |
|
216 |
||
217 |
||
218 |
section "sums" |
|
219 |
||
220 |
hide const In0 In1 |
|
221 |
||
222 |
syntax |
|
223 |
fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80) |
|
224 |
translations |
|
225 |
"fun_sum" == "sum_case" |
|
226 |
||
227 |
consts the_Inl :: "'a + 'b \<Rightarrow> 'a" |
|
228 |
the_Inr :: "'a + 'b \<Rightarrow> 'b" |
|
229 |
primrec "the_Inl (Inl a) = a" |
|
230 |
primrec "the_Inr (Inr b) = b" |
|
231 |
||
232 |
datatype ('a, 'b, 'c) sum3 = In1 'a | In2 'b | In3 'c |
|
233 |
||
234 |
consts the_In1 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'a" |
|
235 |
the_In2 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'b" |
|
236 |
the_In3 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'c" |
|
237 |
primrec "the_In1 (In1 a) = a" |
|
238 |
primrec "the_In2 (In2 b) = b" |
|
239 |
primrec "the_In3 (In3 c) = c" |
|
240 |
||
241 |
syntax |
|
242 |
In1l :: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3" |
|
243 |
In1r :: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3" |
|
244 |
translations |
|
245 |
"In1l e" == "In1 (Inl e)" |
|
246 |
"In1r c" == "In1 (Inr c)" |
|
247 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13462
diff
changeset
|
248 |
syntax the_In1l :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'al" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13462
diff
changeset
|
249 |
the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13462
diff
changeset
|
250 |
translations |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13462
diff
changeset
|
251 |
"the_In1l" == "the_Inl \<circ> the_In1" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13462
diff
changeset
|
252 |
"the_In1r" == "the_Inr \<circ> the_In1" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13462
diff
changeset
|
253 |
|
12854 | 254 |
ML {* |
255 |
fun sum3_instantiate thm = map (fn s => simplify(simpset()delsimps[not_None_eq]) |
|
256 |
(read_instantiate [("t","In"^s^" ?x")] thm)) ["1l","2","3","1r"] |
|
257 |
*} |
|
258 |
(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *) |
|
259 |
||
260 |
translations |
|
12919 | 261 |
"option"<= (type) "Datatype.option" |
12854 | 262 |
"list" <= (type) "List.list" |
263 |
"sum3" <= (type) "Basis.sum3" |
|
264 |
||
265 |
||
266 |
section "quantifiers for option type" |
|
267 |
||
268 |
syntax |
|
269 |
Oall :: "[pttrn, 'a option, bool] => bool" ("(3! _:_:/ _)" [0,0,10] 10) |
|
270 |
Oex :: "[pttrn, 'a option, bool] => bool" ("(3? _:_:/ _)" [0,0,10] 10) |
|
271 |
||
272 |
syntax (symbols) |
|
273 |
Oall :: "[pttrn, 'a option, bool] => bool" ("(3\<forall>_\<in>_:/ _)" [0,0,10] 10) |
|
274 |
Oex :: "[pttrn, 'a option, bool] => bool" ("(3\<exists>_\<in>_:/ _)" [0,0,10] 10) |
|
275 |
||
276 |
translations |
|
277 |
"! x:A: P" == "! x:o2s A. P" |
|
278 |
"? x:A: P" == "? x:o2s A. P" |
|
279 |
||
280 |
||
281 |
section "unique association lists" |
|
282 |
||
283 |
constdefs |
|
284 |
unique :: "('a \<times> 'b) list \<Rightarrow> bool" |
|
12893 | 285 |
"unique \<equiv> distinct \<circ> map fst" |
12854 | 286 |
|
287 |
lemma uniqueD [rule_format (no_asm)]: |
|
288 |
"unique l--> (!x y. (x,y):set l --> (!x' y'. (x',y'):set l --> x=x'--> y=y'))" |
|
289 |
apply (unfold unique_def o_def) |
|
290 |
apply (induct_tac "l") |
|
291 |
apply (auto dest: fst_in_set_lemma) |
|
292 |
done |
|
293 |
||
294 |
lemma unique_Nil [simp]: "unique []" |
|
295 |
apply (unfold unique_def) |
|
296 |
apply (simp (no_asm)) |
|
297 |
done |
|
298 |
||
299 |
lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))" |
|
300 |
apply (unfold unique_def) |
|
301 |
apply (auto dest: fst_in_set_lemma) |
|
302 |
done |
|
303 |
||
304 |
lemmas unique_ConsI = conjI [THEN unique_Cons [THEN iffD2], standard] |
|
305 |
||
306 |
lemma unique_single [simp]: "!!p. unique [p]" |
|
307 |
apply auto |
|
308 |
done |
|
309 |
||
310 |
lemma unique_ConsD: "unique (x#xs) ==> unique xs" |
|
311 |
apply (simp add: unique_def) |
|
312 |
done |
|
313 |
||
314 |
lemma unique_append [rule_format (no_asm)]: "unique l' ==> unique l --> |
|
315 |
(!(x,y):set l. !(x',y'):set l'. x' ~= x) --> unique (l @ l')" |
|
316 |
apply (induct_tac "l") |
|
317 |
apply (auto dest: fst_in_set_lemma) |
|
318 |
done |
|
319 |
||
320 |
lemma unique_map_inj [rule_format (no_asm)]: "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)" |
|
321 |
apply (induct_tac "l") |
|
322 |
apply (auto dest: fst_in_set_lemma simp add: inj_eq) |
|
323 |
done |
|
324 |
||
325 |
lemma map_of_SomeI [rule_format (no_asm)]: "unique l --> (k, x) : set l --> map_of l k = Some x" |
|
326 |
apply (induct_tac "l") |
|
327 |
apply auto |
|
328 |
done |
|
329 |
||
330 |
||
331 |
section "list patterns" |
|
332 |
||
333 |
consts |
|
334 |
lsplit :: "[['a, 'a list] => 'b, 'a list] => 'b" |
|
335 |
defs |
|
336 |
lsplit_def: "lsplit == %f l. f (hd l) (tl l)" |
|
337 |
(* list patterns -- extends pre-defined type "pttrn" used in abstractions *) |
|
338 |
syntax |
|
339 |
"_lpttrn" :: "[pttrn,pttrn] => pttrn" ("_#/_" [901,900] 900) |
|
340 |
translations |
|
341 |
"%y#x#xs. b" == "lsplit (%y x#xs. b)" |
|
342 |
"%x#xs . b" == "lsplit (%x xs . b)" |
|
343 |
||
344 |
lemma lsplit [simp]: "lsplit c (x#xs) = c x xs" |
|
345 |
apply (unfold lsplit_def) |
|
346 |
apply (simp (no_asm)) |
|
347 |
done |
|
348 |
||
349 |
lemma lsplit2 [simp]: "lsplit P (x#xs) y z = P x xs y z" |
|
350 |
apply (unfold lsplit_def) |
|
351 |
apply simp |
|
352 |
done |
|
353 |
||
354 |
||
355 |
section "dummy pattern for quantifiers, let, etc." |
|
356 |
||
357 |
syntax |
|
358 |
"@dummy_pat" :: pttrn ("'_") |
|
359 |
||
360 |
parse_translation {* |
|
361 |
let fun dummy_pat_tr [] = Free ("_",dummyT) |
|
362 |
| dummy_pat_tr ts = raise TERM ("dummy_pat_tr", ts); |
|
363 |
in [("@dummy_pat", dummy_pat_tr)] |
|
364 |
end |
|
365 |
*} |
|
366 |
||
367 |
end |