| author | haftmann | 
| Mon, 12 Jul 2010 08:58:27 +0200 | |
| changeset 37766 | a779f463bae4 | 
| parent 36176 | 3fe7e97ccca8 | 
| child 37956 | ee939247b2fb | 
| 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  | 
||
| 16417 | 6  | 
theory Basis imports Main begin  | 
| 12854 | 7  | 
|
8  | 
||
9  | 
section "misc"  | 
|
10  | 
||
11  | 
declare same_fstI [intro!] (*### TO HOL/Wellfounded_Relations *)  | 
|
12  | 
||
13  | 
declare split_if_asm [split] option.split [split] option.split_asm [split]  | 
|
| 24019 | 14  | 
declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
 | 
| 12854 | 15  | 
declare if_weak_cong [cong del] option.weak_case_cong [cong del]  | 
| 18447 | 16  | 
declare length_Suc_conv [iff]  | 
17  | 
||
| 12854 | 18  | 
lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}"
 | 
19  | 
apply auto  | 
|
20  | 
done  | 
|
21  | 
||
22  | 
lemma subset_insertD:  | 
|
23  | 
"A <= insert x B ==> A <= B & x ~: A | (EX B'. A = insert x B' & B' <= B)"  | 
|
24  | 
apply (case_tac "x:A")  | 
|
25  | 
apply (rule disjI2)  | 
|
26  | 
apply (rule_tac x = "A-{x}" in exI)
 | 
|
27  | 
apply fast+  | 
|
28  | 
done  | 
|
29  | 
||
| 
35067
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
30  | 
abbreviation nat3 :: nat  ("3") where "3 == Suc 2"
 | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
31  | 
abbreviation nat4 :: nat  ("4") where "4 == Suc 3"
 | 
| 12854 | 32  | 
|
33  | 
(*unused*)  | 
|
34  | 
lemma range_bool_domain: "range f = {f True, f False}"
 | 
|
35  | 
apply auto  | 
|
36  | 
apply (case_tac "xa")  | 
|
37  | 
apply auto  | 
|
38  | 
done  | 
|
39  | 
||
| 13867 | 40  | 
(* irrefl_tranclI in Transitive_Closure.thy is more general *)  | 
| 12854 | 41  | 
lemma irrefl_tranclI': "r^-1 Int r^+ = {} ==> !x. (x, x) ~: r^+"
 | 
| 13867 | 42  | 
by(blast elim: tranclE dest: trancl_into_rtrancl)  | 
43  | 
||
| 12854 | 44  | 
|
45  | 
lemma trancl_rtrancl_trancl:  | 
|
46  | 
"\<lbrakk>(x,y)\<in>r^+; (y,z)\<in>r^*\<rbrakk> \<Longrightarrow> (x,z)\<in>r^+"  | 
|
47  | 
by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2)  | 
|
48  | 
||
49  | 
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
 | 
50  | 
"\<lbrakk>(a,b)\<in>r^*; a\<noteq>b\<rbrakk> \<Longrightarrow> (a,b)\<in>r^+"  | 
| 12854 | 51  | 
apply (drule rtranclD)  | 
52  | 
apply auto  | 
|
53  | 
done  | 
|
54  | 
||
55  | 
lemma rtrancl_into_rtrancl2:  | 
|
56  | 
"\<lbrakk> (a, b) \<in> r; (b, c) \<in> r^* \<rbrakk> \<Longrightarrow> (a, c) \<in> r^*"  | 
|
57  | 
by (auto intro: r_into_rtrancl rtrancl_trans)  | 
|
58  | 
||
59  | 
lemma triangle_lemma:  | 
|
| 
32376
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
60  | 
"\<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>  | 
| 
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
61  | 
\<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"  | 
| 12854 | 62  | 
proof -  | 
63  | 
assume unique: "\<And> a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b=c"  | 
|
| 
32376
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
64  | 
assume "(a,x)\<in>r\<^sup>*"  | 
| 
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
65  | 
then show "(a,y)\<in>r\<^sup>* \<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"  | 
| 12854 | 66  | 
proof (induct rule: converse_rtrancl_induct)  | 
| 
32376
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
67  | 
assume "(x,y)\<in>r\<^sup>*"  | 
| 12854 | 68  | 
then show ?thesis  | 
69  | 
by blast  | 
|
70  | 
next  | 
|
71  | 
fix a v  | 
|
72  | 
assume a_v_r: "(a, v) \<in> r" and  | 
|
| 
32376
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
73  | 
v_x_rt: "(v, x) \<in> r\<^sup>*" and  | 
| 
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
74  | 
a_y_rt: "(a, y) \<in> r\<^sup>*" and  | 
| 
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
75  | 
hyp: "(v, y) \<in> r\<^sup>* \<Longrightarrow> (x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"  | 
| 12854 | 76  | 
from a_y_rt  | 
| 
32376
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
77  | 
show "(x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"  | 
| 12854 | 78  | 
proof (cases rule: converse_rtranclE)  | 
79  | 
assume "a=y"  | 
|
| 
32376
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
80  | 
with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32376 
diff
changeset
 | 
