163
|
1 |
theory FSet
|
|
2 |
imports QuotMain
|
|
3 |
begin
|
|
4 |
|
165
|
5 |
(* test for get_fun *)
|
|
6 |
datatype t =
|
|
7 |
vr "string"
|
|
8 |
| ap "t list"
|
|
9 |
| lm "string" "t"
|
|
10 |
|
|
11 |
consts Rt :: "t \<Rightarrow> t \<Rightarrow> bool"
|
|
12 |
axioms t_eq: "EQUIV Rt"
|
|
13 |
|
|
14 |
quotient qt = "t" / "Rt"
|
|
15 |
by (rule t_eq)
|
|
16 |
|
|
17 |
setup {*
|
|
18 |
maps_update @{type_name "list"} {mapfun = @{const_name "map"}, relfun = @{const_name "LIST_REL"}} #>
|
|
19 |
maps_update @{type_name "*"} {mapfun = @{const_name "prod_fun"}, relfun = @{const_name "prod_rel"}} #>
|
|
20 |
maps_update @{type_name "fun"} {mapfun = @{const_name "fun_map"}, relfun = @{const_name "FUN_REL"}}
|
|
21 |
*}
|
|
22 |
|
|
23 |
|
|
24 |
ML {*
|
|
25 |
get_fun repF @{typ t} @{typ qt} @{context} @{typ "((((qt \<Rightarrow> qt) \<Rightarrow> qt) \<Rightarrow> qt) list) * nat"}
|
|
26 |
|> fst
|
|
27 |
|> Syntax.string_of_term @{context}
|
|
28 |
|> writeln
|
|
29 |
*}
|
|
30 |
|
|
31 |
ML {*
|
|
32 |
get_fun absF @{typ t} @{typ qt} @{context} @{typ "qt * nat"}
|
|
33 |
|> fst
|
|
34 |
|> Syntax.string_of_term @{context}
|
|
35 |
|> writeln
|
|
36 |
*}
|
|
37 |
|
|
38 |
ML {*
|
|
39 |
get_fun absF @{typ t} @{typ qt} @{context} @{typ "(qt \<Rightarrow> qt) \<Rightarrow> qt"}
|
|
40 |
|> fst
|
|
41 |
|> Syntax.pretty_term @{context}
|
|
42 |
|> Pretty.string_of
|
|
43 |
|> writeln
|
|
44 |
*}
|
|
45 |
(* end test get_fun *)
|
|
46 |
|
|
47 |
|
163
|
48 |
inductive
|
|
49 |
list_eq (infix "\<approx>" 50)
|
|
50 |
where
|
|
51 |
"a#b#xs \<approx> b#a#xs"
|
|
52 |
| "[] \<approx> []"
|
|
53 |
| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
|
|
54 |
| "a#a#xs \<approx> a#xs"
|
|
55 |
| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
|
|
56 |
| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
|
|
57 |
|
|
58 |
lemma list_eq_refl:
|
|
59 |
shows "xs \<approx> xs"
|
|
60 |
apply (induct xs)
|
|
61 |
apply (auto intro: list_eq.intros)
|
|
62 |
done
|
|
63 |
|
|
64 |
lemma equiv_list_eq:
|
|
65 |
shows "EQUIV list_eq"
|
|
66 |
unfolding EQUIV_REFL_SYM_TRANS REFL_def SYM_def TRANS_def
|
|
67 |
apply(auto intro: list_eq.intros list_eq_refl)
|
|
68 |
done
|
|
69 |
|
|
70 |
quotient fset = "'a list" / "list_eq"
|
|
71 |
apply(rule equiv_list_eq)
|
|
72 |
done
|
|
73 |
|
|
74 |
print_theorems
|
|
75 |
|
|
76 |
typ "'a fset"
|
|
77 |
thm "Rep_fset"
|
|
78 |
thm "ABS_fset_def"
|
|
79 |
|
|
80 |
ML {* @{term "Abs_fset"} *}
|
|
81 |
local_setup {*
|
|
82 |
make_const_def @{binding EMPTY} @{term "[]"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
|
|
83 |
*}
|
|
84 |
|
|
85 |
term Nil
|
|
86 |
term EMPTY
|
|
87 |
thm EMPTY_def
|
|
88 |
|
|
89 |
|
|
90 |
local_setup {*
|
|
91 |
make_const_def @{binding INSERT} @{term "op #"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
|
|
92 |
*}
|
|
93 |
|
|
94 |
term Cons
|
|
95 |
term INSERT
|
|
96 |
thm INSERT_def
|
|
97 |
|
|
98 |
local_setup {*
|
|
99 |
make_const_def @{binding UNION} @{term "op @"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
|
|
100 |
*}
|
|
101 |
|
|
102 |
term append
|
|
103 |
term UNION
|
|
104 |
thm UNION_def
|
|
105 |
|
|
106 |
|
|
107 |
thm QUOTIENT_fset
|
|
108 |
|
|
109 |
thm QUOT_TYPE_I_fset.thm11
|
|
110 |
|
|
111 |
|
|
112 |
fun
|
|
113 |
membship :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" (infix "memb" 100)
|
|
114 |
where
|
|
115 |
m1: "(x memb []) = False"
|
|
116 |
| m2: "(x memb (y#xs)) = ((x=y) \<or> (x memb xs))"
|
|
117 |
|
|
118 |
fun
|
|
119 |
card1 :: "'a list \<Rightarrow> nat"
|
|
120 |
where
|
|
121 |
card1_nil: "(card1 []) = 0"
|
|
122 |
| card1_cons: "(card1 (x # xs)) = (if (x memb xs) then (card1 xs) else (Suc (card1 xs)))"
|
|
123 |
|
|
124 |
local_setup {*
|
|
125 |
make_const_def @{binding card} @{term "card1"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
|
|
126 |
*}
|
|
127 |
|
|
128 |
term card1
|
|
129 |
term card
|
|
130 |
thm card_def
|
|
131 |
|
|
132 |
(* text {*
|
|
133 |
Maybe make_const_def should require a theorem that says that the particular lifted function
|
|
134 |
respects the relation. With it such a definition would be impossible:
|
|
135 |
make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
|
|
136 |
*}*)
|
|
137 |
|
|
138 |
lemma card1_0:
|
|
139 |
fixes a :: "'a list"
|
|
140 |
shows "(card1 a = 0) = (a = [])"
|
|
141 |
apply (induct a)
|
|
142 |
apply (simp)
|
|
143 |
apply (simp_all)
|
|
144 |
apply meson
|
|
145 |
apply (simp_all)
|
|
146 |
done
|
|
147 |
|
|
148 |
lemma not_mem_card1:
|
|
149 |
fixes x :: "'a"
|
|
150 |
fixes xs :: "'a list"
|
|
151 |
shows "~(x memb xs) \<Longrightarrow> card1 (x # xs) = Suc (card1 xs)"
|
|
152 |
by simp
|
|
153 |
|
|
154 |
|
|
155 |
lemma mem_cons:
|
|
156 |
fixes x :: "'a"
|
|
157 |
fixes xs :: "'a list"
|
|
158 |
assumes a : "x memb xs"
|
|
159 |
shows "x # xs \<approx> xs"
|
|
160 |
using a
|
|
161 |
apply (induct xs)
|
|
162 |
apply (auto intro: list_eq.intros)
|
|
163 |
done
|
|
164 |
|
|
165 |
lemma card1_suc:
|
|
166 |
fixes xs :: "'a list"
|
|
167 |
fixes n :: "nat"
|
|
168 |
assumes c: "card1 xs = Suc n"
|
|
169 |
shows "\<exists>a ys. ~(a memb ys) \<and> xs \<approx> (a # ys)"
|
|
170 |
using c
|
|
171 |
apply(induct xs)
|
|
172 |
apply (metis Suc_neq_Zero card1_0)
|
|
173 |
apply (metis QUOT_TYPE_I_fset.R_trans card1_cons list_eq_refl mem_cons)
|
|
174 |
done
|
|
175 |
|
|
176 |
primrec
|
|
177 |
fold1
|
|
178 |
where
|
|
179 |
"fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
|
|
180 |
| "fold1 f g z (a # A) =
|
|
181 |
(if ((!u v. (f u v = f v u))
|
|
182 |
\<and> (!u v w. ((f u (f v w) = f (f u v) w))))
|
|
183 |
then (
|
|
184 |
if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
|
|
185 |
) else z)"
|
|
186 |
|
|
187 |
(* fold1_def is not usable, but: *)
|
|
188 |
thm fold1.simps
|
|
189 |
|
|
190 |
lemma fs1_strong_cases:
|
|
191 |
fixes X :: "'a list"
|
|
192 |
shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a memb Y) \<and> (X \<approx> a # Y)))"
|
|
193 |
apply (induct X)
|
|
194 |
apply (simp)
|
|
195 |
apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons m1)
|
|
196 |
done
|
|
197 |
|
|
198 |
local_setup {*
|
|
199 |
make_const_def @{binding IN} @{term "membship"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
|
|
200 |
*}
|
|
201 |
|
|
202 |
term membship
|
|
203 |
term IN
|
|
204 |
thm IN_def
|
|
205 |
|
|
206 |
ML {*
|
|
207 |
val consts = [@{const_name "Nil"}, @{const_name "Cons"},
|
|
208 |
@{const_name "membship"}, @{const_name "card1"},
|
|
209 |
@{const_name "append"}, @{const_name "fold1"}];
|
|
210 |
*}
|
|
211 |
|
|
212 |
ML {* val fset_defs = @{thms EMPTY_def IN_def UNION_def card_def INSERT_def} *}
|
|
213 |
ML {* val fset_defs_sym = map (fn t => symmetric (unabs_def @{context} t)) fset_defs *}
|
|
214 |
|
|
215 |
thm fun_map.simps
|
|
216 |
text {* Respectfullness *}
|
|
217 |
|
164
|
218 |
lemma memb_rsp:
|
163
|
219 |
fixes z
|
|
220 |
assumes a: "list_eq x y"
|
|
221 |
shows "(z memb x) = (z memb y)"
|
|
222 |
using a by induct auto
|
|
223 |
|
164
|
224 |
lemma ho_memb_rsp:
|
|
225 |
"(op = ===> (op \<approx> ===> op =)) (op memb) (op memb)"
|
|
226 |
apply (simp add: FUN_REL.simps)
|
|
227 |
apply (metis memb_rsp)
|
|
228 |
done
|
|
229 |
|
163
|
230 |
lemma card1_rsp:
|
|
231 |
fixes a b :: "'a list"
|
|
232 |
assumes e: "a \<approx> b"
|
|
233 |
shows "card1 a = card1 b"
|
|
234 |
using e apply induct
|
164
|
235 |
apply (simp_all add:memb_rsp)
|
163
|
236 |
done
|
|
237 |
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
238 |
lemma ho_card1_rsp: "op \<approx> ===> op = card1 card1"
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
239 |
apply (simp add: FUN_REL.simps)
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
240 |
apply (metis card1_rsp)
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
241 |
done
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
242 |
|
164
|
243 |
lemma cons_rsp:
|
163
|
244 |
fixes z
|
|
245 |
assumes a: "xs \<approx> ys"
|
|
246 |
shows "(z # xs) \<approx> (z # ys)"
|
|
247 |
using a by (rule list_eq.intros(5))
|
|
248 |
|
164
|
249 |
lemma ho_cons_rsp:
|
|
250 |
"op = ===> op \<approx> ===> op \<approx> op # op #"
|
|
251 |
apply (simp add: FUN_REL.simps)
|
|
252 |
apply (metis cons_rsp)
|
|
253 |
done
|
|
254 |
|
163
|
255 |
lemma append_respects_fst:
|
|
256 |
assumes a : "list_eq l1 l2"
|
|
257 |
shows "list_eq (l1 @ s) (l2 @ s)"
|
|
258 |
using a
|
|
259 |
apply(induct)
|
|
260 |
apply(auto intro: list_eq.intros)
|
|
261 |
apply(simp add: list_eq_refl)
|
|
262 |
done
|
|
263 |
|
|
264 |
thm list.induct
|
|
265 |
lemma list_induct_hol4:
|
|
266 |
fixes P :: "'a list \<Rightarrow> bool"
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
267 |
assumes a: "((P []) \<and> (\<forall>t. (P t) \<longrightarrow> (\<forall>h. (P (h # t)))))"
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
268 |
shows "\<forall>l. (P l)"
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
269 |
using a
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
270 |
apply (rule_tac allI)
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
271 |
apply (induct_tac "l")
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
272 |
apply (simp)
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
273 |
apply (metis)
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
274 |
done
|
163
|
275 |
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
276 |
ML {*
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
277 |
fun regularize thm rty rel rel_eqv lthy =
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
278 |
let
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
279 |
val g = build_regularize_goal thm rty rel lthy;
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
280 |
val tac =
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
281 |
atac ORELSE' (simp_tac ((Simplifier.context lthy HOL_ss) addsimps
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
282 |
[(@{thm equiv_res_forall} OF [rel_eqv]),
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
283 |
(@{thm equiv_res_exists} OF [rel_eqv])])) THEN_ALL_NEW
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
284 |
((rtac @{thm RIGHT_RES_FORALL_REGULAR}) THEN'
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
285 |
(RANGE [fn _ => all_tac, atac]) THEN' (MetisTools.metis_tac lthy []));
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
286 |
val cgoal = cterm_of (ProofContext.theory_of lthy) g;
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
287 |
val cthm = Goal.prove_internal [] cgoal (fn _ => tac 1);
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
288 |
in
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
289 |
cthm OF [thm]
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
290 |
end
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
291 |
*}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
292 |
|
163
|
293 |
|
|
294 |
prove list_induct_r: {*
|
|
295 |
build_regularize_goal (atomize_thm @{thm list_induct_hol4}) @{typ "'a List.list"} @{term "op \<approx>"} @{context} *}
|
|
296 |
apply (simp only: equiv_res_forall[OF equiv_list_eq])
|
|
297 |
thm RIGHT_RES_FORALL_REGULAR
|
|
298 |
apply (rule RIGHT_RES_FORALL_REGULAR)
|
|
299 |
prefer 2
|
|
300 |
apply (assumption)
|
|
301 |
apply (metis)
|
|
302 |
done
|
|
303 |
|
|
304 |
ML {*
|
|
305 |
fun instantiate_tac thm = Subgoal.FOCUS (fn {concl, ...} =>
|
|
306 |
let
|
|
307 |
val pat = Drule.strip_imp_concl (cprop_of thm)
|
|
308 |
val insts = Thm.match (pat, concl)
|
|
309 |
in
|
|
310 |
rtac (Drule.instantiate insts thm) 1
|
|
311 |
end
|
|
312 |
handle _ => no_tac
|
|
313 |
)
|
|
314 |
*}
|
|
315 |
|
|
316 |
ML {*
|
|
317 |
fun res_forall_rsp_tac ctxt =
|
|
318 |
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
|
|
319 |
THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI}
|
|
320 |
THEN' instantiate_tac @{thm RES_FORALL_RSP} ctxt THEN'
|
|
321 |
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
|
|
322 |
*}
|
|
323 |
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
324 |
ML {*
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
325 |
fun res_exists_rsp_tac ctxt =
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
326 |
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
327 |
THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
328 |
THEN' instantiate_tac @{thm RES_EXISTS_RSP} ctxt THEN'
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
329 |
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
330 |
*}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
331 |
|
163
|
332 |
|
|
333 |
ML {*
|
|
334 |
fun quotient_tac quot_thm =
|
|
335 |
REPEAT_ALL_NEW (FIRST' [
|
|
336 |
rtac @{thm FUN_QUOTIENT},
|
|
337 |
rtac quot_thm,
|
|
338 |
rtac @{thm IDENTITY_QUOTIENT}
|
|
339 |
])
|
|
340 |
*}
|
|
341 |
|
|
342 |
ML {*
|
166
|
343 |
fun LAMBDA_RES_TAC ctxt i st =
|
|
344 |
(case (term_of o #concl o fst) (Subgoal.focus ctxt i st) of
|
164
|
345 |
(_ $ (_ $ (Abs(_,_,_))$(Abs(_,_,_)))) =>
|
|
346 |
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
|
|
347 |
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
|
166
|
348 |
| _ => fn _ => no_tac) i st
|
164
|
349 |
*}
|
|
350 |
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
351 |
ML {*
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
352 |
fun WEAK_LAMBDA_RES_TAC ctxt i st =
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
353 |
(case (term_of o #concl o fst) (Subgoal.focus ctxt i st) of
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
354 |
(_ $ (_ $ _$(Abs(_,_,_)))) =>
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
355 |
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
356 |
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
357 |
| (_ $ (_ $ (Abs(_,_,_))$_)) =>
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
358 |
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
359 |
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
360 |
| _ => fn _ => no_tac) i st
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
361 |
*}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
362 |
|
164
|
363 |
|
|
364 |
ML {*
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
365 |
fun r_mk_comb_tac ctxt quot_thm reflex_thm trans_thm rsp_thms =
|
163
|
366 |
(FIRST' [
|
|
367 |
rtac @{thm FUN_QUOTIENT},
|
|
368 |
rtac quot_thm,
|
|
369 |
rtac @{thm IDENTITY_QUOTIENT},
|
|
370 |
rtac @{thm ext},
|
|
371 |
rtac trans_thm,
|
164
|
372 |
LAMBDA_RES_TAC ctxt,
|
163
|
373 |
res_forall_rsp_tac ctxt,
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
374 |
res_exists_rsp_tac ctxt,
|
166
|
375 |
(
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
376 |
(simp_tac ((Simplifier.context ctxt HOL_ss) addsimps (rsp_thms @ @{thms ho_all_prs ho_ex_prs})))
|
166
|
377 |
THEN_ALL_NEW (fn _ => no_tac)
|
|
378 |
),
|
163
|
379 |
(instantiate_tac @{thm REP_ABS_RSP(1)} ctxt THEN' (RANGE [quotient_tac quot_thm])),
|
164
|
380 |
rtac refl,
|
163
|
381 |
(instantiate_tac @{thm APPLY_RSP} ctxt THEN' (RANGE [quotient_tac quot_thm, quotient_tac quot_thm])),
|
164
|
382 |
rtac reflex_thm,
|
163
|
383 |
atac,
|
|
384 |
(
|
|
385 |
(simp_tac ((Simplifier.context @{context} HOL_ss) addsimps @{thms FUN_REL.simps}))
|
|
386 |
THEN_ALL_NEW (fn _ => no_tac)
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
387 |
),
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
388 |
WEAK_LAMBDA_RES_TAC ctxt
|
163
|
389 |
])
|
|
390 |
*}
|
|
391 |
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
392 |
ML Goal.prove
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
393 |
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
394 |
ML {*
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
395 |
fun repabs_eq lthy thm constructors rty qty quot_thm reflex_thm trans_thm rsp_thms =
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
396 |
let
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
397 |
val rt = build_repabs_term lthy thm constructors rty qty;
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
398 |
val rg = Logic.mk_equals ((Thm.prop_of thm), rt);
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
399 |
fun tac ctxt = (ObjectLogic.full_atomize_tac) THEN'
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
400 |
(REPEAT_ALL_NEW (r_mk_comb_tac ctxt quot_thm reflex_thm trans_thm rsp_thms));
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
401 |
val cthm = Goal.prove lthy [] [] rg (fn x => tac (#context x) 1);
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
402 |
in
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
403 |
(rt, cthm, thm)
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
404 |
end
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
405 |
*}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
406 |
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
407 |
ML {*
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
408 |
fun repabs_eq2 lthy (rt, thm, othm) =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
409 |
let
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
410 |
fun tac2 ctxt =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
411 |
(simp_tac ((Simplifier.context ctxt empty_ss) addsimps [symmetric thm]))
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
412 |
THEN' (rtac othm);
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
413 |
val cthm = Goal.prove lthy [] [] rt (fn x => tac2 (#context x) 1);
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
414 |
in
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
415 |
cthm
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
416 |
end
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
417 |
*}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
418 |
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
419 |
|
163
|
420 |
ML {*
|
|
421 |
fun r_mk_comb_tac_fset ctxt =
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
422 |
r_mk_comb_tac ctxt @{thm QUOTIENT_fset} @{thm list_eq_refl} @{thm QUOT_TYPE_I_fset.R_trans2} @{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp}
|
163
|
423 |
*}
|
|
424 |
|
|
425 |
|
|
426 |
ML {* val thm = @{thm list_induct_r} OF [atomize_thm @{thm list_induct_hol4}] *}
|
|
427 |
ML {* val trm_r = build_repabs_goal @{context} thm consts @{typ "'a list"} @{typ "'a fset"} *}
|
|
428 |
ML {* val trm = build_repabs_term @{context} thm consts @{typ "'a list"} @{typ "'a fset"} *}
|
|
429 |
|
166
|
430 |
ML {* trm_r *}
|
163
|
431 |
prove list_induct_tr: trm_r
|
|
432 |
apply (atomize(full))
|
|
433 |
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
|
|
434 |
done
|
|
435 |
|
|
436 |
prove list_induct_t: trm
|
|
437 |
apply (simp only: list_induct_tr[symmetric])
|
|
438 |
apply (tactic {* rtac thm 1 *})
|
|
439 |
done
|
|
440 |
|
|
441 |
thm list.recs(2)
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
442 |
|
163
|
443 |
thm m2
|
|
444 |
ML {* atomize_thm @{thm m2} *}
|
|
445 |
|
|
446 |
prove m2_r_p: {*
|
|
447 |
build_regularize_goal (atomize_thm @{thm m2}) @{typ "'a List.list"} @{term "op \<approx>"} @{context} *}
|
|
448 |
apply (simp add: equiv_res_forall[OF equiv_list_eq])
|
|
449 |
done
|
|
450 |
|
|
451 |
ML {* val m2_r = @{thm m2_r_p} OF [atomize_thm @{thm m2}] *}
|
|
452 |
ML {* val m2_t_g = build_repabs_goal @{context} m2_r consts @{typ "'a list"} @{typ "'a fset"} *}
|
|
453 |
ML {* val m2_t = build_repabs_term @{context} m2_r consts @{typ "'a list"} @{typ "'a fset"} *}
|
|
454 |
prove m2_t_p: m2_t_g
|
|
455 |
apply (atomize(full))
|
|
456 |
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
|
|
457 |
done
|
|
458 |
|
|
459 |
prove m2_t: m2_t
|
|
460 |
apply (simp only: m2_t_p[symmetric])
|
|
461 |
apply (tactic {* rtac m2_r 1 *})
|
|
462 |
done
|
|
463 |
|
|
464 |
lemma id_apply2 [simp]: "id x \<equiv> x"
|
|
465 |
by (simp add: id_def)
|
|
466 |
|
|
467 |
ML {*
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
468 |
val lpis = @{thm LAMBDA_PRS} OF [@{thm QUOTIENT_fset}, @{thm IDENTITY_QUOTIENT}];
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
469 |
val lpist = @{thm "HOL.sym"} OF [lpis];
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
470 |
val lam_prs = MetaSimplifier.rewrite_rule [@{thm id_apply2}] lpist
|
163
|
471 |
*}
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
472 |
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
473 |
text {* the proper way to do it *}
|
163
|
474 |
ML {*
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
475 |
val lpi = Drule.instantiate'
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
476 |
[SOME @{ctyp "'a list"}, SOME @{ctyp "'a fset"}, SOME @{ctyp "bool"}, SOME @{ctyp "bool"}] [] @{thm LAMBDA_PRS};
|
163
|
477 |
*}
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
478 |
prove asdfasdf : {* concl_of lpi *}
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
479 |
apply (tactic {* compose_tac (false, @{thm LAMBDA_PRS}, 2) 1 *})
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
480 |
apply (tactic {* quotient_tac @{thm QUOTIENT_fset} 1 *})
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
481 |
apply (tactic {* quotient_tac @{thm QUOTIENT_fset} 1 *})
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
482 |
done
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
483 |
thm HOL.sym[OF asdfasdf,simplified]
|
163
|
484 |
|
|
485 |
ML {*
|
|
486 |
fun eqsubst_thm ctxt thms thm =
|
|
487 |
let
|
|
488 |
val goalstate = Goal.init (Thm.cprop_of thm)
|
|
489 |
val a' = case (SINGLE (EqSubst.eqsubst_tac ctxt [0] thms 1) goalstate) of
|
|
490 |
NONE => error "eqsubst_thm"
|
|
491 |
| SOME th => cprem_of th 1
|
|
492 |
val tac = (EqSubst.eqsubst_tac ctxt [0] thms 1) THEN simp_tac HOL_ss 1
|
|
493 |
val cgoal = cterm_of (ProofContext.theory_of ctxt) (Logic.mk_equals (term_of (Thm.cprop_of thm), term_of a'))
|
|
494 |
val rt = Toplevel.program (fn () => Goal.prove_internal [] cgoal (fn _ => tac));
|
|
495 |
in
|
|
496 |
Simplifier.rewrite_rule [rt] thm
|
|
497 |
end
|
|
498 |
*}
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
499 |
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
500 |
(* @{thms HOL.sym[OF asdfasdf,simplified]} *)
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
501 |
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
502 |
ML {*
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
503 |
fun simp_lam_prs lthy thm =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
504 |
simp_lam_prs lthy (eqsubst_thm lthy [lam_prs] thm)
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
505 |
handle _ => thm
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
506 |
*}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
507 |
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
508 |
ML {* val m2_t' = eqsubst_thm @{context} [lam_prs] @{thm m2_t} *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
509 |
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
510 |
ML {*
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
511 |
fun simp_allex_prs lthy thm =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
512 |
let
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
513 |
val rwf = @{thm FORALL_PRS[OF QUOTIENT_fset]};
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
514 |
val rwfs = @{thm "HOL.sym"} OF [spec OF [rwf]];
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
515 |
val rwe = @{thm EXISTS_PRS[OF QUOTIENT_fset]};
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
516 |
val rwes = @{thm "HOL.sym"} OF [spec OF [rwe]]
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
517 |
in
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
518 |
(simp_allex_prs lthy (eqsubst_thm lthy [rwfs, rwes] thm))
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
519 |
end
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
520 |
handle _ => thm
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
521 |
*}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
522 |
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
523 |
ML {* val ithm = simp_allex_prs @{context} m2_t' *}
|
164
|
524 |
ML {* val rthm = MetaSimplifier.rewrite_rule fset_defs_sym ithm *}
|
|
525 |
ML {* ObjectLogic.rulify rthm *}
|
163
|
526 |
|
|
527 |
|
|
528 |
ML {* val card1_suc_f = Thm.freezeT (atomize_thm @{thm card1_suc}) *}
|
|
529 |
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
530 |
prove card1_suc_r_p: {*
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
531 |
build_regularize_goal (atomize_thm @{thm card1_suc}) @{typ "'a List.list"} @{term "op \<approx>"} @{context} *}
|
163
|
532 |
apply (simp add: equiv_res_forall[OF equiv_list_eq] equiv_res_exists[OF equiv_list_eq])
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
533 |
done
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
534 |
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
535 |
ML {* val card1_suc_r = @{thm card1_suc_r_p} OF [atomize_thm @{thm card1_suc}] *}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
536 |
ML {* val card1_suc_t_g = build_repabs_goal @{context} card1_suc_r consts @{typ "'a list"} @{typ "'a fset"} *}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
537 |
ML {* val card1_suc_t = build_repabs_term @{context} card1_suc_r consts @{typ "'a list"} @{typ "'a fset"} *}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
538 |
prove card1_suc_t_p: card1_suc_t_g
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
539 |
apply (atomize(full))
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
540 |
apply (tactic {* REPEAT_ALL_NEW (r_mk_comb_tac_fset @{context}) 1 *})
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
541 |
done
|
163
|
542 |
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
543 |
prove card1_suc_t: card1_suc_t
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
544 |
apply (simp only: card1_suc_t_p[symmetric])
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
545 |
apply (tactic {* rtac card1_suc_r 1 *})
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
546 |
done
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
547 |
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
548 |
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
549 |
ML {* val card1_suc_t_n = @{thm card1_suc_t} *}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
550 |
ML {* val card1_suc_t' = eqsubst_thm @{context} @{thms HOL.sym[OF asdfasdf,simplified]} @{thm card1_suc_t} *}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
551 |
ML {* val card1_suc_t'' = eqsubst_thm @{context} @{thms HOL.sym[OF asdfasdf,simplified]} card1_suc_t' *}
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
552 |
ML {* val ithm = simp_allex_prs @{context} card1_suc_t'' *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
553 |
ML {* val rthm = MetaSimplifier.rewrite_rule fset_defs_sym ithm *}
|
171
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
554 |
ML {* val qthm = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} rthm *}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
555 |
ML {* ObjectLogic.rulify qthm *}
|
13aab4c59096
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
556 |
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
557 |
thm fold1.simps(2)
|
163
|
558 |
|
|
559 |
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
560 |
ML {* val ind_r_a = atomize_thm @{thm fold1.simps(2)} *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
561 |
ML {* val ind_r_r = regularize ind_r_a @{typ "'a List.list"} @{term "op \<approx>"} @{thm equiv_list_eq} @{context} *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
562 |
ML {*
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
563 |
val rt = build_repabs_term @{context} ind_r_r consts @{typ "'a list"} @{typ "'a fset"}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
564 |
val rg = Logic.mk_equals ((Thm.prop_of ind_r_r), rt);
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
565 |
*}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
566 |
(*prove rg
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
567 |
apply(atomize(full))
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
568 |
apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *})*)
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
569 |
ML {* val (g, thm, othm) =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
570 |
Toplevel.program (fn () =>
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
571 |
repabs_eq @{context} ind_r_r consts @{typ "'a list"} @{typ "'a fset"}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
572 |
@{thm QUOTIENT_fset} @{thm list_eq_refl} @{thm QUOT_TYPE_I_fset.R_trans2}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
573 |
@{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
574 |
)
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
575 |
*}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
576 |
ML {*
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
577 |
fun tac2 ctxt =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
578 |
(simp_tac ((Simplifier.context ctxt empty_ss) addsimps [symmetric thm]))
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
579 |
THEN' (rtac othm); *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
580 |
(* ML {* val cthm = Goal.prove @{context} [] [] g (fn x => tac2 (#context x) 1);
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
581 |
*} *)
|
163
|
582 |
|
|
583 |
ML {*
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
584 |
val ind_r_t2 =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
585 |
Toplevel.program (fn () =>
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
586 |
repabs_eq2 @{context} (g, thm, othm)
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
587 |
)
|
163
|
588 |
*}
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
589 |
ML {* val ind_r_l = simp_lam_prs @{context} ind_r_t2 *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
590 |
ML {* val ind_r_a = simp_allex_prs @{context} ind_r_l *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
591 |
ML {* val ind_r_d = MetaSimplifier.rewrite_rule fset_defs_sym ind_r_a *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
592 |
ML {* val ind_r_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} ind_r_d *}
|
163
|
593 |
|
|
594 |
ML {*
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
595 |
fun lift thm =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
596 |
let
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
597 |
val ind_r_a = atomize_thm thm;
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
598 |
val ind_r_r = regularize ind_r_a @{typ "'a List.list"} @{term "op \<approx>"} @{thm equiv_list_eq} @{context};
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
599 |
val (g, t, ot) =
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
600 |
repabs_eq @{context} ind_r_r consts @{typ "'a list"} @{typ "'a fset"}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
601 |
@{thm QUOTIENT_fset} @{thm list_eq_refl} @{thm QUOT_TYPE_I_fset.R_trans2}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
602 |
@{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp};
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
603 |
val ind_r_t = repabs_eq2 @{context} (g, t, ot);
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
604 |
val ind_r_l = simp_lam_prs @{context} ind_r_t;
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
605 |
val ind_r_a = simp_allex_prs @{context} ind_r_l;
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
606 |
val ind_r_d = MetaSimplifier.rewrite_rule fset_defs_sym ind_r_a;
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
607 |
val ind_r_s = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} ind_r_d
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
608 |
in
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
609 |
ObjectLogic.rulify ind_r_s
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
610 |
end
|
163
|
611 |
*}
|
|
612 |
|
172
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
613 |
ML {* lift @{thm m2} *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
614 |
ML {* lift @{thm m1} *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
615 |
ML {* lift @{thm list_eq.intros(4)} *}
|
da38ce2711a6
More infrastructure for automatic lifting of theorems lifted before
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
616 |
ML {* lift @{thm list_eq.intros(5)} *}
|
163
|
617 |
|
|
618 |
(* Has all the theorems about fset plugged in. These should be parameters to the tactic *)
|
|
619 |
ML {*
|
|
620 |
fun transconv_fset_tac' ctxt =
|
|
621 |
(LocalDefs.unfold_tac @{context} fset_defs) THEN
|
|
622 |
ObjectLogic.full_atomize_tac 1 THEN
|
|
623 |
REPEAT_ALL_NEW (FIRST' [
|
|
624 |
rtac @{thm list_eq_refl},
|
|
625 |
rtac @{thm cons_preserves},
|
164
|
626 |
rtac @{thm memb_rsp},
|
163
|
627 |
rtac @{thm card1_rsp},
|
|
628 |
rtac @{thm QUOT_TYPE_I_fset.R_trans2},
|
|
629 |
CHANGED o (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms QUOT_TYPE_I_fset.REP_ABS_rsp})),
|
|
630 |
Cong_Tac.cong_tac @{thm cong},
|
|
631 |
rtac @{thm ext}
|
|
632 |
]) 1
|
|
633 |
*}
|
|
634 |
|
|
635 |
ML {*
|
|
636 |
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm m1}))
|
|
637 |
val goal = build_repabs_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
|
|
638 |
val cgoal = cterm_of @{theory} goal
|
|
639 |
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
|
|
640 |
*}
|
|
641 |
|
|
642 |
(*notation ( output) "prop" ("#_" [1000] 1000) *)
|
|
643 |
notation ( output) "Trueprop" ("#_" [1000] 1000)
|
|
644 |
|
|
645 |
|
|
646 |
prove {* (Thm.term_of cgoal2) *}
|
|
647 |
apply (tactic {* transconv_fset_tac' @{context} *})
|
|
648 |
done
|
|
649 |
|
|
650 |
thm length_append (* Not true but worth checking that the goal is correct *)
|
|
651 |
ML {*
|
|
652 |
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm length_append}))
|
|
653 |
val goal = build_repabs_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
|
|
654 |
val cgoal = cterm_of @{theory} goal
|
|
655 |
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
|
|
656 |
*}
|
|
657 |
prove {* Thm.term_of cgoal2 *}
|
|
658 |
apply (tactic {* transconv_fset_tac' @{context} *})
|
|
659 |
sorry
|
|
660 |
|
|
661 |
thm m2
|
|
662 |
ML {*
|
|
663 |
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm m2}))
|
|
664 |
val goal = build_repabs_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
|
|
665 |
val cgoal = cterm_of @{theory} goal
|
|
666 |
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
|
|
667 |
*}
|
|
668 |
prove {* Thm.term_of cgoal2 *}
|
|
669 |
apply (tactic {* transconv_fset_tac' @{context} *})
|
|
670 |
done
|
|
671 |
|
|
672 |
thm list_eq.intros(4)
|
|
673 |
ML {*
|
|
674 |
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm list_eq.intros(4)}))
|
|
675 |
val goal = build_repabs_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
|
|
676 |
val cgoal = cterm_of @{theory} goal
|
|
677 |
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite false fset_defs_sym cgoal)
|
|
678 |
val cgoal3 = Thm.rhs_of (MetaSimplifier.rewrite true @{thms QUOT_TYPE_I_fset.thm10} cgoal2)
|
|
679 |
*}
|
|
680 |
|
|
681 |
(* It is the same, but we need a name for it. *)
|
|
682 |
prove zzz : {* Thm.term_of cgoal3 *}
|
|
683 |
apply (tactic {* transconv_fset_tac' @{context} *})
|
|
684 |
done
|
|
685 |
|
|
686 |
(*lemma zzz' :
|
|
687 |
"(REP_fset (INSERT a (INSERT a (ABS_fset xs))) \<approx> REP_fset (INSERT a (ABS_fset xs)))"
|
|
688 |
using list_eq.intros(4) by (simp only: zzz)
|
|
689 |
|
|
690 |
thm QUOT_TYPE_I_fset.REPS_same
|
|
691 |
ML {* val zzz'' = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same} @{thm zzz'} *}
|
|
692 |
*)
|
|
693 |
|
|
694 |
thm list_eq.intros(5)
|
|
695 |
(* prove {* build_repabs_goal @{context} (atomize_thm @{thm list_eq.intros(5)}) consts @{typ "'a list"} @{typ "'a fset"} *} *)
|
|
696 |
ML {*
|
|
697 |
val m1_novars = snd(no_vars ((Context.Theory @{theory}), @{thm list_eq.intros(5)}))
|
|
698 |
val goal = build_repabs_goal @{context} m1_novars consts @{typ "'a list"} @{typ "'a fset"}
|
|
699 |
val cgoal = cterm_of @{theory} goal
|
|
700 |
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite true fset_defs_sym cgoal)
|
|
701 |
*}
|
|
702 |
prove {* Thm.term_of cgoal2 *}
|
|
703 |
apply (tactic {* transconv_fset_tac' @{context} *})
|
|
704 |
done
|
|
705 |
|
|
706 |
ML {*
|
|
707 |
fun lift_theorem_fset_aux thm lthy =
|
|
708 |
let
|
|
709 |
val ((_, [novars]), lthy2) = Variable.import true [thm] lthy;
|
|
710 |
val goal = build_repabs_goal @{context} novars consts @{typ "'a list"} @{typ "'a fset"};
|
|
711 |
val cgoal = cterm_of @{theory} goal;
|
|
712 |
val cgoal2 = Thm.rhs_of (MetaSimplifier.rewrite true fset_defs_sym cgoal);
|
|
713 |
val tac = transconv_fset_tac' @{context};
|
|
714 |
val cthm = Goal.prove_internal [] cgoal2 (fn _ => tac);
|
|
715 |
val nthm = MetaSimplifier.rewrite_rule [symmetric cthm] (snd (no_vars (Context.Theory @{theory}, thm)))
|
|
716 |
val nthm2 = MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.REPS_same QUOT_TYPE_I_fset.thm10} nthm;
|
|
717 |
val [nthm3] = ProofContext.export lthy2 lthy [nthm2]
|
|
718 |
in
|
|
719 |
nthm3
|
|
720 |
end
|
|
721 |
*}
|
|
722 |
|
|
723 |
ML {* lift_theorem_fset_aux @{thm m1} @{context} *}
|
|
724 |
ML {*
|
|
725 |
fun lift_theorem_fset name thm lthy =
|
|
726 |
let
|
|
727 |
val lifted_thm = lift_theorem_fset_aux thm lthy;
|
|
728 |
val (_, lthy2) = note (name, lifted_thm) lthy;
|
|
729 |
in
|
|
730 |
lthy2
|
|
731 |
end;
|
|
732 |
*}
|
|
733 |
|
|
734 |
(* These do not work without proper definitions to rewrite back *)
|
|
735 |
local_setup {* lift_theorem_fset @{binding "m1_lift"} @{thm m1} *}
|
|
736 |
local_setup {* lift_theorem_fset @{binding "leqi4_lift"} @{thm list_eq.intros(4)} *}
|
|
737 |
local_setup {* lift_theorem_fset @{binding "leqi5_lift"} @{thm list_eq.intros(5)} *}
|
|
738 |
local_setup {* lift_theorem_fset @{binding "m2_lift"} @{thm m2} *}
|
|
739 |
thm m1_lift
|
|
740 |
thm leqi4_lift
|
|
741 |
thm leqi5_lift
|
|
742 |
thm m2_lift
|
|
743 |
ML {* @{thm card1_suc_r} OF [card1_suc_f] *}
|
|
744 |
(*ML {* Toplevel.program (fn () => lift_theorem_fset @{binding "card_suc"}
|
|
745 |
(@{thm card1_suc_r} OF [card1_suc_f]) @{context}) *}*)
|
|
746 |
(*local_setup {* lift_theorem_fset @{binding "card_suc"} @{thm card1_suc} *}*)
|
|
747 |
|
|
748 |
thm leqi4_lift
|
|
749 |
ML {*
|
|
750 |
val (nam, typ) = hd (Term.add_vars (prop_of @{thm leqi4_lift}) [])
|
|
751 |
val (_, l) = dest_Type typ
|
|
752 |
val t = Type ("FSet.fset", l)
|
|
753 |
val v = Var (nam, t)
|
|
754 |
val cv = cterm_of @{theory} ((term_of @{cpat "REP_fset"}) $ v)
|
|
755 |
*}
|
|
756 |
|
|
757 |
ML {*
|
|
758 |
Toplevel.program (fn () =>
|
|
759 |
MetaSimplifier.rewrite_rule @{thms QUOT_TYPE_I_fset.thm10} (
|
|
760 |
Drule.instantiate' [] [NONE, SOME (cv)] @{thm leqi4_lift}
|
|
761 |
)
|
|
762 |
)
|
|
763 |
*}
|
|
764 |
|
|
765 |
|
|
766 |
|
|
767 |
(*prove aaa: {* (Thm.term_of cgoal2) *}
|
|
768 |
apply (tactic {* LocalDefs.unfold_tac @{context} fset_defs *} )
|
|
769 |
apply (atomize(full))
|
|
770 |
apply (tactic {* transconv_fset_tac' @{context} 1 *})
|
|
771 |
done*)
|
|
772 |
|
|
773 |
(*
|
|
774 |
datatype obj1 =
|
|
775 |
OVAR1 "string"
|
|
776 |
| OBJ1 "(string * (string \<Rightarrow> obj1)) list"
|
|
777 |
| INVOKE1 "obj1 \<Rightarrow> string"
|
|
778 |
| UPDATE1 "obj1 \<Rightarrow> string \<Rightarrow> (string \<Rightarrow> obj1)"
|
|
779 |
*)
|
|
780 |
|
|
781 |
end
|