5 *) |
5 *) |
6 |
6 |
7 val strip_tac1 = SELECT_GOAL (safe_tac (HOL_cs delrules [conjI, disjE, impCE])); |
7 val strip_tac1 = SELECT_GOAL (safe_tac (HOL_cs delrules [conjI, disjE, impCE])); |
8 |
8 |
9 Goalw [image_def] |
9 Goalw [image_def] |
10 "x \\<in> f``A \\<Longrightarrow> \\<exists>y. y \\<in> A \\<and> x = f y"; |
10 "x \\<in> f``A ==> \\<exists>y. y \\<in> A \\<and> x = f y"; |
11 by(Auto_tac); |
11 by(Auto_tac); |
12 qed "image_rev"; |
12 qed "image_rev"; |
13 |
13 |
14 fun case_tac1 s i = EVERY [case_tac s i, rotate_tac ~1 i, rotate_tac ~1 (i+1)]; |
14 fun case_tac1 s i = EVERY [case_tac s i, rotate_tac ~1 i, rotate_tac ~1 (i+1)]; |
15 |
15 |
16 val select_split = prove_goalw Prod.thy [split_def] |
16 val select_split = prove_goalw Prod.thy [split_def] |
17 "(\\<epsilon>(x,y). P x y) = (\\<epsilon>xy. P (fst xy) (snd xy))" (K [rtac refl 1]); |
17 "(SOME (x,y). P x y) = (SOME xy. P (fst xy) (snd xy))" (K [rtac refl 1]); |
18 |
18 |
19 |
19 |
20 val split_beta = prove_goal Prod.thy "(\\<lambda>(x,y). P x y) z = P (fst z) (snd z)" |
20 val split_beta = prove_goal Prod.thy "(\\<lambda>(x,y). P x y) z = P (fst z) (snd z)" |
21 (fn _ => [stac surjective_pairing 1, stac split 1, rtac refl 1]); |
21 (fn _ => [stac surjective_pairing 1, stac split 1, rtac refl 1]); |
22 val split_beta2 = prove_goal Prod.thy "(\\<lambda>(x,y). P x y) (w,z) = P w z" |
22 val split_beta2 = prove_goal Prod.thy "(\\<lambda>(x,y). P x y) (w,z) = P w z" |
23 (fn _ => [Auto_tac]); |
23 (fn _ => [Auto_tac]); |
24 val splitE2 = prove_goal Prod.thy "\\<lbrakk>Q (split P z); \\<And>x y. \\<lbrakk>z = (x, y); Q (P x y)\\<rbrakk> \\<Longrightarrow> R\\<rbrakk> \\<Longrightarrow> R" (fn prems => [ |
24 val splitE2 = prove_goal Prod.thy "[|Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R|] ==> R" (fn prems => [ |
25 REPEAT (resolve_tac (prems@[surjective_pairing]) 1), |
25 REPEAT (resolve_tac (prems@[surjective_pairing]) 1), |
26 rtac (split_beta RS subst) 1, |
26 rtac (split_beta RS subst) 1, |
27 rtac (hd prems) 1]); |
27 rtac (hd prems) 1]); |
28 val splitE2' = prove_goal Prod.thy "\\<lbrakk>((\\<lambda>(x,y). P x y) z) w; \\<And>x y. \\<lbrakk>z = (x, y); (P x y) w\\<rbrakk> \\<Longrightarrow> R\\<rbrakk> \\<Longrightarrow> R" (fn prems => [ |
28 val splitE2' = prove_goal Prod.thy "[|((\\<lambda>(x,y). P x y) z) w; !!x y. [|z = (x, y); (P x y) w|] ==> R|] ==> R" (fn prems => [ |
29 REPEAT (resolve_tac (prems@[surjective_pairing]) 1), |
29 REPEAT (resolve_tac (prems@[surjective_pairing]) 1), |
30 res_inst_tac [("P1","P")] (split_beta RS subst) 1, |
30 res_inst_tac [("P1","P")] (split_beta RS subst) 1, |
31 rtac (hd prems) 1]); |
31 rtac (hd prems) 1]); |
32 |
32 |
33 |
33 |
34 fun pair_tac s = res_inst_tac [("p",s)] PairE THEN' hyp_subst_tac; |
34 fun pair_tac s = res_inst_tac [("p",s)] PairE THEN' hyp_subst_tac; |
35 |
35 |
36 val BallE = prove_goal thy "\\<lbrakk>Ball A P; x \\<notin> A \\<Longrightarrow> Q; P x \\<Longrightarrow> Q \\<rbrakk> \\<Longrightarrow> Q" |
36 val BallE = prove_goal thy "[|Ball A P; x \\<notin> A ==> Q; P x ==> Q |] ==> Q" |
37 (fn prems => [rtac ballE 1, resolve_tac prems 1, |
37 (fn prems => [rtac ballE 1, resolve_tac prems 1, |
38 eresolve_tac prems 1, eresolve_tac prems 1]); |
38 eresolve_tac prems 1, eresolve_tac prems 1]); |
39 |
39 |
40 val set_cs2 = set_cs delrules [ballE] addSEs [BallE]; |
40 val set_cs2 = set_cs delrules [ballE] addSEs [BallE]; |
41 |
41 |
42 Addsimps [Let_def]; |
42 Addsimps [Let_def]; |
43 |
43 |
44 (* To HOL.ML *) |
44 (* To HOL.ML *) |
45 |
45 |
46 val ex1_some_eq_trivial = prove_goal HOL.thy "\\<lbrakk> \\<exists>!x. P x; P y \\<rbrakk> \\<Longrightarrow> Eps P = y" |
46 val ex1_some_eq_trivial = prove_goal HOL.thy "[| \\<exists>!x. P x; P y |] ==> Eps P = y" |
47 (fn prems => [ |
47 (fn prems => [ |
48 cut_facts_tac prems 1, |
48 cut_facts_tac prems 1, |
49 rtac some_equality 1, |
49 rtac some_equality 1, |
50 atac 1, |
50 atac 1, |
51 etac ex1E 1, |
51 etac ex1E 1, |
118 Goal "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)"; |
118 Goal "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)"; |
119 by (induct_tac "l" 1); |
119 by (induct_tac "l" 1); |
120 by (auto_tac (claset() addDs [fst_in_set_lemma],simpset()addsimps[inj_eq])); |
120 by (auto_tac (claset() addDs [fst_in_set_lemma],simpset()addsimps[inj_eq])); |
121 qed_spec_mp "unique_map_inj"; |
121 qed_spec_mp "unique_map_inj"; |
122 |
122 |
123 Goal "\\<And>l. unique l \\<Longrightarrow> unique (map (split (\\<lambda>k. Pair (k, C))) l)"; |
123 Goal "!!l. unique l ==> unique (map (split (\\<lambda>k. Pair (k, C))) l)"; |
124 by(etac unique_map_inj 1); |
124 by(etac unique_map_inj 1); |
125 by(rtac injI 1); |
125 by(rtac injI 1); |
126 by Auto_tac; |
126 by Auto_tac; |
127 qed "unique_map_Pair"; |
127 qed "unique_map_Pair"; |
128 |
128 |
129 Goal "\\<lbrakk>M = N; \\<And>x. x\\<in>N \\<Longrightarrow> f x = g x\\<rbrakk> \\<Longrightarrow> f``M = g``N"; |
129 Goal "[|M = N; !!x. x\\<in>N ==> f x = g x|] ==> f``M = g``N"; |
130 by(rtac set_ext 1); |
130 by(rtac set_ext 1); |
131 by(simp_tac (simpset() addsimps image_def::premises()) 1); |
131 by(simp_tac (simpset() addsimps image_def::premises()) 1); |
132 qed "image_cong"; |
132 qed "image_cong"; |
133 |
133 |
134 val split_Pair_eq = prove_goal Prod.thy |
134 val split_Pair_eq = prove_goal Prod.thy |
135 "\\<And>X. ((x, y), z) \\<in> split (\\<lambda>x. Pair (x, Y)) `` A \\<Longrightarrow> y = Y" (K [ |
135 "!!X. ((x, y), z) \\<in> split (\\<lambda>x. Pair (x, Y)) `` A ==> y = Y" (K [ |
136 etac imageE 1, |
136 etac imageE 1, |
137 split_all_tac 1, |
137 split_all_tac 1, |
138 auto_tac(claset_of Prod.thy,simpset_of Prod.thy)]); |
138 auto_tac(claset_of Prod.thy,simpset_of Prod.thy)]); |
139 |
139 |
140 |
140 |
141 (* More about Maps *) |
141 (* More about Maps *) |
142 |
142 |
143 val override_SomeD = prove_goalw thy [override_def] "(s \\<oplus> t) k = Some x \\<Longrightarrow> \ |
143 val override_SomeD = prove_goalw thy [override_def] "(s \\<oplus> t) k = Some x ==> \ |
144 \ t k = Some x | t k = None \\<and> s k = Some x" (fn prems => [ |
144 \ t k = Some x | t k = None \\<and> s k = Some x" (fn prems => [ |
145 cut_facts_tac prems 1, |
145 cut_facts_tac prems 1, |
146 case_tac "\\<exists>x. t k = Some x" 1, |
146 case_tac "\\<exists>x. t k = Some x" 1, |
147 etac exE 1, |
147 etac exE 1, |
148 rotate_tac ~1 1, |
148 rotate_tac ~1 1, |
151 rotate_tac ~1 1, |
151 rotate_tac ~1 1, |
152 Asm_full_simp_tac 1]); |
152 Asm_full_simp_tac 1]); |
153 |
153 |
154 Addsimps [fun_upd_same, fun_upd_other]; |
154 Addsimps [fun_upd_same, fun_upd_other]; |
155 |
155 |
156 Goal "unique xys \\<longrightarrow> (map_of xys x = Some y) = ((x,y) \\<in> set xys)"; |
156 Goal "unique xys --> (map_of xys x = Some y) = ((x,y) \\<in> set xys)"; |
157 by(induct_tac "xys" 1); |
157 by(induct_tac "xys" 1); |
158 by(Simp_tac 1); |
158 by(Simp_tac 1); |
159 by(force_tac (claset(), simpset() addsimps [unique_Cons]) 1); |
159 by(force_tac (claset(), simpset() addsimps [unique_Cons]) 1); |
160 qed_spec_mp "unique_map_of_Some_conv"; |
160 qed_spec_mp "unique_map_of_Some_conv"; |
161 |
161 |
162 val in_set_get = unique_map_of_Some_conv RS iffD2; |
162 val in_set_get = unique_map_of_Some_conv RS iffD2; |
163 val get_in_set = unique_map_of_Some_conv RS iffD1; |
163 val get_in_set = unique_map_of_Some_conv RS iffD1; |
164 |
164 |
165 Goal "(\\<forall>(x,y)\\<in>set l. P x y) \\<longrightarrow> (\\<forall>x. \\<forall>y. map_of l x = Some y \\<longrightarrow> P x y)"; |
165 Goal "(\\<forall>(x,y)\\<in>set l. P x y) --> (\\<forall>x. \\<forall>y. map_of l x = Some y --> P x y)"; |
166 by(induct_tac "l" 1); |
166 by(induct_tac "l" 1); |
167 by(ALLGOALS Simp_tac); |
167 by(ALLGOALS Simp_tac); |
168 by Safe_tac; |
168 by Safe_tac; |
169 by Auto_tac; |
169 by Auto_tac; |
170 bind_thm("Ball_set_table",result() RS mp); |
170 bind_thm("Ball_set_table",result() RS mp); |
171 |
171 |
172 val table_mono = prove_goal thy |
172 val table_mono = prove_goal thy |
173 "unique l' \\<longrightarrow> (\\<forall>xy. (xy)\\<in>set l \\<longrightarrow> (xy)\\<in>set l') \\<longrightarrow>\ |
173 "unique l' --> (\\<forall>xy. (xy)\\<in>set l --> (xy)\\<in>set l') -->\ |
174 \ (\\<forall>k y. map_of l k = Some y \\<longrightarrow> map_of l' k = Some y)" (fn _ => [ |
174 \ (\\<forall>k y. map_of l k = Some y --> map_of l' k = Some y)" (fn _ => [ |
175 induct_tac "l" 1, |
175 induct_tac "l" 1, |
176 Auto_tac, |
176 Auto_tac, |
177 fast_tac (HOL_cs addSIs [in_set_get]) 1]) |
177 fast_tac (HOL_cs addSIs [in_set_get]) 1]) |
178 RS mp RS mp RS spec RS spec RS mp; |
178 RS mp RS mp RS spec RS spec RS mp; |
179 |
179 |
180 val table_map_Some' = prove_goal thy "map_of t k = Some (k', x) \\<longrightarrow> \ |
180 val table_map_Some' = prove_goal thy "map_of t k = Some (k', x) --> \ |
181 \ map_of (map (\\<lambda>u. ((fst u, fst (snd u)), snd (snd u))) t) (k, k') = Some x" (K [ |
181 \ map_of (map (\\<lambda>u. ((fst u, fst (snd u)), snd (snd u))) t) (k, k') = Some x" (K [ |
182 induct_tac "t" 1, |
182 induct_tac "t" 1, |
183 ALLGOALS Simp_tac, |
183 ALLGOALS Simp_tac, |
184 case_tac1 "k = fst a" 1, |
184 case_tac1 "k = fst a" 1, |
185 Auto_tac]) RS mp; |
185 Auto_tac]) RS mp; |
186 val table_map_Some = prove_goal thy |
186 val table_map_Some = prove_goal thy |
187 "map_of (map (\\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) \\<longrightarrow> \ |
187 "map_of (map (\\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) --> \ |
188 \map_of t (k, k') = Some x" (K [ |
188 \map_of t (k, k') = Some x" (K [ |
189 induct_tac "t" 1, |
189 induct_tac "t" 1, |
190 Auto_tac]) RS mp; |
190 Auto_tac]) RS mp; |
191 |
191 |
192 |
192 |
193 val table_mapf_Some = prove_goal thy "\\<And>f. \\<forall>x y. f x = f y \\<longrightarrow> x = y \\<Longrightarrow> \ |
193 val table_mapf_Some = prove_goal thy "!!f. \\<forall>x y. f x = f y --> x = y ==> \ |
194 \ map_of (map (\\<lambda>(k,x). (k,f x)) t) k = Some (f x) \\<longrightarrow> map_of t k = Some x" (K [ |
194 \ map_of (map (\\<lambda>(k,x). (k,f x)) t) k = Some (f x) --> map_of t k = Some x" (K [ |
195 induct_tac "t" 1, |
195 induct_tac "t" 1, |
196 Auto_tac]) RS mp; |
196 Auto_tac]) RS mp; |
197 val table_mapf_SomeD = prove_goal thy |
197 val table_mapf_SomeD = prove_goal thy |
198 "map_of (map (\\<lambda>(k,x). (k, f x)) t) k = Some z \\<longrightarrow> (\\<exists>y. (k,y)\\<in>set t \\<and> z = f y)"(K [ |
198 "map_of (map (\\<lambda>(k,x). (k, f x)) t) k = Some z --> (\\<exists>y. (k,y)\\<in>set t \\<and> z = f y)"(K [ |
199 induct_tac "t" 1, |
199 induct_tac "t" 1, |
200 Auto_tac]) RS mp; |
200 Auto_tac]) RS mp; |
201 |
201 |
202 val table_mapf_Some2 = prove_goal thy |
202 val table_mapf_Some2 = prove_goal thy |
203 "\\<And>k. map_of (map (\\<lambda>(k,x). (k,C,x)) t) k = Some (D,x) \\<Longrightarrow> C = D \\<and> map_of t k = Some x" (K [ |
203 "!!k. map_of (map (\\<lambda>(k,x). (k,C,x)) t) k = Some (D,x) ==> C = D \\<and> map_of t k = Some x" (K [ |
204 forward_tac [table_mapf_SomeD] 1, |
204 forward_tac [table_mapf_SomeD] 1, |
205 Auto_tac, |
205 Auto_tac, |
206 rtac table_mapf_Some 1, |
206 rtac table_mapf_Some 1, |
207 atac 2, |
207 atac 2, |
208 Fast_tac 1]); |
208 Fast_tac 1]); |