81  | 
by (auto intro: r_into_rtrancl rtrancl_trans)  | 
| 12854 | 82  | 
then show ?thesis  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32376 
diff
changeset
 | 
83  | 
by blast  | 
| 12854 | 84  | 
next  | 
85  | 
fix w  | 
|
86  | 
assume a_w_r: "(a, w) \<in> r" and  | 
|
| 
32376
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
87  | 
w_y_rt: "(w, y) \<in> r\<^sup>*"  | 
| 12854 | 88  | 
from a_v_r a_w_r unique  | 
89  | 
have "v=w"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32376 
diff
changeset
 | 
90  | 
by auto  | 
| 12854 | 91  | 
with w_y_rt hyp  | 
92  | 
show ?thesis  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32376 
diff
changeset
 | 
93  | 
by blast  | 
| 12854 | 94  | 
qed  | 
95  | 
qed  | 
|
96  | 
qed  | 
|
97  | 
||
98  | 
||
99  | 
lemma rtrancl_cases [consumes 1, case_names Refl Trancl]:  | 
|
| 
32376
 
66b4946d9483
reverted accidential corruption of superscripts introduced in a508148f7c25
 
krauss 
parents: 
32367 
diff
changeset
 | 
100  | 
"\<lbrakk>(a,b)\<in>r\<^sup>*; a = b \<Longrightarrow> P; (a,b)\<in>r\<^sup>+ \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"  | 
| 12854 | 101  | 
apply (erule rtranclE)  | 
102  | 
apply (auto dest: rtrancl_into_trancl1)  | 
|
103  | 
done  | 
|
104  | 
||
105  | 
(* context (theory "Set") *)  | 
|
106  | 
lemma Ball_weaken:"\<lbrakk>Ball s P;\<And> x. P x\<longrightarrow>Q x\<rbrakk>\<Longrightarrow>Ball s Q"  | 
|
107  | 
by auto  | 
|
108  | 
||
109  | 
(* context (theory "Finite") *)  | 
|
110  | 
lemma finite_SetCompr2: "[| finite (Collect P); !y. P y --> finite (range (f y)) |] ==>  | 
|
111  | 
  finite {f y x |x y. P y}"
 | 
|
112  | 
apply (subgoal_tac "{f y x |x y. P y} = UNION (Collect P) (%y. range (f y))")
 | 
