author | blanchet |
Mon, 03 Mar 2014 12:48:20 +0100 | |
changeset 55855 | 98ad5680173a |
parent 55518 | 1ddb2edf5ceb |
child 57983 | 6edc3529bb4e |
permissions | -rw-r--r-- |
12857 | 1 |
(* Title: HOL/Bali/Basis.thy |
12854 | 2 |
Author: David von Oheimb |
3 |
*) |
|
4 |
header {* Definitions extending HOL as logical basis of Bali *} |
|
5 |
||
45151 | 6 |
theory Basis |
7 |
imports Main "~~/src/HOL/Library/Old_Recdef" |
|
8 |
begin |
|
12854 | 9 |
|
10 |
section "misc" |
|
11 |
||
51304 | 12 |
ML {* fun strip_tac i = REPEAT (resolve_tac [impI, allI] i) *} |
13 |
||
12854 | 14 |
declare split_if_asm [split] option.split [split] option.split_asm [split] |
52037 | 15 |
setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *} |
12854 | 16 |
declare if_weak_cong [cong del] option.weak_case_cong [cong del] |
18447 | 17 |
declare length_Suc_conv [iff] |
18 |
||
12854 | 19 |
lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}" |
45151 | 20 |
by auto |
12854 | 21 |
|
45151 | 22 |
lemma subset_insertD: "A \<subseteq> insert x B \<Longrightarrow> A \<subseteq> B \<and> x \<notin> A \<or> (\<exists>B'. A = insert x B' \<and> B' \<subseteq> B)" |
23 |
apply (case_tac "x \<in> A") |
|
24 |
apply (rule disjI2) |
|
25 |
apply (rule_tac x = "A - {x}" in exI) |
|
26 |
apply fast+ |
|
27 |
done |
|
12854 | 28 |
|
45151 | 29 |
abbreviation nat3 :: nat ("3") where "3 \<equiv> Suc 2" |
30 |
abbreviation nat4 :: nat ("4") where "4 \<equiv> Suc 3" |
|
12854 | 31 |
|
13867 | 32 |
(* irrefl_tranclI in Transitive_Closure.thy is more general *) |
45151 | 33 |
lemma irrefl_tranclI': "r\<inverse> \<inter> r\<^sup>+ = {} \<Longrightarrow> \<forall>x. (x, x) \<notin> r\<^sup>+" |
34 |
by (blast elim: tranclE dest: trancl_into_rtrancl) |
|
13867 | 35 |
|
12854 | 36 |
|
45151 | 37 |
lemma trancl_rtrancl_trancl: "\<lbrakk>(x, y) \<in> r\<^sup>+; (y, z) \<in> r\<^sup>*\<rbrakk> \<Longrightarrow> (x, z) \<in> r\<^sup>+" |
38 |
by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2) |
|
12854 | 39 |
|
45151 | 40 |
lemma rtrancl_into_trancl3: "\<lbrakk>(a, b) \<in> r\<^sup>*; a \<noteq> b\<rbrakk> \<Longrightarrow> (a, b) \<in> r\<^sup>+" |
41 |
apply (drule rtranclD) |
|
42 |
apply auto |
|
43 |
done |
|
12854 | 44 |
|
45151 | 45 |
lemma rtrancl_into_rtrancl2: "\<lbrakk>(a, b) \<in> r; (b, c) \<in> r\<^sup>*\<rbrakk> \<Longrightarrow> (a, c) \<in> r\<^sup>*" |
46 |
by (auto intro: rtrancl_trans) |
|
12854 | 47 |
|
48 |
lemma triangle_lemma: |
|
45151 | 49 |
assumes unique: "\<And>a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b = c" |
50 |
and ax: "(a,x)\<in>r\<^sup>*" and ay: "(a,y)\<in>r\<^sup>*" |
|
51 |
shows "(x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*" |
|
52 |
using ax ay |
|
53 |
proof (induct rule: converse_rtrancl_induct) |
|
54 |
assume "(x,y)\<in>r\<^sup>*" |
|
55 |
then show ?thesis by blast |
|
56 |
next |
|
57 |
fix a v |
|
58 |
assume a_v_r: "(a, v) \<in> r" |
|
59 |
and v_x_rt: "(v, x) \<in> r\<^sup>*" |
|
60 |
and a_y_rt: "(a, y) \<in> r\<^sup>*" |
|
61 |
and hyp: "(v, y) \<in> r\<^sup>* \<Longrightarrow> (x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*" |
|
62 |
from a_y_rt show "(x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*" |
|
63 |
proof (cases rule: converse_rtranclE) |
|
64 |
assume "a = y" |
|
65 |
with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*" |
|
66 |
by (auto intro: rtrancl_trans) |
|
67 |
then show ?thesis by blast |
|
12854 | 68 |
next |
45151 | 69 |
fix w |
70 |
assume a_w_r: "(a, w) \<in> r" |
|
71 |
and w_y_rt: "(w, y) \<in> r\<^sup>*" |
|
72 |
from a_v_r a_w_r unique have "v=w" by auto |
|
73 |
with w_y_rt hyp show ?thesis by blast |
|
12854 | 74 |
qed |
75 |
qed |
|
76 |
||
77 |
||
45151 | 78 |
lemma rtrancl_cases: |
79 |
assumes "(a,b)\<in>r\<^sup>*" |
|
80 |
obtains (Refl) "a = b" |
|
81 |
| (Trancl) "(a,b)\<in>r\<^sup>+" |
|
82 |
apply (rule rtranclE [OF assms]) |
|
83 |
apply (auto dest: rtrancl_into_trancl1) |
|
84 |
done |
|
12854 | 85 |
|
45151 | 86 |
lemma Ball_weaken: "\<lbrakk>Ball s P; \<And> x. P x\<longrightarrow>Q x\<rbrakk>\<Longrightarrow>Ball s Q" |
87 |
by auto |
|
12854 | 88 |
|
45151 | 89 |
lemma finite_SetCompr2: |
90 |
"finite (Collect P) \<Longrightarrow> \<forall>y. P y \<longrightarrow> finite (range (f y)) \<Longrightarrow> |
|
91 |
finite {f y x |x y. P y}" |
|
92 |
apply (subgoal_tac "{f y x |x y. P y} = UNION (Collect P) (\<lambda>y. range (f y))") |
|
93 |
prefer 2 apply fast |
|
94 |
apply (erule ssubst) |
|
95 |
apply (erule finite_UN_I) |
|
96 |
apply fast |
|
97 |
done |
|
12854 | 98 |
|
45151 | 99 |
lemma list_all2_trans: "\<forall>a b c. P1 a b \<longrightarrow> P2 b c \<longrightarrow> P3 a c \<Longrightarrow> |
100 |
\<forall>xs2 xs3. list_all2 P1 xs1 xs2 \<longrightarrow> list_all2 P2 xs2 xs3 \<longrightarrow> list_all2 P3 xs1 xs3" |
|
101 |
apply (induct_tac xs1) |
|
102 |
apply simp |
|
103 |
apply (rule allI) |
|
104 |
apply (induct_tac xs2) |
|
105 |
apply simp |
|
106 |
apply (rule allI) |
|
107 |
apply (induct_tac xs3) |
|
108 |
apply auto |
|
109 |
done |
|
12854 | 110 |
|
111 |
||
112 |
section "pairs" |
|
113 |
||
45151 | 114 |
lemma surjective_pairing5: |
115 |
"p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))), |
|
116 |
snd (snd (snd (snd p))))" |
|
117 |
by auto |
|
12854 | 118 |
|
45151 | 119 |
lemma fst_splitE [elim!]: |
120 |
assumes "fst s' = x'" |
|
121 |
obtains x s where "s' = (x,s)" and "x = x'" |
|
122 |
using assms by (cases s') auto |
|
12854 | 123 |
|
45151 | 124 |
lemma fst_in_set_lemma: "(x, y) : set l \<Longrightarrow> x : fst ` set l" |
125 |
by (induct l) auto |
|
12854 | 126 |
|
127 |
||
128 |
section "quantifiers" |
|
129 |
||
45151 | 130 |
lemma All_Ex_refl_eq2 [simp]: "(\<forall>x. (\<exists>b. x = f b \<and> Q b) \<longrightarrow> P x) = (\<forall>b. Q b \<longrightarrow> P (f b))" |
131 |
by auto |
|
12854 | 132 |
|
45151 | 133 |
lemma ex_ex_miniscope1 [simp]: "(\<exists>w v. P w v \<and> Q v) = (\<exists>v. (\<exists>w. P w v) \<and> Q v)" |
134 |
by auto |
|
12854 | 135 |
|
45151 | 136 |
lemma ex_miniscope2 [simp]: "(\<exists>v. P v \<and> Q \<and> R v) = (Q \<and> (\<exists>v. P v \<and> R v))" |
137 |
by auto |
|
12854 | 138 |
|
139 |
lemma ex_reorder31: "(\<exists>z x y. P x y z) = (\<exists>x y z. P x y z)" |
|
45151 | 140 |
by auto |
12854 | 141 |
|
45151 | 142 |
lemma All_Ex_refl_eq1 [simp]: "(\<forall>x. (\<exists>b. x = f b) \<longrightarrow> P x) = (\<forall>b. P (f b))" |
143 |
by auto |
|
12854 | 144 |
|
145 |
||
146 |
section "sums" |
|
147 |
||
36176
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents:
35431
diff
changeset
|
148 |
hide_const In0 In1 |
12854 | 149 |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
55151
diff
changeset
|
150 |
notation case_sum (infixr "'(+')"80) |
12854 | 151 |
|
45151 | 152 |
primrec the_Inl :: "'a + 'b \<Rightarrow> 'a" |
37956 | 153 |
where "the_Inl (Inl a) = a" |
154 |
||
45151 | 155 |
primrec the_Inr :: "'a + 'b \<Rightarrow> 'b" |
37956 | 156 |
where "the_Inr (Inr b) = b" |
12854 | 157 |
|
158 |
datatype ('a, 'b, 'c) sum3 = In1 'a | In2 'b | In3 'c |
|
159 |
||
45151 | 160 |
primrec the_In1 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'a" |
37956 | 161 |
where "the_In1 (In1 a) = a" |
162 |
||
45151 | 163 |
primrec the_In2 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'b" |
37956 | 164 |
where "the_In2 (In2 b) = b" |
165 |
||
45151 | 166 |
primrec the_In3 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'c" |
37956 | 167 |
where "the_In3 (In3 c) = c" |
12854 | 168 |
|
45151 | 169 |
abbreviation In1l :: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3" |
170 |
where "In1l e \<equiv> In1 (Inl e)" |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
34915
diff
changeset
|
171 |
|
45151 | 172 |
abbreviation In1r :: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3" |
173 |
where "In1r c \<equiv> In1 (Inr c)" |
|
12854 | 174 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
34915
diff
changeset
|
175 |
abbreviation the_In1l :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'al" |
45151 | 176 |
where "the_In1l \<equiv> the_Inl \<circ> the_In1" |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
34915
diff
changeset
|
177 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
34915
diff
changeset
|
178 |
abbreviation the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar" |
45151 | 179 |
where "the_In1r \<equiv> the_Inr \<circ> the_In1" |
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
|
180 |
|
12854 | 181 |
ML {* |
55151 | 182 |
fun sum3_instantiate ctxt thm = |
183 |
map (fn s => |
|
184 |
simplify (ctxt delsimps @{thms not_None_eq}) |
|
185 |
(Rule_Insts.read_instantiate ctxt [(("t", 0), "In" ^ s ^ " x")] ["x"] thm)) |
|
186 |
["1l","2","3","1r"] |
|
12854 | 187 |
*} |
188 |
(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *) |
|
189 |
||
190 |
||
191 |
section "quantifiers for option type" |
|
192 |
||
193 |
syntax |
|
45151 | 194 |
"_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" ("(3! _:_:/ _)" [0,0,10] 10) |
195 |
"_Oex" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" ("(3? _:_:/ _)" [0,0,10] 10) |
|
12854 | 196 |
|
197 |
syntax (symbols) |
|
45151 | 198 |
"_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" ("(3\<forall>_\<in>_:/ _)" [0,0,10] 10) |
199 |
"_Oex" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" ("(3\<exists>_\<in>_:/ _)" [0,0,10] 10) |
|
12854 | 200 |
|
201 |
translations |
|
55518
1ddb2edf5ceb
folded 'Option.set' into BNF-generated 'set_option'
blanchet
parents:
55414
diff
changeset
|
202 |
"\<forall>x\<in>A: P" \<rightleftharpoons> "\<forall>x\<in>CONST set_option A. P" |
1ddb2edf5ceb
folded 'Option.set' into BNF-generated 'set_option'
blanchet
parents:
55414
diff
changeset
|
203 |
"\<exists>x\<in>A: P" \<rightleftharpoons> "\<exists>x\<in>CONST set_option A. P" |
45151 | 204 |
|
12854 | 205 |
|
19323 | 206 |
section "Special map update" |
207 |
||
208 |
text{* Deemed too special for theory Map. *} |
|
209 |
||
45151 | 210 |
definition chg_map :: "('b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> ('a \<rightharpoonup> 'b) \<Rightarrow> ('a \<rightharpoonup> 'b)" |
211 |
where "chg_map f a m = (case m a of None \<Rightarrow> m | Some b \<Rightarrow> m(a\<mapsto>f b))" |
|
19323 | 212 |
|
45151 | 213 |
lemma chg_map_new[simp]: "m a = None \<Longrightarrow> chg_map f a m = m" |
214 |
unfolding chg_map_def by auto |
|
19323 | 215 |
|
45151 | 216 |
lemma chg_map_upd[simp]: "m a = Some b \<Longrightarrow> chg_map f a m = m(a\<mapsto>f b)" |
217 |
unfolding chg_map_def by auto |
|
19323 | 218 |
|
219 |
lemma chg_map_other [simp]: "a \<noteq> b \<Longrightarrow> chg_map f a m b = m b" |
|
45151 | 220 |
by (auto simp: chg_map_def) |
19323 | 221 |
|
12854 | 222 |
|
223 |
section "unique association lists" |
|
224 |
||
45151 | 225 |
definition unique :: "('a \<times> 'b) list \<Rightarrow> bool" |
37956 | 226 |
where "unique = distinct \<circ> map fst" |
12854 | 227 |
|
45151 | 228 |
lemma uniqueD: "unique l \<Longrightarrow> (x, y) \<in> set l \<Longrightarrow> (x', y') \<in> set l \<Longrightarrow> x = x' \<Longrightarrow> y = y'" |
229 |
unfolding unique_def o_def |
|
230 |
by (induct l) (auto dest: fst_in_set_lemma) |
|
12854 | 231 |
|
232 |
lemma unique_Nil [simp]: "unique []" |
|
45151 | 233 |
by (simp add: unique_def) |
12854 | 234 |
|
45151 | 235 |
lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l \<and> (\<forall>y. (x,y) \<notin> set l))" |
236 |
by (auto simp: unique_def dest: fst_in_set_lemma) |
|
12854 | 237 |
|
45151 | 238 |
lemma unique_ConsD: "unique (x#xs) \<Longrightarrow> unique xs" |
239 |
by (simp add: unique_def) |
|
12854 | 240 |
|
45151 | 241 |
lemma unique_single [simp]: "\<And>p. unique [p]" |
242 |
by simp |
|
12854 | 243 |
|
45151 | 244 |
lemma unique_append [rule_format (no_asm)]: "unique l' \<Longrightarrow> unique l \<Longrightarrow> |
245 |
(\<forall>(x,y)\<in>set l. \<forall>(x',y')\<in>set l'. x' \<noteq> x) \<longrightarrow> unique (l @ l')" |
|
246 |
by (induct l) (auto dest: fst_in_set_lemma) |
|
12854 | 247 |
|
45151 | 248 |
lemma unique_map_inj: "unique l \<Longrightarrow> inj f \<Longrightarrow> unique (map (\<lambda>(k,x). (f k, g k x)) l)" |
249 |
by (induct l) (auto dest: fst_in_set_lemma simp add: inj_eq) |
|
250 |
||
251 |
lemma map_of_SomeI: "unique l \<Longrightarrow> (k, x) : set l \<Longrightarrow> map_of l k = Some x" |
|
252 |
by (induct l) auto |
|
12854 | 253 |
|
254 |
||
255 |
section "list patterns" |
|
256 |
||
45151 | 257 |
definition lsplit :: "[['a, 'a list] \<Rightarrow> 'b, 'a list] \<Rightarrow> 'b" |
258 |
where "lsplit = (\<lambda>f l. f (hd l) (tl l))" |
|
37956 | 259 |
|
260 |
text {* list patterns -- extends pre-defined type "pttrn" used in abstractions *} |
|
12854 | 261 |
syntax |
45151 | 262 |
"_lpttrn" :: "[pttrn, pttrn] \<Rightarrow> pttrn" ("_#/_" [901,900] 900) |
12854 | 263 |
translations |
45151 | 264 |
"\<lambda>y # x # xs. b" \<rightleftharpoons> "CONST lsplit (\<lambda>y x # xs. b)" |
265 |
"\<lambda>x # xs. b" \<rightleftharpoons> "CONST lsplit (\<lambda>x xs. b)" |
|
12854 | 266 |
|
267 |
lemma lsplit [simp]: "lsplit c (x#xs) = c x xs" |
|
45151 | 268 |
by (simp add: lsplit_def) |
12854 | 269 |
|
270 |
lemma lsplit2 [simp]: "lsplit P (x#xs) y z = P x xs y z" |
|
45151 | 271 |
by (simp add: lsplit_def) |
12854 | 272 |
|
273 |
end |