|
113  | 
prefer 2 apply fast  | 
|
114  | 
apply (erule ssubst)  | 
|
115  | 
apply (erule finite_UN_I)  | 
|
116  | 
apply fast  | 
|
117  | 
done  | 
|
118  | 
||
119  | 
||
120  | 
(* ### TO theory "List" *)  | 
|
121  | 
lemma list_all2_trans: "\<forall> a b c. P1 a b \<longrightarrow> P2 b c \<longrightarrow> P3 a c \<Longrightarrow>  | 
|
122  | 
\<forall>xs2 xs3. list_all2 P1 xs1 xs2 \<longrightarrow> list_all2 P2 xs2 xs3 \<longrightarrow> list_all2 P3 xs1 xs3"  | 
|
123  | 
apply (induct_tac "xs1")  | 
|
124  | 
apply simp  | 
|
125  | 
apply (rule allI)  | 
|
126  | 
apply (induct_tac "xs2")  | 
|
127  | 
apply simp  | 
|
128  | 
apply (rule allI)  | 
|
129  | 
apply (induct_tac "xs3")  | 
|
130  | 
apply auto  | 
|
131  | 
done  | 
|
132  | 
||
133  | 
||
134  | 
section "pairs"  | 
|
135  | 
||
136  | 
lemma surjective_pairing5: "p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))),  | 
|
137  | 
snd (snd (snd (snd p))))"  | 
|
138  | 
apply auto  | 
|
139  | 
done  | 
|
140  | 
||
141  | 
lemma fst_splitE [elim!]:  | 
|
142  | 
"[| fst s' = x'; !!x s. [| s' = (x,s); x = x' |] ==> Q |] ==> Q"  | 
|
| 26349 | 143  | 
by (cases s') auto  | 
| 12854 | 144  | 
|
145  | 
lemma fst_in_set_lemma [rule_format (no_asm)]: "(x, y) : set l --> x : fst ` set l"  | 
|
146  | 
apply (induct_tac "l")  | 
|
147  | 
apply auto  | 
|
148  | 
done  | 
|
149  | 
||
150  | 
||
151  | 
section "quantifiers"  | 
|
152  | 
||
153  | 
lemma All_Ex_refl_eq2 [simp]:  | 
|
154  | 
"(!x. (? b. x = f b & Q b) \<longrightarrow> P x) = (!b. Q b --> P (f b))"  | 
|
155  | 
apply auto  | 
|
156  | 
done  | 
|
157  | 
||
158  | 
lemma ex_ex_miniscope1 [simp]:  | 
|
159  | 
"(EX w v. P w v & Q v) = (EX v. (EX w. P w v) & Q v)"  | 
|
160  | 
apply auto  | 
|
161  | 
done  | 
|
162  | 
||
163  | 
lemma ex_miniscope2 [simp]:  | 
|
164  | 
"(EX v. P v & Q & R v) = (Q & (EX v. P v & R v))"  | 
|
165  | 
apply auto  | 
|
166  | 
done  | 
|
167  | 
||
168  | 
lemma ex_reorder31: "(\<exists>z x y. P x y z) = (\<exists>x y z. P x y z)"  | 
|
169  | 
apply auto  | 
|
170  | 
done  | 
|
171  | 
||
172  | 
lemma All_Ex_refl_eq1 [simp]: "(!x. (? b. x = f b) --> P x) = (!b. P (f b))"  | 
|
173  | 
apply auto  | 
|
174  | 
done  | 
|
175  | 
||
176  | 
||
177  | 
section "sums"  | 
|
178  | 
||
| 
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
 | 
179  | 
hide_const In0 In1  | 
| 12854 | 180  | 
|
| 
35067
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
181  | 
notation sum_case (infixr "'(+')"80)  | 
| 12854 | 182  | 
|
183  | 
consts the_Inl :: "'a + 'b \<Rightarrow> 'a"  | 
|
184  | 
the_Inr :: "'a + 'b \<Rightarrow> 'b"  | 
|
185  | 
primrec "the_Inl (Inl a) = a"  | 
|
186  | 
primrec "the_Inr (Inr b) = b"  | 
|
187  | 
||
188  | 
datatype ('a, 'b, 'c) sum3 = In1 'a | In2 'b | In3 'c
 | 
|
189  | 
||
190  | 
consts    the_In1  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'a"
 | 
|
191  | 
          the_In2  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'b"
 | 
|
192  | 
          the_In3  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'c"
 | 
|
193  | 
primrec "the_In1 (In1 a) = a"  | 
|
194  | 
primrec "the_In2 (In2 b) = b"  | 
|
195  | 
primrec "the_In3 (In3 c) = c"  | 
|
196  | 
||
| 
35067
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
197  | 
abbreviation In1l   :: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
 | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
198  | 
where "In1l e == In1 (Inl e)"  | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
199  | 
|
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
200  | 
abbreviation In1r   :: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
 | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
201  | 
where "In1r c == In1 (Inr c)"  | 
| 12854 | 202  | 
|
| 
35067
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
203  | 
abbreviation the_In1l :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'al"
 | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
204  | 
where "the_In1l == the_Inl \<circ> the_In1"  | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
205  | 
|
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
206  | 
abbreviation the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar"
 | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
207  | 
where "the_In1r == 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
 | 
208  | 
|
| 12854 | 209  | 
ML {*
 | 
| 27226 | 210  | 
fun sum3_instantiate ctxt thm = map (fn s =>  | 
| 
32149
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
30235 
diff
changeset
 | 
211  | 
  simplify (simpset_of ctxt delsimps[@{thm not_None_eq}])
 | 
| 27239 | 212  | 
    (read_instantiate ctxt [(("t", 0), "In" ^ s ^ " ?x")] thm)) ["1l","2","3","1r"]
 | 
| 12854 | 213  | 
*}  | 
214  | 
(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *)  | 
|
215  | 
||
216  | 
||
217  | 
section "quantifiers for option type"  | 
|
218  | 
||
219  | 
syntax  | 
|
| 
35355
 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 
wenzelm 
parents: 
35067 
diff
changeset
 | 
220  | 
  "_Oall" :: "[pttrn, 'a option, bool] => bool"   ("(3! _:_:/ _)" [0,0,10] 10)
 | 
| 
 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 
wenzelm 
parents: 
35067 
diff
changeset
 | 
221  | 
  "_Oex"  :: "[pttrn, 'a option, bool] => bool"   ("(3? _:_:/ _)" [0,0,10] 10)
 | 
| 12854 | 222  | 
|
223  | 
syntax (symbols)  | 
|
| 
35355
 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 
wenzelm 
parents: 
35067 
diff
changeset
 | 
224  | 
  "_Oall" :: "[pttrn, 'a option, bool] => bool"   ("(3\<forall>_\<in>_:/ _)"  [0,0,10] 10)
 | 
| 
 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 
wenzelm 
parents: 
35067 
diff
changeset
 | 
225  | 
  "_Oex"  :: "[pttrn, 'a option, bool] => bool"   ("(3\<exists>_\<in>_:/ _)"  [0,0,10] 10)
 | 
| 12854 | 226  | 
|
227  | 
translations  | 
|
| 
30235
 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 
nipkow 
parents: 
27239 
diff
changeset
 | 
228  | 
"! x:A: P" == "! x:CONST Option.set A. P"  | 
| 
 
58d147683393
Made Option a separate theory and renamed option_map to Option.map
 
nipkow 
parents: 
27239 
diff
changeset
 | 
229  | 
"? x:A: P" == "? x:CONST Option.set A. P"  | 
| 12854 | 230  | 
|
| 19323 | 231  | 
section "Special map update"  | 
232  | 
||
233  | 
text{* Deemed too special for theory Map. *}
 | 
|
234  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35067 
diff
changeset
 | 
235  | 
definition chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)" where
 | 
| 19323 | 236  | 
"chg_map f a m == case m a of None => m | Some b => m(a|->f b)"  | 
237  | 
||
238  | 
lemma chg_map_new[simp]: "m a = None ==> chg_map f a m = m"  | 
|
239  | 
by (unfold chg_map_def, auto)  | 
|
240  | 
||
241  | 
lemma chg_map_upd[simp]: "m a = Some b ==> chg_map f a m = m(a|->f b)"  | 
|
242  | 
by (unfold chg_map_def, auto)  | 
|
243  | 
||
244  | 
lemma chg_map_other [simp]: "a \<noteq> b \<Longrightarrow> chg_map f a m b = m b"  | 
|
245  | 
by (auto simp: chg_map_def split add: option.split)  | 
|
246  | 
||
| 12854 | 247  | 
|
248  | 
section "unique association lists"  | 
|
249  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35067 
diff
changeset
 | 
250  | 
definition unique :: "('a \<times> 'b) list \<Rightarrow> bool" where
 | 
| 12893 | 251  | 
"unique \<equiv> distinct \<circ> map fst"  | 
| 12854 | 252  | 
|
253  | 
lemma uniqueD [rule_format (no_asm)]:  | 
|
254  | 
"unique l--> (!x y. (x,y):set l --> (!x' y'. (x',y'):set l --> x=x'--> y=y'))"  | 
|
255  | 
apply (unfold unique_def o_def)  | 
|
256  | 
apply (induct_tac "l")  | 
|
257  | 
apply (auto dest: fst_in_set_lemma)  | 
|
258  | 
done  | 
|
259  | 
||
260  | 
lemma unique_Nil [simp]: "unique []"  | 
|
261  | 
apply (unfold unique_def)  | 
|
262  | 
apply (simp (no_asm))  | 
|
263  | 
done  | 
|
264  | 
||
265  | 
lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))"  | 
|
266  | 
apply (unfold unique_def)  | 
|
267  | 
apply (auto dest: fst_in_set_lemma)  | 
|
268  | 
done  | 
|
269  | 
||
270  | 
lemmas unique_ConsI = conjI [THEN unique_Cons [THEN iffD2], standard]  | 
|
271  | 
||
272  | 
lemma unique_single [simp]: "!!p. unique [p]"  | 
|
273  | 
apply auto  | 
|
274  | 
done  | 
|
275  | 
||
276  | 
lemma unique_ConsD: "unique (x#xs) ==> unique xs"  | 
|
277  | 
apply (simp add: unique_def)  | 
|
278  | 
done  | 
|
279  | 
||
280  | 
lemma unique_append [rule_format (no_asm)]: "unique l' ==> unique l -->  | 
|
281  | 
(!(x,y):set l. !(x',y'):set l'. x' ~= x) --> unique (l @ l')"  | 
|
282  | 
apply (induct_tac "l")  | 
|
283  | 
apply (auto dest: fst_in_set_lemma)  | 
|
284  | 
done  | 
|
285  | 
||
286  | 
lemma unique_map_inj [rule_format (no_asm)]: "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)"  | 
|
287  | 
apply (induct_tac "l")  | 
|
288  | 
apply (auto dest: fst_in_set_lemma simp add: inj_eq)  | 
|
289  | 
done  | 
|
290  | 
||
291  | 
lemma map_of_SomeI [rule_format (no_asm)]: "unique l --> (k, x) : set l --> map_of l k = Some x"  | 
|
292  | 
apply (induct_tac "l")  | 
|
293  | 
apply auto  | 
|
294  | 
done  | 
|
295  | 
||
296  | 
||
297  | 
section "list patterns"  | 
|
298  | 
||
299  | 
consts  | 
|
300  | 
lsplit :: "[['a, 'a list] => 'b, 'a list] => 'b"  | 
|
301  | 
defs  | 
|
302  | 
lsplit_def: "lsplit == %f l. f (hd l) (tl l)"  | 
|
303  | 
(* list patterns -- extends pre-defined type "pttrn" used in abstractions *)  | 
|
304  | 
syntax  | 
|
305  | 
  "_lpttrn"    :: "[pttrn,pttrn] => pttrn"     ("_#/_" [901,900] 900)
 | 
|
306  | 
translations  | 
|
| 
35067
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
307  | 
"%y#x#xs. b" == "CONST lsplit (%y x#xs. b)"  | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
34915 
diff
changeset
 | 
308  | 
"%x#xs . b" == "CONST lsplit (%x xs . b)"  | 
| 12854 | 309  | 
|
310  | 
lemma lsplit [simp]: "lsplit c (x#xs) = c x xs"  | 
|
311  | 
apply (unfold lsplit_def)  | 
|
312  | 
apply (simp (no_asm))  | 
|
313  | 
done  | 
|
314  | 
||
315  | 
lemma lsplit2 [simp]: "lsplit P (x#xs) y z = P x xs y z"  | 
|
316  | 
apply (unfold lsplit_def)  | 
|
317  | 
apply simp  | 
|
318  | 
done  | 
|
319  | 
||
320  | 
||
321  | 
end  |