0
|
1 |
theory QuotMain
|
6
|
2 |
imports QuotScript QuotList Prove
|
264
|
3 |
uses ("quotient_info.ML")
|
|
4 |
("quotient.ML")
|
277
|
5 |
("quotient_def.ML")
|
0
|
6 |
begin
|
|
7 |
|
297
|
8 |
|
0
|
9 |
locale QUOT_TYPE =
|
|
10 |
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
|
|
11 |
and Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
|
|
12 |
and Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
|
|
13 |
assumes equiv: "EQUIV R"
|
|
14 |
and rep_prop: "\<And>y. \<exists>x. Rep y = R x"
|
|
15 |
and rep_inverse: "\<And>x. Abs (Rep x) = x"
|
|
16 |
and abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)"
|
|
17 |
and rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)"
|
15
|
18 |
begin
|
0
|
19 |
|
|
20 |
definition
|
200
|
21 |
ABS::"'a \<Rightarrow> 'b"
|
|
22 |
where
|
0
|
23 |
"ABS x \<equiv> Abs (R x)"
|
|
24 |
|
|
25 |
definition
|
200
|
26 |
REP::"'b \<Rightarrow> 'a"
|
|
27 |
where
|
0
|
28 |
"REP a = Eps (Rep a)"
|
|
29 |
|
15
|
30 |
lemma lem9:
|
0
|
31 |
shows "R (Eps (R x)) = R x"
|
|
32 |
proof -
|
|
33 |
have a: "R x x" using equiv by (simp add: EQUIV_REFL_SYM_TRANS REFL_def)
|
|
34 |
then have "R x (Eps (R x))" by (rule someI)
|
15
|
35 |
then show "R (Eps (R x)) = R x"
|
0
|
36 |
using equiv unfolding EQUIV_def by simp
|
|
37 |
qed
|
|
38 |
|
|
39 |
theorem thm10:
|
24
|
40 |
shows "ABS (REP a) \<equiv> a"
|
|
41 |
apply (rule eq_reflection)
|
|
42 |
unfolding ABS_def REP_def
|
0
|
43 |
proof -
|
15
|
44 |
from rep_prop
|
0
|
45 |
obtain x where eq: "Rep a = R x" by auto
|
|
46 |
have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp
|
|
47 |
also have "\<dots> = Abs (R x)" using lem9 by simp
|
|
48 |
also have "\<dots> = Abs (Rep a)" using eq by simp
|
|
49 |
also have "\<dots> = a" using rep_inverse by simp
|
|
50 |
finally
|
|
51 |
show "Abs (R (Eps (Rep a))) = a" by simp
|
|
52 |
qed
|
|
53 |
|
15
|
54 |
lemma REP_refl:
|
0
|
55 |
shows "R (REP a) (REP a)"
|
|
56 |
unfolding REP_def
|
|
57 |
by (simp add: equiv[simplified EQUIV_def])
|
|
58 |
|
|
59 |
lemma lem7:
|
22
|
60 |
shows "(R x = R y) = (Abs (R x) = Abs (R y))"
|
0
|
61 |
apply(rule iffI)
|
|
62 |
apply(simp)
|
|
63 |
apply(drule rep_inject[THEN iffD2])
|
|
64 |
apply(simp add: abs_inverse)
|
|
65 |
done
|
15
|
66 |
|
0
|
67 |
theorem thm11:
|
|
68 |
shows "R r r' = (ABS r = ABS r')"
|
|
69 |
unfolding ABS_def
|
|
70 |
by (simp only: equiv[simplified EQUIV_def] lem7)
|
|
71 |
|
4
|
72 |
|
2
|
73 |
lemma REP_ABS_rsp:
|
4
|
74 |
shows "R f (REP (ABS g)) = R f g"
|
|
75 |
and "R (REP (ABS g)) f = R g f"
|
23
|
76 |
by (simp_all add: thm10 thm11)
|
4
|
77 |
|
0
|
78 |
lemma QUOTIENT:
|
|
79 |
"QUOTIENT R ABS REP"
|
|
80 |
apply(unfold QUOTIENT_def)
|
|
81 |
apply(simp add: thm10)
|
|
82 |
apply(simp add: REP_refl)
|
|
83 |
apply(subst thm11[symmetric])
|
|
84 |
apply(simp add: equiv[simplified EQUIV_def])
|
|
85 |
done
|
|
86 |
|
21
|
87 |
lemma R_trans:
|
49
|
88 |
assumes ab: "R a b"
|
|
89 |
and bc: "R b c"
|
22
|
90 |
shows "R a c"
|
21
|
91 |
proof -
|
|
92 |
have tr: "TRANS R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
|
|
93 |
moreover have ab: "R a b" by fact
|
|
94 |
moreover have bc: "R b c" by fact
|
22
|
95 |
ultimately show "R a c" unfolding TRANS_def by blast
|
21
|
96 |
qed
|
|
97 |
|
|
98 |
lemma R_sym:
|
49
|
99 |
assumes ab: "R a b"
|
22
|
100 |
shows "R b a"
|
21
|
101 |
proof -
|
|
102 |
have re: "SYM R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
|
22
|
103 |
then show "R b a" using ab unfolding SYM_def by blast
|
21
|
104 |
qed
|
|
105 |
|
49
|
106 |
lemma R_trans2:
|
|
107 |
assumes ac: "R a c"
|
22
|
108 |
and bd: "R b d"
|
21
|
109 |
shows "R a b = R c d"
|
200
|
110 |
using ac bd
|
|
111 |
by (blast intro: R_trans R_sym)
|
21
|
112 |
|
|
113 |
lemma REPS_same:
|
25
|
114 |
shows "R (REP a) (REP b) \<equiv> (a = b)"
|
38
|
115 |
proof -
|
|
116 |
have "R (REP a) (REP b) = (a = b)"
|
|
117 |
proof
|
|
118 |
assume as: "R (REP a) (REP b)"
|
|
119 |
from rep_prop
|
|
120 |
obtain x y
|
|
121 |
where eqs: "Rep a = R x" "Rep b = R y" by blast
|
|
122 |
from eqs have "R (Eps (R x)) (Eps (R y))" using as unfolding REP_def by simp
|
|
123 |
then have "R x (Eps (R y))" using lem9 by simp
|
|
124 |
then have "R (Eps (R y)) x" using R_sym by blast
|
|
125 |
then have "R y x" using lem9 by simp
|
|
126 |
then have "R x y" using R_sym by blast
|
|
127 |
then have "ABS x = ABS y" using thm11 by simp
|
|
128 |
then have "Abs (Rep a) = Abs (Rep b)" using eqs unfolding ABS_def by simp
|
|
129 |
then show "a = b" using rep_inverse by simp
|
|
130 |
next
|
|
131 |
assume ab: "a = b"
|
|
132 |
have "REFL R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp
|
|
133 |
then show "R (REP a) (REP b)" unfolding REFL_def using ab by auto
|
|
134 |
qed
|
|
135 |
then show "R (REP a) (REP b) \<equiv> (a = b)" by simp
|
21
|
136 |
qed
|
|
137 |
|
0
|
138 |
end
|
|
139 |
|
|
140 |
section {* type definition for the quotient type *}
|
|
141 |
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
142 |
(* the auxiliary data for the quotient types *)
|
264
|
143 |
use "quotient_info.ML"
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
144 |
|
185
|
145 |
declare [[map list = (map, LIST_REL)]]
|
|
146 |
declare [[map * = (prod_fun, prod_rel)]]
|
|
147 |
declare [[map "fun" = (fun_map, FUN_REL)]]
|
|
148 |
|
|
149 |
ML {* maps_lookup @{theory} "List.list" *}
|
|
150 |
ML {* maps_lookup @{theory} "*" *}
|
|
151 |
ML {* maps_lookup @{theory} "fun" *}
|
174
|
152 |
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
153 |
|
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
154 |
(* definition of the quotient types *)
|
277
|
155 |
(* FIXME: should be called quotient_typ.ML *)
|
268
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
156 |
use "quotient.ML"
|
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
157 |
|
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
158 |
|
277
|
159 |
(* lifting of constants *)
|
|
160 |
use "quotient_def.ML"
|
|
161 |
|
310
fec6301a1989
added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
162 |
|
fec6301a1989
added a container for quotient constants (does not work yet though)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
163 |
|
139
|
164 |
section {* ATOMIZE *}
|
|
165 |
|
|
166 |
lemma atomize_eqv[atomize]:
|
|
167 |
shows "(Trueprop A \<equiv> Trueprop B) \<equiv> (A \<equiv> B)"
|
|
168 |
proof
|
|
169 |
assume "A \<equiv> B"
|
|
170 |
then show "Trueprop A \<equiv> Trueprop B" by unfold
|
|
171 |
next
|
|
172 |
assume *: "Trueprop A \<equiv> Trueprop B"
|
|
173 |
have "A = B"
|
|
174 |
proof (cases A)
|
|
175 |
case True
|
|
176 |
have "A" by fact
|
|
177 |
then show "A = B" using * by simp
|
|
178 |
next
|
|
179 |
case False
|
|
180 |
have "\<not>A" by fact
|
|
181 |
then show "A = B" using * by auto
|
|
182 |
qed
|
|
183 |
then show "A \<equiv> B" by (rule eq_reflection)
|
|
184 |
qed
|
|
185 |
|
|
186 |
ML {*
|
|
187 |
fun atomize_thm thm =
|
|
188 |
let
|
221
|
189 |
val thm' = Thm.freezeT (forall_intr_vars thm)
|
139
|
190 |
val thm'' = ObjectLogic.atomize (cprop_of thm')
|
|
191 |
in
|
221
|
192 |
@{thm Pure.equal_elim_rule1} OF [thm'', thm']
|
139
|
193 |
end
|
|
194 |
*}
|
|
195 |
|
140
|
196 |
ML {* atomize_thm @{thm list.induct} *}
|
139
|
197 |
|
|
198 |
section {* REGULARIZE *}
|
282
|
199 |
(*
|
|
200 |
|
|
201 |
Regularizing a theorem means:
|
|
202 |
- Quantifiers over a type that needs lifting are replaced by
|
|
203 |
bounded quantifiers, for example:
|
|
204 |
\<forall>x. P \<Longrightarrow> \<forall>x\<in>(Respects R). P
|
|
205 |
- Abstractions over a type that needs lifting are replaced
|
303
|
206 |
by bounded abstractions:
|
282
|
207 |
\<lambda>x. P \<Longrightarrow> Ball (Respects R) (\<lambda>x. P)
|
|
208 |
|
|
209 |
- Equalities over the type being lifted are replaced by
|
|
210 |
appropriate relations:
|
|
211 |
A = B \<Longrightarrow> A \<approx> B
|
|
212 |
Example with more complicated types of A, B:
|
|
213 |
A = B \<Longrightarrow> (op = \<Longrightarrow> op \<approx>) A B
|
|
214 |
|
|
215 |
Regularizing is done in 3 phases:
|
|
216 |
- First a regularized term is created
|
|
217 |
- Next we prove that the original theorem implies the new one
|
|
218 |
- Finally using MP we get the new theorem.
|
|
219 |
|
|
220 |
To prove that the old theorem implies the new one, we first
|
|
221 |
atomize it and then try:
|
|
222 |
- Reflexivity of the relation
|
|
223 |
- Assumption
|
|
224 |
- Elimnating quantifiers on both sides of toplevel implication
|
|
225 |
- Simplifying implications on both sides of toplevel implication
|
|
226 |
- Ball (Respects ?E) ?P = All ?P
|
|
227 |
- (\<And>x. ?R x \<Longrightarrow> ?P x \<longrightarrow> ?Q x) \<Longrightarrow> All ?P \<longrightarrow> Ball ?R ?Q
|
|
228 |
|
|
229 |
*)
|
139
|
230 |
|
|
231 |
text {* tyRel takes a type and builds a relation that a quantifier over this
|
|
232 |
type needs to respect. *}
|
|
233 |
ML {*
|
|
234 |
fun tyRel ty rty rel lthy =
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
235 |
if Sign.typ_instance (ProofContext.theory_of lthy) (ty, rty)
|
139
|
236 |
then rel
|
|
237 |
else (case ty of
|
|
238 |
Type (s, tys) =>
|
|
239 |
let
|
|
240 |
val tys_rel = map (fn ty => ty --> ty --> @{typ bool}) tys;
|
|
241 |
val ty_out = ty --> ty --> @{typ bool};
|
|
242 |
val tys_out = tys_rel ---> ty_out;
|
|
243 |
in
|
|
244 |
(case (maps_lookup (ProofContext.theory_of lthy) s) of
|
303
|
245 |
SOME (info) => list_comb (Const (#relfun info, tys_out),
|
|
246 |
map (fn ty => tyRel ty rty rel lthy) tys)
|
139
|
247 |
| NONE => HOLogic.eq_const ty
|
|
248 |
)
|
|
249 |
end
|
|
250 |
| _ => HOLogic.eq_const ty)
|
|
251 |
*}
|
|
252 |
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
253 |
(*
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
254 |
ML {* cterm_of @{theory}
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
255 |
(tyRel @{typ "'a \<Rightarrow> 'a list \<Rightarrow> 't \<Rightarrow> 't"} (Logic.varifyT @{typ "'a list"})
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
256 |
@{term "f::('a list \<Rightarrow> 'a list \<Rightarrow> bool)"} @{context})
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
257 |
*}
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
258 |
*)
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
259 |
|
139
|
260 |
definition
|
|
261 |
Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
|
|
262 |
where
|
|
263 |
"(x \<in> p) \<Longrightarrow> (Babs p m x = m x)"
|
|
264 |
(* TODO: Consider defining it with an "if"; sth like:
|
|
265 |
Babs p m = \<lambda>x. if x \<in> p then m x else undefined
|
|
266 |
*)
|
|
267 |
|
|
268 |
ML {*
|
|
269 |
fun needs_lift (rty as Type (rty_s, _)) ty =
|
|
270 |
case ty of
|
|
271 |
Type (s, tys) =>
|
|
272 |
(s = rty_s) orelse (exists (needs_lift rty) tys)
|
|
273 |
| _ => false
|
|
274 |
|
|
275 |
*}
|
|
276 |
|
|
277 |
ML {*
|
|
278 |
fun mk_babs ty ty' = Const (@{const_name "Babs"}, [ty' --> @{typ bool}, ty] ---> ty)
|
|
279 |
fun mk_ball ty = Const (@{const_name "Ball"}, [ty, ty] ---> @{typ bool})
|
|
280 |
fun mk_bex ty = Const (@{const_name "Bex"}, [ty, ty] ---> @{typ bool})
|
|
281 |
fun mk_resp ty = Const (@{const_name Respects}, [[ty, ty] ---> @{typ bool}, ty] ---> @{typ bool})
|
|
282 |
*}
|
|
283 |
|
|
284 |
(* applies f to the subterm of an abstractions, otherwise to the given term *)
|
|
285 |
ML {*
|
|
286 |
fun apply_subt f trm =
|
|
287 |
case trm of
|
145
|
288 |
Abs (x, T, t) =>
|
|
289 |
let
|
|
290 |
val (x', t') = Term.dest_abs (x, T, t)
|
|
291 |
in
|
|
292 |
Term.absfree (x', T, f t')
|
|
293 |
end
|
139
|
294 |
| _ => f trm
|
|
295 |
*}
|
|
296 |
|
|
297 |
(* FIXME: if there are more than one quotient, then you have to look up the relation *)
|
|
298 |
ML {*
|
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
299 |
fun my_reg lthy rel rty trm =
|
139
|
300 |
case trm of
|
|
301 |
Abs (x, T, t) =>
|
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
302 |
if (needs_lift rty T) then
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
303 |
let
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
304 |
val rrel = tyRel T rty rel lthy
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
305 |
in
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
306 |
(mk_babs (fastype_of trm) T) $ (mk_resp T $ rrel) $ (apply_subt (my_reg lthy rel rty) trm)
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
307 |
end
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
308 |
else
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
309 |
Abs(x, T, (apply_subt (my_reg lthy rel rty) t))
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
310 |
| Const (@{const_name "All"}, ty) $ (t as Abs (x, T, _)) =>
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
311 |
let
|
139
|
312 |
val ty1 = domain_type ty
|
|
313 |
val ty2 = domain_type ty1
|
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
314 |
val rrel = tyRel T rty rel lthy
|
139
|
315 |
in
|
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
316 |
if (needs_lift rty T) then
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
317 |
(mk_ball ty1) $ (mk_resp ty2 $ rrel) $ (apply_subt (my_reg lthy rel rty) t)
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
318 |
else
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
319 |
Const (@{const_name "All"}, ty) $ apply_subt (my_reg lthy rel rty) t
|
139
|
320 |
end
|
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
321 |
| Const (@{const_name "Ex"}, ty) $ (t as Abs (x, T, _)) =>
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
322 |
let
|
139
|
323 |
val ty1 = domain_type ty
|
|
324 |
val ty2 = domain_type ty1
|
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
325 |
val rrel = tyRel T rty rel lthy
|
139
|
326 |
in
|
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
327 |
if (needs_lift rty T) then
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
328 |
(mk_bex ty1) $ (mk_resp ty2 $ rrel) $ (apply_subt (my_reg lthy rel rty) t)
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
329 |
else
|
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
330 |
Const (@{const_name "Ex"}, ty) $ apply_subt (my_reg lthy rel rty) t
|
139
|
331 |
end
|
251
|
332 |
| Const (@{const_name "op ="}, ty) $ t =>
|
|
333 |
if needs_lift rty (fastype_of t) then
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
334 |
(tyRel (fastype_of t) rty rel lthy) $ t (* FIXME: t should be regularised too *)
|
251
|
335 |
else Const (@{const_name "op ="}, ty) $ (my_reg lthy rel rty t)
|
248
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
336 |
| t1 $ t2 => (my_reg lthy rel rty t1) $ (my_reg lthy rel rty t2)
|
139
|
337 |
| _ => trm
|
|
338 |
*}
|
|
339 |
|
301
|
340 |
(* For polymorphic types we need to find the type of the Relation term. *)
|
|
341 |
(* TODO: we assume that the relation is a Constant. Is this always true? *)
|
283
|
342 |
ML {*
|
|
343 |
fun my_reg_inst lthy rel rty trm =
|
|
344 |
case rel of
|
|
345 |
Const (n, _) => Syntax.check_term lthy
|
285
|
346 |
(my_reg lthy (Const (n, dummyT)) (Logic.varifyT rty) trm)
|
283
|
347 |
*}
|
|
348 |
|
288
|
349 |
(*
|
|
350 |
ML {*
|
283
|
351 |
val r = Free ("R", dummyT);
|
303
|
352 |
val t = (my_reg_inst @{context} r @{typ "'a list"} @{term "\<forall>(x::'b list). P x"});
|
283
|
353 |
val t2 = Syntax.check_term @{context} t;
|
303
|
354 |
cterm_of @{theory} t2
|
|
355 |
*}
|
|
356 |
*)
|
283
|
357 |
|
141
|
358 |
text {* Assumes that the given theorem is atomized *}
|
140
|
359 |
ML {*
|
|
360 |
fun build_regularize_goal thm rty rel lthy =
|
|
361 |
Logic.mk_implies
|
|
362 |
((prop_of thm),
|
283
|
363 |
(my_reg_inst lthy rel rty (prop_of thm)))
|
140
|
364 |
*}
|
|
365 |
|
303
|
366 |
lemma universal_twice:
|
|
367 |
assumes *: "\<And>x. (P x \<longrightarrow> Q x)"
|
|
368 |
shows "(\<forall>x. P x) \<longrightarrow> (\<forall>x. Q x)"
|
|
369 |
using * by auto
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
370 |
|
303
|
371 |
lemma implication_twice:
|
|
372 |
assumes a: "c \<longrightarrow> a"
|
|
373 |
assumes b: "a \<Longrightarrow> b \<longrightarrow> d"
|
|
374 |
shows "(a \<longrightarrow> b) \<longrightarrow> (c \<longrightarrow> d)"
|
|
375 |
using a b by auto
|
251
|
376 |
|
187
|
377 |
ML {*
|
251
|
378 |
fun regularize thm rty rel rel_eqv rel_refl lthy =
|
187
|
379 |
let
|
288
|
380 |
val goal = build_regularize_goal thm rty rel lthy;
|
187
|
381 |
fun tac ctxt =
|
252
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
382 |
(ObjectLogic.full_atomize_tac) THEN'
|
301
|
383 |
REPEAT_ALL_NEW (FIRST' [
|
|
384 |
rtac rel_refl,
|
|
385 |
atac,
|
|
386 |
rtac @{thm universal_twice},
|
|
387 |
(rtac @{thm impI} THEN' atac),
|
|
388 |
rtac @{thm implication_twice},
|
|
389 |
EqSubst.eqsubst_tac ctxt [0]
|
|
390 |
[(@{thm equiv_res_forall} OF [rel_eqv]),
|
|
391 |
(@{thm equiv_res_exists} OF [rel_eqv])],
|
303
|
392 |
(* For a = b \<longrightarrow> a \<approx> b *)
|
301
|
393 |
(rtac @{thm impI} THEN' (asm_full_simp_tac HOL_ss) THEN' rtac rel_refl),
|
|
394 |
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
|
|
395 |
]);
|
|
396 |
val cthm = Goal.prove lthy [] [] goal
|
|
397 |
(fn {context, ...} => tac context 1);
|
187
|
398 |
in
|
|
399 |
cthm OF [thm]
|
|
400 |
end
|
|
401 |
*}
|
|
402 |
|
140
|
403 |
section {* RepAbs injection *}
|
301
|
404 |
(*
|
|
405 |
|
303
|
406 |
RepAbs injection is done in the following phases:
|
|
407 |
1) build_repabs_term inserts rep-abs pairs in the term
|
|
408 |
2) we prove the equality between the original theorem and this one
|
|
409 |
3) we use Pure.equal_elim_rule1 to get the new theorem.
|
|
410 |
|
|
411 |
build_repabs_term does:
|
301
|
412 |
|
|
413 |
For abstractions:
|
|
414 |
* If the type of the abstraction doesn't need lifting we recurse.
|
|
415 |
* If it does we add RepAbs around the whole term and check if the
|
|
416 |
variable needs lifting.
|
|
417 |
* If it doesn't then we recurse
|
|
418 |
* If it does we recurse and put 'RepAbs' around all occurences
|
303
|
419 |
of the variable in the obtained subterm. This in combination
|
|
420 |
with the RepAbs above will let us change the type of the
|
|
421 |
abstraction with rewriting.
|
301
|
422 |
For applications:
|
|
423 |
* If the term is 'Respects' applied to anything we leave it unchanged
|
|
424 |
* If the term needs lifting and the head is a constant that we know
|
|
425 |
how to lift, we put a RepAbs and recurse
|
|
426 |
* If the term needs lifting and the head is a free applied to subterms
|
|
427 |
(if it is not applied we treated it in Abs branch) then we
|
|
428 |
put RepAbs and recurse
|
|
429 |
* Otherwise just recurse.
|
|
430 |
|
|
431 |
|
|
432 |
To prove that the old theorem implies the new one, we first
|
|
433 |
atomize it and then try:
|
|
434 |
|
303
|
435 |
1) theorems 'trans2' from the appropriate QUOT_TYPE
|
301
|
436 |
2) remove lambdas from both sides (LAMBDA_RES_TAC)
|
303
|
437 |
3) remove Ball/Bex from the right hand side
|
|
438 |
4) use user-supplied RSP theorems
|
|
439 |
5) remove rep_abs from the right side
|
|
440 |
6) reflexivity of equality
|
301
|
441 |
7) split applications of lifted type (apply_rsp)
|
|
442 |
8) split applications of non-lifted type (cong_tac)
|
|
443 |
9) apply extentionality
|
303
|
444 |
10) reflexivity of the relation
|
301
|
445 |
11) assumption
|
303
|
446 |
(Lambdas under respects may have left us some assumptions)
|
301
|
447 |
12) proving obvious higher order equalities by simplifying fun_rel
|
303
|
448 |
(not sure if it is still needed?)
|
301
|
449 |
13) unfolding lambda on one side
|
|
450 |
14) simplifying (= ===> =) for simpler respectfullness
|
|
451 |
|
|
452 |
*)
|
|
453 |
|
139
|
454 |
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
455 |
(* changes (?'a ?'b raw) (?'a ?'b quo) (int 'b raw \<Rightarrow> bool) to (int 'b quo \<Rightarrow> bool) *)
|
139
|
456 |
ML {*
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
457 |
fun exchange_ty lthy rty qty ty =
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
458 |
let
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
459 |
val thy = ProofContext.theory_of lthy
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
460 |
in
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
461 |
if Sign.typ_instance thy (ty, rty) then
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
462 |
let
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
463 |
val inst = Sign.typ_match thy (rty, ty) Vartab.empty
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
464 |
in
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
465 |
Envir.subst_type inst qty
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
466 |
end
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
467 |
else
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
468 |
let
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
469 |
val (s, tys) = dest_Type ty
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
470 |
in
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
471 |
Type (s, map (exchange_ty lthy rty qty) tys)
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
472 |
end
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
473 |
end
|
302
|
474 |
handle TYPE _ => ty (* for dest_Type *)
|
218
|
475 |
*}
|
|
476 |
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
477 |
(*consts Rl :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
478 |
axioms Rl_eq: "EQUIV Rl"
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
479 |
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
480 |
quotient ql = "'a list" / "Rl"
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
481 |
by (rule Rl_eq)
|
302
|
482 |
ML {*
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
483 |
ctyp_of @{theory} (exchange_ty @{context} (Logic.varifyT @{typ "'a list"}) (Logic.varifyT @{typ "'a ql"}) @{typ "nat list \<Rightarrow> (nat \<times> nat) list"});
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
484 |
ctyp_of @{theory} (exchange_ty @{context} (Logic.varifyT @{typ "nat \<times> nat"}) (Logic.varifyT @{typ "int"}) @{typ "nat list \<Rightarrow> (nat \<times> nat) list"})
|
285
|
485 |
*}
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
486 |
*)
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
487 |
|
285
|
488 |
ML {*
|
|
489 |
fun find_matching_types rty ty =
|
300
|
490 |
if Type.raw_instance (Logic.varifyT ty, rty)
|
|
491 |
then [ty]
|
|
492 |
else
|
|
493 |
let val (s, tys) = dest_Type ty in
|
|
494 |
flat (map (find_matching_types rty) tys)
|
|
495 |
end
|
302
|
496 |
handle TYPE _ => []
|
285
|
497 |
*}
|
|
498 |
|
|
499 |
ML {*
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
500 |
fun negF absF = repF
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
501 |
| negF repF = absF
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
502 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
503 |
fun get_fun flag qenv lthy ty =
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
504 |
let
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
505 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
506 |
fun get_fun_aux s fs =
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
507 |
(case (maps_lookup (ProofContext.theory_of lthy) s) of
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
508 |
SOME info => list_comb (Const (#mapfun info, dummyT), fs)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
509 |
| NONE => error ("no map association for type " ^ s))
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
510 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
511 |
fun get_const flag qty =
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
512 |
let
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
513 |
val thy = ProofContext.theory_of lthy
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
514 |
val qty_name = Long_Name.base_name (fst (dest_Type qty))
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
515 |
in
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
516 |
case flag of
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
517 |
absF => Const (Sign.full_bname thy ("ABS_" ^ qty_name), dummyT)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
518 |
| repF => Const (Sign.full_bname thy ("REP_" ^ qty_name), dummyT)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
519 |
end
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
520 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
521 |
fun mk_identity ty = Abs ("", ty, Bound 0)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
522 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
523 |
in
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
524 |
if (AList.defined (op=) qenv ty)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
525 |
then (get_const flag ty)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
526 |
else (case ty of
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
527 |
TFree _ => mk_identity ty
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
528 |
| Type (_, []) => mk_identity ty
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
529 |
| Type ("fun" , [ty1, ty2]) =>
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
530 |
let
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
531 |
val fs_ty1 = get_fun (negF flag) qenv lthy ty1
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
532 |
val fs_ty2 = get_fun flag qenv lthy ty2
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
533 |
in
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
534 |
get_fun_aux "fun" [fs_ty1, fs_ty2]
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
535 |
end
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
536 |
| Type (s, tys) => get_fun_aux s (map (get_fun flag qenv lthy) tys)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
537 |
| _ => error ("no type variables allowed"))
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
538 |
end
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
539 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
540 |
(* returns all subterms where two types differ *)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
541 |
fun diff (T, S) Ds =
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
542 |
case (T, S) of
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
543 |
(TVar v, TVar u) => if v = u then Ds else (T, S)::Ds
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
544 |
| (TFree x, TFree y) => if x = y then Ds else (T, S)::Ds
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
545 |
| (Type (a, Ts), Type (b, Us)) =>
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
546 |
if a = b then diffs (Ts, Us) Ds else (T, S)::Ds
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
547 |
| _ => (T, S)::Ds
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
548 |
and diffs (T::Ts, U::Us) Ds = diffs (Ts, Us) (diff (T, U) Ds)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
549 |
| diffs ([], []) Ds = Ds
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
550 |
| diffs _ _ = error "Unequal length of type arguments"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
551 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
552 |
*}
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
553 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
554 |
ML {*
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
555 |
fun get_fun_OLD flag (rty, qty) lthy ty =
|
285
|
556 |
let
|
|
557 |
val tys = find_matching_types rty ty;
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
558 |
val qenv = map (fn t => (exchange_ty lthy rty qty t, t)) tys;
|
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
559 |
val xchg_ty = exchange_ty lthy rty qty ty
|
285
|
560 |
in
|
|
561 |
get_fun flag qenv lthy xchg_ty
|
|
562 |
end
|
|
563 |
*}
|
|
564 |
|
235
|
565 |
text {* Does the same as 'subst' in a given prop or theorem *}
|
|
566 |
ML {*
|
|
567 |
fun eqsubst_prop ctxt thms t =
|
|
568 |
let
|
|
569 |
val goalstate = Goal.init (cterm_of (ProofContext.theory_of ctxt) t)
|
|
570 |
val a' = case (SINGLE (EqSubst.eqsubst_tac ctxt [0] thms 1) goalstate) of
|
|
571 |
NONE => error "eqsubst_prop"
|
|
572 |
| SOME th => cprem_of th 1
|
|
573 |
in term_of a' end
|
|
574 |
*}
|
|
575 |
|
|
576 |
ML {*
|
|
577 |
fun repeat_eqsubst_prop ctxt thms t =
|
|
578 |
repeat_eqsubst_prop ctxt thms (eqsubst_prop ctxt thms t)
|
|
579 |
handle _ => t
|
|
580 |
*}
|
|
581 |
|
|
582 |
|
|
583 |
ML {*
|
|
584 |
fun eqsubst_thm ctxt thms thm =
|
|
585 |
let
|
|
586 |
val goalstate = Goal.init (Thm.cprop_of thm)
|
|
587 |
val a' = case (SINGLE (EqSubst.eqsubst_tac ctxt [0] thms 1) goalstate) of
|
|
588 |
NONE => error "eqsubst_thm"
|
|
589 |
| SOME th => cprem_of th 1
|
|
590 |
val tac = (EqSubst.eqsubst_tac ctxt [0] thms 1) THEN simp_tac HOL_ss 1
|
302
|
591 |
val goal = Logic.mk_equals (term_of (Thm.cprop_of thm), term_of a');
|
|
592 |
val cgoal = cterm_of (ProofContext.theory_of ctxt) goal
|
|
593 |
val rt = Goal.prove_internal [] cgoal (fn _ => tac);
|
235
|
594 |
in
|
301
|
595 |
@{thm Pure.equal_elim_rule1} OF [rt, thm]
|
235
|
596 |
end
|
|
597 |
*}
|
|
598 |
|
|
599 |
ML {*
|
|
600 |
fun repeat_eqsubst_thm ctxt thms thm =
|
|
601 |
repeat_eqsubst_thm ctxt thms (eqsubst_thm ctxt thms thm)
|
|
602 |
handle _ => thm
|
|
603 |
*}
|
|
604 |
|
302
|
605 |
(* Needed to have a meta-equality *)
|
|
606 |
lemma id_def_sym: "(\<lambda>x. x) \<equiv> id"
|
|
607 |
by (simp add: id_def)
|
|
608 |
|
303
|
609 |
(* TODO: can be also obtained with: *)
|
|
610 |
ML {* symmetric (eq_reflection OF @{thms id_def}) *}
|
|
611 |
|
218
|
612 |
ML {*
|
301
|
613 |
fun build_repabs_term lthy thm consts rty qty =
|
139
|
614 |
let
|
303
|
615 |
(* TODO: The rty and qty stored in the quotient_info should
|
|
616 |
be varified, so this will soon not be needed *)
|
285
|
617 |
val rty = Logic.varifyT rty;
|
|
618 |
val qty = Logic.varifyT qty;
|
139
|
619 |
|
285
|
620 |
fun mk_abs tm =
|
|
621 |
let
|
|
622 |
val ty = fastype_of tm
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
623 |
in Syntax.check_term lthy ((get_fun_OLD absF (rty, qty) lthy ty) $ tm) end
|
285
|
624 |
fun mk_repabs tm =
|
|
625 |
let
|
|
626 |
val ty = fastype_of tm
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
627 |
in Syntax.check_term lthy ((get_fun_OLD repF (rty, qty) lthy ty) $ (mk_abs tm)) end
|
139
|
628 |
|
301
|
629 |
fun is_lifted_const (Const (x, _)) = member (op =) consts x
|
|
630 |
| is_lifted_const _ = false;
|
139
|
631 |
|
|
632 |
fun build_aux lthy tm =
|
|
633 |
case tm of
|
301
|
634 |
Abs (a as (_, vty, _)) =>
|
139
|
635 |
let
|
301
|
636 |
val (vs, t) = Term.dest_abs a;
|
|
637 |
val v = Free(vs, vty);
|
|
638 |
val t' = lambda v (build_aux lthy t)
|
139
|
639 |
in
|
301
|
640 |
if (not (needs_lift rty (fastype_of tm))) then t'
|
|
641 |
else mk_repabs (
|
|
642 |
if not (needs_lift rty vty) then t'
|
|
643 |
else
|
|
644 |
let
|
|
645 |
val v' = mk_repabs v;
|
303
|
646 |
(* TODO: I believe 'beta' is not needed any more *)
|
|
647 |
val t1 = (* Envir.beta_norm *) (t' $ v')
|
301
|
648 |
in
|
|
649 |
lambda v t1
|
|
650 |
end)
|
139
|
651 |
end
|
301
|
652 |
| x =>
|
|
653 |
case Term.strip_comb tm of
|
|
654 |
(Const(@{const_name Respects}, _), _) => tm
|
|
655 |
| (opp, tms0) =>
|
|
656 |
let
|
|
657 |
val tms = map (build_aux lthy) tms0
|
|
658 |
val ty = fastype_of tm
|
|
659 |
in
|
|
660 |
if (is_lifted_const opp andalso needs_lift rty ty) then
|
|
661 |
mk_repabs (list_comb (opp, tms))
|
|
662 |
else if ((Term.is_Free opp) andalso (length tms > 0) andalso (needs_lift rty ty)) then
|
|
663 |
mk_repabs (list_comb (opp, tms))
|
|
664 |
else if tms = [] then opp
|
|
665 |
else list_comb(opp, tms)
|
|
666 |
end
|
139
|
667 |
in
|
235
|
668 |
repeat_eqsubst_prop lthy @{thms id_def_sym}
|
161
|
669 |
(build_aux lthy (Thm.prop_of thm))
|
139
|
670 |
end
|
|
671 |
*}
|
|
672 |
|
303
|
673 |
text {* Builds provable goals for regularized theorems *}
|
139
|
674 |
ML {*
|
140
|
675 |
fun build_repabs_goal ctxt thm cons rty qty =
|
|
676 |
Logic.mk_equals ((Thm.prop_of thm), (build_repabs_term ctxt thm cons rty qty))
|
139
|
677 |
*}
|
|
678 |
|
187
|
679 |
ML {*
|
|
680 |
fun instantiate_tac thm = Subgoal.FOCUS (fn {concl, ...} =>
|
303
|
681 |
let
|
|
682 |
val pat = Drule.strip_imp_concl (cprop_of thm)
|
|
683 |
val insts = Thm.match (pat, concl)
|
|
684 |
in
|
|
685 |
rtac (Drule.instantiate insts thm) 1
|
|
686 |
end
|
|
687 |
handle _ => no_tac)
|
|
688 |
*}
|
|
689 |
|
|
690 |
ML {*
|
|
691 |
fun CHANGED' tac = (fn i => CHANGED (tac i))
|
187
|
692 |
*}
|
|
693 |
|
359
|
694 |
lemma prod_fun_id: "prod_fun id id \<equiv> id"
|
|
695 |
by (rule eq_reflection) (simp add: prod_fun_def)
|
|
696 |
|
|
697 |
lemma map_id: "map id \<equiv> id"
|
|
698 |
apply (rule eq_reflection)
|
|
699 |
apply (rule ext)
|
|
700 |
apply (rule_tac list="x" in list.induct)
|
|
701 |
apply (simp_all)
|
|
702 |
done
|
|
703 |
|
187
|
704 |
ML {*
|
|
705 |
fun quotient_tac quot_thm =
|
|
706 |
REPEAT_ALL_NEW (FIRST' [
|
|
707 |
rtac @{thm FUN_QUOTIENT},
|
|
708 |
rtac quot_thm,
|
292
|
709 |
rtac @{thm IDENTITY_QUOTIENT},
|
303
|
710 |
(* For functional identity quotients, (op = ---> op =) *)
|
|
711 |
CHANGED' (
|
359
|
712 |
(simp_tac (HOL_ss addsimps @{thms eq_reflection[OF FUN_MAP_I] eq_reflection[OF id_apply] id_def_sym prod_fun_id map_id}
|
|
713 |
)))
|
187
|
714 |
])
|
|
715 |
*}
|
|
716 |
|
|
717 |
ML {*
|
|
718 |
fun LAMBDA_RES_TAC ctxt i st =
|
|
719 |
(case (term_of o #concl o fst) (Subgoal.focus ctxt i st) of
|
303
|
720 |
(_ $ (_ $ (Abs(_, _, _)) $ (Abs(_, _, _)))) =>
|
187
|
721 |
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
|
|
722 |
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
|
|
723 |
| _ => fn _ => no_tac) i st
|
|
724 |
*}
|
|
725 |
|
|
726 |
ML {*
|
|
727 |
fun WEAK_LAMBDA_RES_TAC ctxt i st =
|
|
728 |
(case (term_of o #concl o fst) (Subgoal.focus ctxt i st) of
|
301
|
729 |
(_ $ (_ $ _ $ (Abs(_, _, _)))) =>
|
187
|
730 |
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
|
|
731 |
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
|
301
|
732 |
| (_ $ (_ $ (Abs(_, _, _)) $ _)) =>
|
187
|
733 |
(EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
|
|
734 |
(rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
|
|
735 |
| _ => fn _ => no_tac) i st
|
|
736 |
*}
|
|
737 |
|
206
|
738 |
ML {*
|
|
739 |
fun APPLY_RSP_TAC rty = Subgoal.FOCUS (fn {concl, ...} =>
|
|
740 |
let
|
|
741 |
val (_ $ (R $ (f $ _) $ (_ $ _))) = term_of concl;
|
|
742 |
val pat = Drule.strip_imp_concl (cprop_of @{thm APPLY_RSP});
|
|
743 |
val insts = Thm.match (pat, concl)
|
303
|
744 |
in
|
|
745 |
if needs_lift rty (type_of f) then
|
|
746 |
rtac (Drule.instantiate insts @{thm APPLY_RSP}) 1
|
|
747 |
else no_tac
|
|
748 |
end
|
|
749 |
handle _ => no_tac)
|
206
|
750 |
*}
|
187
|
751 |
|
|
752 |
ML {*
|
301
|
753 |
val ball_rsp_tac = Subgoal.FOCUS (fn {concl, context = ctxt, ...} =>
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
754 |
let
|
301
|
755 |
val _ $ (_ $ (Const (@{const_name Ball}, _) $ _) $
|
|
756 |
(Const (@{const_name Ball}, _) $ _)) = term_of concl
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
757 |
in
|
302
|
758 |
((simp_tac (HOL_ss addsimps @{thms FUN_REL.simps}))
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
759 |
THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
760 |
THEN' instantiate_tac @{thm RES_FORALL_RSP} ctxt THEN'
|
302
|
761 |
(simp_tac (HOL_ss addsimps @{thms FUN_REL.simps}))) 1
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
762 |
end
|
303
|
763 |
handle _ => no_tac)
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
764 |
*}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
765 |
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
766 |
ML {*
|
301
|
767 |
val bex_rsp_tac = Subgoal.FOCUS (fn {concl, context = ctxt, ...} =>
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
768 |
let
|
301
|
769 |
val _ $ (_ $ (Const (@{const_name Bex}, _) $ _) $
|
|
770 |
(Const (@{const_name Bex}, _) $ _)) = term_of concl
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
771 |
in
|
302
|
772 |
((simp_tac (HOL_ss addsimps @{thms FUN_REL.simps}))
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
773 |
THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
774 |
THEN' instantiate_tac @{thm RES_EXISTS_RSP} ctxt THEN'
|
302
|
775 |
(simp_tac (HOL_ss addsimps @{thms FUN_REL.simps}))) 1
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
776 |
end
|
303
|
777 |
handle _ => no_tac)
|
|
778 |
*}
|
|
779 |
|
|
780 |
ML {*
|
|
781 |
fun SOLVES' tac = tac THEN_ALL_NEW (fn _ => no_tac)
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
782 |
*}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
783 |
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
784 |
ML {*
|
206
|
785 |
fun r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms =
|
187
|
786 |
(FIRST' [
|
|
787 |
rtac trans_thm,
|
|
788 |
LAMBDA_RES_TAC ctxt,
|
301
|
789 |
ball_rsp_tac ctxt,
|
|
790 |
bex_rsp_tac ctxt,
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
791 |
FIRST' (map rtac rsp_thms),
|
187
|
792 |
(instantiate_tac @{thm REP_ABS_RSP(1)} ctxt THEN' (RANGE [quotient_tac quot_thm])),
|
|
793 |
rtac refl,
|
206
|
794 |
(APPLY_RSP_TAC rty ctxt THEN' (RANGE [quotient_tac quot_thm, quotient_tac quot_thm])),
|
|
795 |
Cong_Tac.cong_tac @{thm cong},
|
|
796 |
rtac @{thm ext},
|
187
|
797 |
rtac reflex_thm,
|
|
798 |
atac,
|
303
|
799 |
SOLVES' (simp_tac (HOL_ss addsimps @{thms FUN_REL.simps})),
|
292
|
800 |
WEAK_LAMBDA_RES_TAC ctxt,
|
303
|
801 |
CHANGED' (asm_full_simp_tac (HOL_ss addsimps @{thms FUN_REL_EQ}))
|
187
|
802 |
])
|
|
803 |
*}
|
|
804 |
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
805 |
ML {*
|
301
|
806 |
fun repabs lthy thm consts rty qty quot_thm reflex_thm trans_thm rsp_thms =
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
807 |
let
|
301
|
808 |
val rt = build_repabs_term lthy thm consts rty qty;
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
809 |
val rg = Logic.mk_equals ((Thm.prop_of thm), rt);
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
810 |
fun tac ctxt = (ObjectLogic.full_atomize_tac) THEN'
|
206
|
811 |
(REPEAT_ALL_NEW (r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms));
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
812 |
val cthm = Goal.prove lthy [] [] rg (fn x => tac (#context x) 1);
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
813 |
in
|
210
|
814 |
@{thm Pure.equal_elim_rule1} OF [cthm, thm]
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
815 |
end
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
816 |
*}
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
817 |
|
187
|
818 |
section {* Cleaning the goal *}
|
|
819 |
|
236
|
820 |
|
292
|
821 |
ML {*
|
|
822 |
fun simp_ids lthy thm =
|
307
|
823 |
MetaSimplifier.rewrite_rule @{thms eq_reflection[OF FUN_MAP_I] eq_reflection[OF id_apply] id_def_sym prod_fun_id map_id} thm
|
|
824 |
*}
|
|
825 |
|
|
826 |
ML {*
|
|
827 |
fun simp_ids_trm trm =
|
|
828 |
trm |>
|
|
829 |
MetaSimplifier.rewrite false @{thms eq_reflection[OF FUN_MAP_I] eq_reflection[OF id_apply] id_def_sym prod_fun_id map_id}
|
|
830 |
|> cprop_of |> Thm.dest_equals |> snd
|
|
831 |
|
292
|
832 |
*}
|
|
833 |
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
834 |
text {* expects atomized definition *}
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
835 |
ML {*
|
301
|
836 |
fun add_lower_defs_aux lthy thm =
|
|
837 |
let
|
|
838 |
val e1 = @{thm fun_cong} OF [thm];
|
|
839 |
val f = eqsubst_thm lthy @{thms fun_map.simps} e1;
|
|
840 |
val g = simp_ids lthy f
|
|
841 |
in
|
|
842 |
(simp_ids lthy thm) :: (add_lower_defs_aux lthy g)
|
|
843 |
end
|
|
844 |
handle _ => [simp_ids lthy thm]
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
845 |
*}
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
846 |
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
847 |
ML {*
|
292
|
848 |
fun add_lower_defs lthy def =
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
849 |
let
|
292
|
850 |
val def_pre_sym = symmetric def
|
|
851 |
val def_atom = atomize_thm def_pre_sym
|
|
852 |
val defs_all = add_lower_defs_aux lthy def_atom
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
853 |
in
|
214
|
854 |
map Thm.varifyT defs_all
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
855 |
end
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
856 |
*}
|
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
857 |
|
303
|
858 |
(* TODO: Check if it behaves properly with varifyed rty *)
|
191
|
859 |
ML {*
|
301
|
860 |
fun findabs_all rty tm =
|
|
861 |
case tm of
|
|
862 |
Abs(_, T, b) =>
|
|
863 |
let
|
|
864 |
val b' = subst_bound ((Free ("x", T)), b);
|
|
865 |
val tys = findabs_all rty b'
|
|
866 |
val ty = fastype_of tm
|
|
867 |
in if needs_lift rty ty then (ty :: tys) else tys
|
|
868 |
end
|
|
869 |
| f $ a => (findabs_all rty f) @ (findabs_all rty a)
|
|
870 |
| _ => [];
|
|
871 |
fun findabs rty tm = distinct (op =) (findabs_all rty tm)
|
241
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
872 |
*}
|
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
873 |
|
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
874 |
ML {*
|
301
|
875 |
fun findaps_all rty tm =
|
|
876 |
case tm of
|
|
877 |
Abs(_, T, b) =>
|
|
878 |
findaps_all rty (subst_bound ((Free ("x", T)), b))
|
|
879 |
| (f $ a) => (findaps_all rty f @ findaps_all rty a)
|
|
880 |
| Free (_, (T as (Type ("fun", (_ :: _))))) =>
|
|
881 |
(if needs_lift rty T then [T] else [])
|
|
882 |
| _ => [];
|
|
883 |
fun findaps rty tm = distinct (op =) (findaps_all rty tm)
|
191
|
884 |
*}
|
190
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
885 |
|
303
|
886 |
(* Currently useful only for LAMBDA_PRS *)
|
197
|
887 |
ML {*
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
888 |
fun make_simp_prs_thm lthy quot_thm thm typ =
|
197
|
889 |
let
|
|
890 |
val (_, [lty, rty]) = dest_Type typ;
|
|
891 |
val thy = ProofContext.theory_of lthy;
|
|
892 |
val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty)
|
|
893 |
val inst = [SOME lcty, NONE, SOME rcty];
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
894 |
val lpi = Drule.instantiate' inst [] thm;
|
197
|
895 |
val tac =
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
896 |
(compose_tac (false, lpi, 2)) THEN_ALL_NEW
|
197
|
897 |
(quotient_tac quot_thm);
|
259
|
898 |
val gc = Drule.strip_imp_concl (cprop_of lpi);
|
|
899 |
val t = Goal.prove_internal [] gc (fn _ => tac 1)
|
197
|
900 |
in
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
901 |
MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] t
|
197
|
902 |
end
|
|
903 |
*}
|
|
904 |
|
|
905 |
ML {*
|
301
|
906 |
fun findallex_all rty qty tm =
|
|
907 |
case tm of
|
|
908 |
Const (@{const_name All}, T) $ (s as (Abs(_, _, b))) =>
|
|
909 |
let
|
|
910 |
val (tya, tye) = findallex_all rty qty s
|
|
911 |
in if needs_lift rty T then
|
|
912 |
((T :: tya), tye)
|
|
913 |
else (tya, tye) end
|
|
914 |
| Const (@{const_name Ex}, T) $ (s as (Abs(_, _, b))) =>
|
|
915 |
let
|
|
916 |
val (tya, tye) = findallex_all rty qty s
|
|
917 |
in if needs_lift rty T then
|
|
918 |
(tya, (T :: tye))
|
|
919 |
else (tya, tye) end
|
|
920 |
| Abs(_, T, b) =>
|
|
921 |
findallex_all rty qty (subst_bound ((Free ("x", T)), b))
|
|
922 |
| f $ a =>
|
|
923 |
let
|
|
924 |
val (a1, e1) = findallex_all rty qty f;
|
|
925 |
val (a2, e2) = findallex_all rty qty a;
|
|
926 |
in (a1 @ a2, e1 @ e2) end
|
|
927 |
| _ => ([], []);
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
928 |
*}
|
303
|
929 |
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
930 |
ML {*
|
301
|
931 |
fun findallex lthy rty qty tm =
|
|
932 |
let
|
|
933 |
val (a, e) = findallex_all rty qty tm;
|
|
934 |
val (ad, ed) = (map domain_type a, map domain_type e);
|
|
935 |
val (au, eu) = (distinct (op =) ad, distinct (op =) ed);
|
|
936 |
val (rty, qty) = (Logic.varifyT rty, Logic.varifyT qty)
|
|
937 |
in
|
|
938 |
(map (exchange_ty lthy rty qty) au, map (exchange_ty lthy rty qty) eu)
|
|
939 |
end
|
197
|
940 |
*}
|
|
941 |
|
198
|
942 |
ML {*
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
943 |
fun make_allex_prs_thm lthy quot_thm thm typ =
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
944 |
let
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
945 |
val (_, [lty, rty]) = dest_Type typ;
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
946 |
val thy = ProofContext.theory_of lthy;
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
947 |
val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty)
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
948 |
val inst = [NONE, SOME lcty];
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
949 |
val lpi = Drule.instantiate' inst [] thm;
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
950 |
val tac =
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
951 |
(compose_tac (false, lpi, 1)) THEN_ALL_NEW
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
952 |
(quotient_tac quot_thm);
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
953 |
val gc = Drule.strip_imp_concl (cprop_of lpi);
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
954 |
val t = Goal.prove_internal [] gc (fn _ => tac 1)
|
301
|
955 |
val t_noid = MetaSimplifier.rewrite_rule
|
|
956 |
[@{thm eq_reflection} OF @{thms id_apply}] t;
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
957 |
val t_sym = @{thm "HOL.sym"} OF [t_noid];
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
958 |
val t_eq = @{thm "eq_reflection"} OF [t_sym]
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
959 |
in
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
960 |
t_eq
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
961 |
end
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
962 |
*}
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
963 |
|
269
|
964 |
ML {*
|
|
965 |
fun applic_prs lthy rty qty absrep ty =
|
301
|
966 |
let
|
285
|
967 |
val rty = Logic.varifyT rty;
|
|
968 |
val qty = Logic.varifyT qty;
|
301
|
969 |
fun absty ty =
|
|
970 |
exchange_ty lthy rty qty ty
|
|
971 |
fun mk_rep tm =
|
|
972 |
let
|
|
973 |
val ty = exchange_ty lthy qty rty (fastype_of tm)
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
974 |
in Syntax.check_term lthy ((get_fun_OLD repF (rty, qty) lthy ty) $ tm) end;
|
301
|
975 |
fun mk_abs tm =
|
|
976 |
let
|
|
977 |
val ty = fastype_of tm
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
978 |
in Syntax.check_term lthy ((get_fun_OLD absF (rty, qty) lthy ty) $ tm) end
|
301
|
979 |
val (l, ltl) = Term.strip_type ty;
|
|
980 |
val nl = map absty l;
|
|
981 |
val vs = map (fn _ => "x") l;
|
|
982 |
val ((fname :: vfs), lthy') = Variable.variant_fixes ("f" :: vs) lthy;
|
|
983 |
val args = map Free (vfs ~~ nl);
|
|
984 |
val lhs = list_comb((Free (fname, nl ---> ltl)), args);
|
|
985 |
val rargs = map mk_rep args;
|
|
986 |
val f = Free (fname, nl ---> ltl);
|
|
987 |
val rhs = mk_abs (list_comb((mk_rep f), rargs));
|
|
988 |
val eq = Logic.mk_equals (rhs, lhs);
|
|
989 |
val ceq = cterm_of (ProofContext.theory_of lthy') eq;
|
302
|
990 |
val sctxt = HOL_ss addsimps (absrep :: @{thms fun_map.simps});
|
301
|
991 |
val t = Goal.prove_internal [] ceq (fn _ => simp_tac sctxt 1)
|
|
992 |
val t_id = MetaSimplifier.rewrite_rule @{thms id_def_sym} t;
|
|
993 |
in
|
|
994 |
singleton (ProofContext.export lthy' lthy) t_id
|
|
995 |
end
|
269
|
996 |
*}
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
997 |
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
998 |
ML {*
|
239
|
999 |
fun lookup_quot_data lthy qty =
|
|
1000 |
let
|
311
|
1001 |
val qty_name = fst (dest_Type qty)
|
314
|
1002 |
val SOME quotdata = quotdata_lookup lthy qty_name
|
311
|
1003 |
(* cu: Changed the lookup\<dots>not sure whether this works *)
|
303
|
1004 |
(* TODO: Should no longer be needed *)
|
257
|
1005 |
val rty = Logic.unvarifyT (#rtyp quotdata)
|
239
|
1006 |
val rel = #rel quotdata
|
|
1007 |
val rel_eqv = #equiv_thm quotdata
|
|
1008 |
val rel_refl_pre = @{thm EQUIV_REFL} OF [rel_eqv]
|
|
1009 |
val rel_refl = @{thm spec} OF [MetaSimplifier.rewrite_rule [@{thm REFL_def}] rel_refl_pre]
|
|
1010 |
in
|
|
1011 |
(rty, rel, rel_refl, rel_eqv)
|
|
1012 |
end
|
|
1013 |
*}
|
|
1014 |
|
|
1015 |
ML {*
|
|
1016 |
fun lookup_quot_thms lthy qty_name =
|
|
1017 |
let
|
|
1018 |
val thy = ProofContext.theory_of lthy;
|
|
1019 |
val trans2 = PureThy.get_thm thy ("QUOT_TYPE_I_" ^ qty_name ^ ".R_trans2")
|
|
1020 |
val reps_same = PureThy.get_thm thy ("QUOT_TYPE_I_" ^ qty_name ^ ".REPS_same")
|
269
|
1021 |
val absrep = PureThy.get_thm thy ("QUOT_TYPE_I_" ^ qty_name ^ ".thm10")
|
239
|
1022 |
val quot = PureThy.get_thm thy ("QUOTIENT_" ^ qty_name)
|
|
1023 |
in
|
269
|
1024 |
(trans2, reps_same, absrep, quot)
|
239
|
1025 |
end
|
|
1026 |
*}
|
|
1027 |
|
|
1028 |
ML {*
|
|
1029 |
fun lookup_quot_consts defs =
|
|
1030 |
let
|
|
1031 |
fun dest_term (a $ b) = (a, b);
|
|
1032 |
val def_terms = map (snd o Logic.dest_equals o concl_of) defs;
|
|
1033 |
in
|
|
1034 |
map (fst o dest_Const o snd o dest_term) def_terms
|
|
1035 |
end
|
|
1036 |
*}
|
|
1037 |
|
275
|
1038 |
|
239
|
1039 |
ML {*
|
319
|
1040 |
fun lift_thm lthy qty qty_name rsp_thms defs rthm =
|
|
1041 |
let
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1042 |
val _ = tracing ("raw theorem:\n" ^ Syntax.string_of_term lthy (prop_of rthm))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1043 |
|
239
|
1044 |
val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data lthy qty;
|
269
|
1045 |
val (trans2, reps_same, absrep, quot) = lookup_quot_thms lthy qty_name;
|
239
|
1046 |
val consts = lookup_quot_consts defs;
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1047 |
val t_a = atomize_thm rthm;
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1048 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1049 |
val _ = tracing ("raw atomized theorem:\n" ^ Syntax.string_of_term lthy (prop_of t_a))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1050 |
|
251
|
1051 |
val t_r = regularize t_a rty rel rel_eqv rel_refl lthy;
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1052 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1053 |
val _ = tracing ("regularised theorem:\n" ^ Syntax.string_of_term lthy (prop_of t_r))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1054 |
|
210
|
1055 |
val t_t = repabs lthy t_r consts rty qty quot rel_refl trans2 rsp_thms;
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1056 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1057 |
val _ = tracing ("rep/abs injected theorem:\n" ^ Syntax.string_of_term lthy (prop_of t_t))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1058 |
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
1059 |
val (alls, exs) = findallex lthy rty qty (prop_of t_a);
|
267
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
1060 |
val allthms = map (make_allex_prs_thm lthy quot @{thm FORALL_PRS}) alls
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
1061 |
val exthms = map (make_allex_prs_thm lthy quot @{thm EXISTS_PRS}) exs
|
3764566c1151
Automatic FORALL_PRS. 'list.induct' lifts automatically. Faster ALLEX_RSP
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
1062 |
val t_a = MetaSimplifier.rewrite_rule (allthms @ exthms) t_t
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1063 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1064 |
val _ = tracing ("??:\n" ^ Syntax.string_of_term lthy (prop_of t_a))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1065 |
|
198
|
1066 |
val abs = findabs rty (prop_of t_a);
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
1067 |
val aps = findaps rty (prop_of t_a);
|
269
|
1068 |
val app_prs_thms = map (applic_prs lthy rty qty absrep) aps;
|
253
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
1069 |
val lam_prs_thms = map (make_simp_prs_thm lthy quot @{thm LAMBDA_PRS}) abs;
|
259
|
1070 |
val t_l = repeat_eqsubst_thm lthy (lam_prs_thms @ app_prs_thms) t_a;
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1071 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1072 |
val _ = tracing ("??:\n" ^ Syntax.string_of_term lthy (prop_of t_l))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1073 |
|
292
|
1074 |
val defs_sym = flat (map (add_lower_defs lthy) defs);
|
259
|
1075 |
val defs_sym_eq = map (fn x => eq_reflection OF [x]) defs_sym;
|
292
|
1076 |
val t_id = simp_ids lthy t_l;
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1077 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1078 |
val _ = tracing ("??:\n" ^ Syntax.string_of_term lthy (prop_of t_id))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1079 |
|
292
|
1080 |
val t_d0 = MetaSimplifier.rewrite_rule defs_sym_eq t_id;
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1081 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1082 |
val _ = tracing ("??:\n" ^ Syntax.string_of_term lthy (prop_of t_d0))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1083 |
|
259
|
1084 |
val t_d = repeat_eqsubst_thm lthy defs_sym t_d0;
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1085 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1086 |
val _ = tracing ("??:\n" ^ Syntax.string_of_term lthy (prop_of t_d))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1087 |
|
198
|
1088 |
val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d;
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1089 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1090 |
val _ = tracing ("??:\n" ^ Syntax.string_of_term lthy (prop_of t_r))
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1091 |
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
1092 |
val t_rv = ObjectLogic.rulify t_r
|
316
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1093 |
|
13ea9a34c269
added some tracing information to all phases of lifting to the function lift_thm
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1094 |
val _ = tracing ("lifted theorem:\n" ^ Syntax.string_of_term lthy (prop_of t_rv))
|
198
|
1095 |
in
|
289
7e8617f20b59
Remaining fixes for polymorphic types. map_append now lifts properly with 'a list and 'b list.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
1096 |
Thm.varifyT t_rv
|
187
|
1097 |
end
|
198
|
1098 |
*}
|
|
1099 |
|
273
|
1100 |
ML {*
|
301
|
1101 |
fun lift_thm_note qty qty_name rsp_thms defs thm name lthy =
|
|
1102 |
let
|
|
1103 |
val lifted_thm = lift_thm lthy qty qty_name rsp_thms defs thm;
|
|
1104 |
val (_, lthy2) = note (name, lifted_thm) lthy;
|
|
1105 |
in
|
|
1106 |
lthy2
|
|
1107 |
end
|
273
|
1108 |
*}
|
|
1109 |
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1110 |
(******************************************)
|
325
|
1111 |
(******************************************)
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1112 |
(* version with explicit qtrm *)
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1113 |
(******************************************)
|
325
|
1114 |
(******************************************)
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1115 |
|
319
|
1116 |
ML {*
|
338
|
1117 |
fun atomize_goal thy gl =
|
|
1118 |
let
|
|
1119 |
val vars = map Free (Term.add_frees gl []);
|
348
|
1120 |
val all = if fastype_of gl = @{typ bool} then HOLogic.all_const else Term.all;
|
|
1121 |
fun lambda_all (var as Free(_, T)) trm = (all T) $ lambda var trm;
|
|
1122 |
val glv = fold lambda_all vars gl
|
|
1123 |
val gla = (term_of o snd o Thm.dest_equals o cprop_of) (ObjectLogic.atomize (cterm_of thy glv))
|
338
|
1124 |
val glf = Type.legacy_freeze gla
|
|
1125 |
in
|
348
|
1126 |
if fastype_of gl = @{typ bool} then @{term Trueprop} $ glf else glf
|
338
|
1127 |
end
|
|
1128 |
*}
|
|
1129 |
|
|
1130 |
|
348
|
1131 |
ML {* atomize_goal @{theory} @{term "x memb [] = False"} *}
|
|
1132 |
ML {* atomize_goal @{theory} @{term "x = xa \<Longrightarrow> a # x = a # xa"} *}
|
|
1133 |
|
|
1134 |
|
338
|
1135 |
ML {*
|
330
|
1136 |
(* builds the relation for respects *)
|
|
1137 |
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1138 |
fun mk_resp_arg lthy (rty, qty) =
|
319
|
1139 |
let
|
|
1140 |
val thy = ProofContext.theory_of lthy
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1141 |
in
|
334
|
1142 |
if rty = qty
|
|
1143 |
then HOLogic.eq_const rty
|
|
1144 |
else
|
|
1145 |
case (rty, qty) of
|
|
1146 |
(Type (s, tys), Type (s', tys')) =>
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1147 |
if s = s'
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1148 |
then let
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1149 |
val SOME map_info = maps_lookup thy s
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1150 |
val args = map (mk_resp_arg lthy) (tys ~~ tys')
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1151 |
in
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1152 |
list_comb (Const (#relfun map_info, dummyT), args)
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1153 |
end
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1154 |
else let
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1155 |
val SOME qinfo = quotdata_lookup_thy thy s'
|
330
|
1156 |
(* FIXME: check in this case that the rty and qty *)
|
|
1157 |
(* FIXME: correspond to each other *)
|
353
9a0e8ab42ee8
fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1158 |
val (s, _) = dest_Const (#rel qinfo)
|
9a0e8ab42ee8
fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1159 |
(* FIXME: the relation should only be the string *)
|
9a0e8ab42ee8
fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1160 |
(* FIXME: and the type needs to be calculated as below *)
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1161 |
in
|
353
9a0e8ab42ee8
fixed the error by a temporary fix (the data of the eqivalence relation should be only its name)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1162 |
Const (s, rty --> rty --> @{typ bool})
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1163 |
end
|
334
|
1164 |
| _ => HOLogic.eq_const dummyT
|
351
|
1165 |
(* FIXME: check that the types correspond to each other? *)
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1166 |
end
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1167 |
*}
|
239
|
1168 |
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1169 |
ML {*
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1170 |
val mk_babs = Const (@{const_name "Babs"}, dummyT)
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1171 |
val mk_ball = Const (@{const_name "Ball"}, dummyT)
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1172 |
val mk_bex = Const (@{const_name "Bex"}, dummyT)
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1173 |
val mk_resp = Const (@{const_name Respects}, dummyT)
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1174 |
*}
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1175 |
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1176 |
ML {*
|
323
|
1177 |
fun trm_lift_error lthy rtrm qtrm =
|
|
1178 |
let
|
|
1179 |
val rtrm_str = quote (Syntax.string_of_term lthy rtrm)
|
|
1180 |
val qtrm_str = quote (Syntax.string_of_term lthy qtrm)
|
325
|
1181 |
val msg = ["The quotient theorem", qtrm_str, "and lifted theorem", rtrm_str, "do not match."]
|
323
|
1182 |
in
|
365
|
1183 |
raise error (space_implode " " msg)
|
323
|
1184 |
end
|
|
1185 |
*}
|
|
1186 |
|
|
1187 |
ML {*
|
330
|
1188 |
(* - applies f to the subterm of an abstraction, *)
|
|
1189 |
(* otherwise to the given term, *)
|
351
|
1190 |
(* - used by REGULARIZE, therefore abstracted *)
|
330
|
1191 |
(* variables do not have to be treated specially *)
|
|
1192 |
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1193 |
fun apply_subt f trm1 trm2 =
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1194 |
case (trm1, trm2) of
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1195 |
(Abs (x, T, t), Abs (x', T', t')) => Abs (x, T, f t t')
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1196 |
| _ => f trm1 trm2
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1197 |
|
330
|
1198 |
(* the major type of All and Ex quantifiers *)
|
334
|
1199 |
fun qnt_typ ty = domain_type (domain_type ty)
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1200 |
*}
|
319
|
1201 |
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1202 |
(*
|
330
|
1203 |
Regularizing an rtrm means:
|
|
1204 |
- quantifiers over a type that needs lifting are replaced by
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1205 |
bounded quantifiers, for example:
|
330
|
1206 |
\<forall>x. P \<Longrightarrow> \<forall>x \<in> (Respects R). P / All (Respects R) P
|
|
1207 |
|
|
1208 |
the relation R is given by the rty and qty;
|
|
1209 |
|
|
1210 |
- abstractions over a type that needs lifting are replaced
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1211 |
by bounded abstractions:
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1212 |
\<lambda>x. P \<Longrightarrow> Ball (Respects R) (\<lambda>x. P)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1213 |
|
330
|
1214 |
- equalities over the type being lifted are replaced by
|
|
1215 |
corresponding relations:
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1216 |
A = B \<Longrightarrow> A \<approx> B
|
330
|
1217 |
|
|
1218 |
example with more complicated types of A, B:
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1219 |
A = B \<Longrightarrow> (op = \<Longrightarrow> op \<approx>) A B
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1220 |
*)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1221 |
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1222 |
ML {*
|
330
|
1223 |
(* produces a regularized version of rtm *)
|
|
1224 |
(* - the result is still not completely typed *)
|
|
1225 |
(* - does not need any special treatment of *)
|
|
1226 |
(* bound variables *)
|
|
1227 |
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1228 |
fun REGULARIZE_trm lthy rtrm qtrm =
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1229 |
case (rtrm, qtrm) of
|
325
|
1230 |
(Abs (x, ty, t), Abs (x', ty', t')) =>
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1231 |
let
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1232 |
val subtrm = REGULARIZE_trm lthy t t'
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1233 |
in
|
325
|
1234 |
if ty = ty'
|
|
1235 |
then Abs (x, ty, subtrm)
|
326
|
1236 |
else mk_babs $ (mk_resp $ mk_resp_arg lthy (ty, ty')) $ subtrm
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1237 |
end
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1238 |
| (Const (@{const_name "All"}, ty) $ t, Const (@{const_name "All"}, ty') $ t') =>
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1239 |
let
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1240 |
val subtrm = apply_subt (REGULARIZE_trm lthy) t t'
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1241 |
in
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1242 |
if ty = ty'
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1243 |
then Const (@{const_name "All"}, ty) $ subtrm
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1244 |
else mk_ball $ (mk_resp $ mk_resp_arg lthy (qnt_typ ty, qnt_typ ty')) $ subtrm
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1245 |
end
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1246 |
| (Const (@{const_name "Ex"}, ty) $ t, Const (@{const_name "Ex"}, ty') $ t') =>
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1247 |
let
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1248 |
val subtrm = apply_subt (REGULARIZE_trm lthy) t t'
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1249 |
in
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1250 |
if ty = ty'
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1251 |
then Const (@{const_name "Ex"}, ty) $ subtrm
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1252 |
else mk_bex $ (mk_resp $ mk_resp_arg lthy (qnt_typ ty, qnt_typ ty')) $ subtrm
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1253 |
end
|
351
|
1254 |
(* FIXME: Should = only be replaced, when fully applied? *)
|
|
1255 |
(* Then there must be a 2nd argument *)
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1256 |
| (Const (@{const_name "op ="}, ty) $ t, Const (@{const_name "op ="}, ty') $ t') =>
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1257 |
let
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1258 |
val subtrm = REGULARIZE_trm lthy t t'
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1259 |
in
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1260 |
if ty = ty'
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1261 |
then Const (@{const_name "op ="}, ty) $ subtrm
|
349
|
1262 |
else mk_resp_arg lthy (domain_type ty, domain_type ty') $ subtrm
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1263 |
end
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1264 |
| (t1 $ t2, t1' $ t2') =>
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1265 |
(REGULARIZE_trm lthy t1 t1') $ (REGULARIZE_trm lthy t2 t2')
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1266 |
| (Free (x, ty), Free (x', ty')) =>
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1267 |
if x = x'
|
330
|
1268 |
then rtrm (* FIXME: check whether types corresponds *)
|
323
|
1269 |
else trm_lift_error lthy rtrm qtrm
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1270 |
| (Bound i, Bound i') =>
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1271 |
if i = i'
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1272 |
then rtrm
|
323
|
1273 |
else trm_lift_error lthy rtrm qtrm
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1274 |
| (Const (s, ty), Const (s', ty')) =>
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1275 |
if s = s' andalso ty = ty'
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1276 |
then rtrm
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1277 |
else rtrm (* FIXME: check correspondence according to definitions *)
|
323
|
1278 |
| _ => trm_lift_error lthy rtrm qtrm
|
320
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1279 |
*}
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1280 |
|
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1281 |
ML {*
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1282 |
fun mk_REGULARIZE_goal lthy rtrm qtrm =
|
330
|
1283 |
Logic.mk_implies (rtrm, Syntax.check_term lthy (REGULARIZE_trm lthy rtrm qtrm))
|
319
|
1284 |
*}
|
293
|
1285 |
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1286 |
(*
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1287 |
To prove that the old theorem implies the new one, we first
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1288 |
atomize it and then try:
|
330
|
1289 |
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1290 |
- Reflexivity of the relation
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1291 |
- Assumption
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1292 |
- Elimnating quantifiers on both sides of toplevel implication
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1293 |
- Simplifying implications on both sides of toplevel implication
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1294 |
- Ball (Respects ?E) ?P = All ?P
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1295 |
- (\<And>x. ?R x \<Longrightarrow> ?P x \<longrightarrow> ?Q x) \<Longrightarrow> All ?P \<longrightarrow> Ball ?R ?Q
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1296 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1297 |
*)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1298 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1299 |
lemma my_equiv_res_forall2:
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1300 |
fixes P::"'a \<Rightarrow> bool"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1301 |
fixes Q::"'b \<Rightarrow> bool"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1302 |
assumes a: "(All Q) \<longrightarrow> (All P)"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1303 |
shows "(All Q) \<longrightarrow> Ball (Respects E) P"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1304 |
using a by auto
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1305 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1306 |
lemma my_equiv_res_exists:
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1307 |
fixes P::"'a \<Rightarrow> bool"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1308 |
fixes Q::"'b \<Rightarrow> bool"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1309 |
assumes a: "EQUIV E"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1310 |
and b: "(Ex Q) \<longrightarrow> (Ex P)"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1311 |
shows "(Ex Q) \<longrightarrow> Bex (Respects E) P"
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1312 |
apply(subst equiv_res_exists)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1313 |
apply(rule a)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1314 |
apply(rule b)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1315 |
done
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1316 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1317 |
ML {*
|
330
|
1318 |
(* FIXME: get_rid of rel_refl rel_eqv *)
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1319 |
fun REGULARIZE_tac lthy rel_refl rel_eqv =
|
325
|
1320 |
(REPEAT1 o FIRST1)
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1321 |
[rtac rel_refl,
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1322 |
atac,
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1323 |
rtac @{thm universal_twice},
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1324 |
rtac @{thm impI} THEN' atac,
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1325 |
rtac @{thm implication_twice},
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1326 |
rtac @{thm my_equiv_res_forall2},
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1327 |
rtac (rel_eqv RS @{thm my_equiv_res_exists}),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1328 |
(* For a = b \<longrightarrow> a \<approx> b *)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1329 |
rtac @{thm impI} THEN' (asm_full_simp_tac HOL_ss) THEN' rtac rel_refl,
|
325
|
1330 |
rtac @{thm RIGHT_RES_FORALL_REGULAR}]
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1331 |
*}
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1332 |
|
330
|
1333 |
(* version of REGULARIZE_tac including debugging information *)
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1334 |
ML {*
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1335 |
fun my_print_tac ctxt s thm =
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1336 |
let
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1337 |
val prems_str = prems_of thm
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1338 |
|> map (Syntax.string_of_term ctxt)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1339 |
|> cat_lines
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1340 |
val _ = tracing (s ^ "\n" ^ prems_str)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1341 |
in
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1342 |
Seq.single thm
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1343 |
end
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1344 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1345 |
fun DT ctxt s tac = EVERY' [tac, K (my_print_tac ctxt ("after " ^ s))]
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1346 |
*}
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1347 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1348 |
ML {*
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1349 |
fun REGULARIZE_tac' lthy rel_refl rel_eqv =
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1350 |
(REPEAT1 o FIRST1)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1351 |
[(K (print_tac "start")) THEN' (K no_tac),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1352 |
DT lthy "1" (rtac rel_refl),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1353 |
DT lthy "2" atac,
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1354 |
DT lthy "3" (rtac @{thm universal_twice}),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1355 |
DT lthy "4" (rtac @{thm impI} THEN' atac),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1356 |
DT lthy "5" (rtac @{thm implication_twice}),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1357 |
DT lthy "6" (rtac @{thm my_equiv_res_forall2}),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1358 |
DT lthy "7" (rtac (rel_eqv RS @{thm my_equiv_res_exists})),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1359 |
(* For a = b \<longrightarrow> a \<approx> b *)
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1360 |
DT lthy "8" (rtac @{thm impI} THEN' (asm_full_simp_tac HOL_ss) THEN' rtac rel_refl),
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1361 |
DT lthy "9" (rtac @{thm RIGHT_RES_FORALL_REGULAR})]
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1362 |
*}
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1363 |
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1364 |
ML {*
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1365 |
fun REGULARIZE_prove rtrm qtrm rel_eqv rel_refl lthy =
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1366 |
let
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1367 |
val goal = mk_REGULARIZE_goal lthy rtrm qtrm
|
330
|
1368 |
in
|
|
1369 |
Goal.prove lthy [] [] goal
|
325
|
1370 |
(fn {context, ...} => REGULARIZE_tac' context rel_refl rel_eqv)
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1371 |
end
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1372 |
*}
|
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1373 |
|
336
|
1374 |
(* rep-abs injection *)
|
|
1375 |
|
|
1376 |
ML {*
|
|
1377 |
fun mk_repabs lthy (T, T') trm =
|
|
1378 |
Quotient_Def.get_fun repF lthy (T, T')
|
354
2eb6d527dfe4
addded a tactic, which sets up the three goals of the `algorithm'
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1379 |
$ (Quotient_Def.get_fun absF lthy (T, T') $ trm)
|
336
|
1380 |
*}
|
|
1381 |
|
|
1382 |
|
|
1383 |
ML {*
|
|
1384 |
(* bound variables need to be treated properly, *)
|
|
1385 |
(* as the type of subterms need to be calculated *)
|
|
1386 |
|
345
|
1387 |
|
336
|
1388 |
fun inj_REPABS lthy (rtrm, qtrm) =
|
|
1389 |
let
|
|
1390 |
val rty = fastype_of rtrm
|
|
1391 |
val qty = fastype_of qtrm
|
|
1392 |
in
|
|
1393 |
case (rtrm, qtrm) of
|
|
1394 |
(Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name "All"}, _) $ t') =>
|
|
1395 |
Const (@{const_name "Ball"}, T) $ r $ (inj_REPABS lthy (t, t'))
|
|
1396 |
| (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name "Ex"}, _) $ t') =>
|
|
1397 |
Const (@{const_name "Bex"}, T) $ r $ (inj_REPABS lthy (t, t'))
|
|
1398 |
| (Const (@{const_name "Babs"}, T) $ r $ t, t') =>
|
|
1399 |
Const (@{const_name "Babs"}, T) $ r $ (inj_REPABS lthy (t, t'))
|
|
1400 |
| (Abs (x, T, t), Abs (x', T', t')) =>
|
|
1401 |
let
|
|
1402 |
val (y, s) = Term.dest_abs (x, T, t)
|
|
1403 |
val (_, s') = Term.dest_abs (x', T', t')
|
|
1404 |
val yvar = Free (y, T)
|
345
|
1405 |
val result = lambda yvar (inj_REPABS lthy (s, s'))
|
336
|
1406 |
in
|
345
|
1407 |
if rty = qty
|
|
1408 |
then result
|
|
1409 |
else mk_repabs lthy (rty, qty) result
|
336
|
1410 |
end
|
|
1411 |
| _ =>
|
345
|
1412 |
(* FIXME / TODO: this is a case that needs to be looked at *)
|
|
1413 |
(* - variables get a rep-abs insde and outside an application *)
|
|
1414 |
(* - constants only get a rep-abs on the outside of the application *)
|
|
1415 |
(* - applications get a rep-abs insde and outside an application *)
|
336
|
1416 |
let
|
|
1417 |
val (rhead, rargs) = strip_comb rtrm
|
|
1418 |
val (qhead, qargs) = strip_comb qtrm
|
345
|
1419 |
val rargs' = map (inj_REPABS lthy) (rargs ~~ qargs)
|
336
|
1420 |
in
|
345
|
1421 |
if rty = qty
|
355
|
1422 |
then
|
|
1423 |
case (rhead, qhead) of
|
|
1424 |
(Free (_, T), Free (_, T')) =>
|
|
1425 |
if T = T' then list_comb (rhead, rargs')
|
|
1426 |
else list_comb (mk_repabs lthy (T, T') rhead, rargs')
|
|
1427 |
| _ => list_comb (rhead, rargs')
|
345
|
1428 |
else
|
|
1429 |
case (rhead, qhead, length rargs') of
|
|
1430 |
(Const _, Const _, 0) => mk_repabs lthy (rty, qty) rhead
|
355
|
1431 |
| (Free (_, T), Free (_, T'), 0) => mk_repabs lthy (T, T') rhead
|
345
|
1432 |
| (Const _, Const _, _) => mk_repabs lthy (rty, qty) (list_comb (rhead, rargs'))
|
|
1433 |
| (Free (x, T), Free (x', T'), _) =>
|
|
1434 |
mk_repabs lthy (rty, qty) (list_comb (mk_repabs lthy (T, T') rhead, rargs'))
|
|
1435 |
| (Abs _, Abs _, _ ) =>
|
|
1436 |
mk_repabs lthy (rty, qty) (list_comb (inj_REPABS lthy (rhead, qhead), rargs'))
|
|
1437 |
| _ => trm_lift_error lthy rtrm qtrm
|
336
|
1438 |
end
|
|
1439 |
end
|
|
1440 |
*}
|
|
1441 |
|
|
1442 |
ML {*
|
|
1443 |
fun mk_inj_REPABS_goal lthy (rtrm, qtrm) =
|
|
1444 |
Logic.mk_equals (rtrm, Syntax.check_term lthy (inj_REPABS lthy (rtrm, qtrm)))
|
|
1445 |
*}
|
|
1446 |
|
347
|
1447 |
(* Final wrappers *)
|
|
1448 |
|
357
|
1449 |
ML {*
|
|
1450 |
fun regularize_tac ctxt rel_eqv rel_refl =
|
|
1451 |
(ObjectLogic.full_atomize_tac) THEN'
|
|
1452 |
REPEAT_ALL_NEW (FIRST' [
|
|
1453 |
rtac rel_refl,
|
|
1454 |
atac,
|
|
1455 |
rtac @{thm universal_twice},
|
|
1456 |
(rtac @{thm impI} THEN' atac),
|
|
1457 |
rtac @{thm implication_twice},
|
|
1458 |
EqSubst.eqsubst_tac ctxt [0]
|
|
1459 |
[(@{thm equiv_res_forall} OF [rel_eqv]),
|
|
1460 |
(@{thm equiv_res_exists} OF [rel_eqv])],
|
|
1461 |
(* For a = b \<longrightarrow> a \<approx> b *)
|
|
1462 |
(rtac @{thm impI} THEN' (asm_full_simp_tac HOL_ss) THEN' rtac rel_refl),
|
|
1463 |
(rtac @{thm RIGHT_RES_FORALL_REGULAR})
|
|
1464 |
]);
|
|
1465 |
*}
|
347
|
1466 |
|
|
1467 |
ML {*
|
|
1468 |
fun regularize_goal lthy thm rel_eqv rel_refl qtrm =
|
|
1469 |
let
|
|
1470 |
val reg_trm = mk_REGULARIZE_goal lthy (prop_of thm) qtrm;
|
357
|
1471 |
fun tac lthy = regularize_tac lthy rel_eqv rel_refl;
|
347
|
1472 |
val cthm = Goal.prove lthy [] [] reg_trm
|
|
1473 |
(fn {context, ...} => tac context 1);
|
|
1474 |
in
|
|
1475 |
cthm OF [thm]
|
|
1476 |
end
|
|
1477 |
*}
|
|
1478 |
|
|
1479 |
ML {*
|
|
1480 |
fun repabs_goal lthy thm rty quot_thm reflex_thm trans_thm rsp_thms qtrm =
|
|
1481 |
let
|
|
1482 |
val rg = Syntax.check_term lthy (mk_inj_REPABS_goal lthy ((prop_of thm), qtrm));
|
|
1483 |
fun tac ctxt = (ObjectLogic.full_atomize_tac) THEN'
|
|
1484 |
(REPEAT_ALL_NEW (r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms));
|
|
1485 |
val cthm = Goal.prove lthy [] [] rg (fn x => tac (#context x) 1);
|
|
1486 |
in
|
|
1487 |
@{thm Pure.equal_elim_rule1} OF [cthm, thm]
|
|
1488 |
end
|
|
1489 |
*}
|
|
1490 |
|
|
1491 |
ML {*
|
|
1492 |
fun lift_thm_goal lthy qty qty_name rsp_thms defs rthm goal =
|
|
1493 |
let
|
|
1494 |
val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data lthy qty;
|
|
1495 |
val (trans2, reps_same, absrep, quot) = lookup_quot_thms lthy qty_name;
|
|
1496 |
val t_a = atomize_thm rthm;
|
|
1497 |
val goal_a = atomize_goal (ProofContext.theory_of lthy) goal;
|
|
1498 |
val t_r = regularize_goal lthy t_a rel_eqv rel_refl goal_a;
|
|
1499 |
val t_t = repabs_goal lthy t_r rty quot rel_refl trans2 rsp_thms goal_a;
|
|
1500 |
val (alls, exs) = findallex lthy rty qty (prop_of t_a);
|
|
1501 |
val allthms = map (make_allex_prs_thm lthy quot @{thm FORALL_PRS}) alls
|
|
1502 |
val exthms = map (make_allex_prs_thm lthy quot @{thm EXISTS_PRS}) exs
|
|
1503 |
val t_a = MetaSimplifier.rewrite_rule (allthms @ exthms) t_t
|
|
1504 |
val abs = findabs rty (prop_of t_a);
|
|
1505 |
val aps = findaps rty (prop_of t_a);
|
|
1506 |
val app_prs_thms = map (applic_prs lthy rty qty absrep) aps;
|
|
1507 |
val lam_prs_thms = map (make_simp_prs_thm lthy quot @{thm LAMBDA_PRS}) abs;
|
|
1508 |
val t_l = repeat_eqsubst_thm lthy (lam_prs_thms @ app_prs_thms) t_a;
|
|
1509 |
val defs_sym = flat (map (add_lower_defs lthy) defs);
|
|
1510 |
val defs_sym_eq = map (fn x => eq_reflection OF [x]) defs_sym;
|
|
1511 |
val t_id = simp_ids lthy t_l;
|
|
1512 |
val t_d0 = MetaSimplifier.rewrite_rule defs_sym_eq t_id;
|
|
1513 |
val t_d = repeat_eqsubst_thm lthy defs_sym t_d0;
|
|
1514 |
val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d;
|
|
1515 |
val t_rv = ObjectLogic.rulify t_r
|
|
1516 |
in
|
|
1517 |
Thm.varifyT t_rv
|
|
1518 |
end
|
|
1519 |
*}
|
|
1520 |
|
348
|
1521 |
ML {*
|
|
1522 |
fun lift_thm_goal_note lthy qty qty_name rsp_thms defs thm name goal =
|
|
1523 |
let
|
|
1524 |
val lifted_thm = lift_thm_goal lthy qty qty_name rsp_thms defs thm goal;
|
|
1525 |
val (_, lthy2) = note (name, lifted_thm) lthy;
|
|
1526 |
in
|
|
1527 |
lthy2
|
|
1528 |
end
|
|
1529 |
*}
|
|
1530 |
|
362
|
1531 |
ML {*
|
|
1532 |
fun spec_frees_tac [] = []
|
|
1533 |
| spec_frees_tac (x::xs) =
|
|
1534 |
let
|
|
1535 |
val spec' = Drule.instantiate' [SOME (ctyp_of_term x)] [NONE, SOME x] @{thm spec}
|
|
1536 |
in
|
|
1537 |
(rtac spec')::(spec_frees_tac xs)
|
|
1538 |
end
|
|
1539 |
*}
|
|
1540 |
|
|
1541 |
ML {*
|
|
1542 |
val prepare_tac = CSUBGOAL (fn (concl, i) =>
|
|
1543 |
let
|
|
1544 |
val vrs = Thm.add_cterm_frees concl []
|
|
1545 |
in
|
|
1546 |
EVERY' (ObjectLogic.full_atomize_tac::spec_frees_tac vrs) i
|
|
1547 |
end)
|
|
1548 |
*}
|
|
1549 |
|
360
|
1550 |
lemma procedure:
|
|
1551 |
assumes a: "A"
|
|
1552 |
and b: "A \<Longrightarrow> B"
|
|
1553 |
and c: "B = C"
|
|
1554 |
and d: "C = D"
|
|
1555 |
shows "D"
|
|
1556 |
using a b c d
|
|
1557 |
by simp
|
|
1558 |
|
|
1559 |
ML {*
|
|
1560 |
fun procedure_inst ctxt rtrm qtrm =
|
|
1561 |
let
|
|
1562 |
val thy = ProofContext.theory_of ctxt
|
|
1563 |
val rtrm' = HOLogic.dest_Trueprop rtrm
|
|
1564 |
val qtrm' = HOLogic.dest_Trueprop qtrm
|
|
1565 |
val reg_goal = Syntax.check_term ctxt (REGULARIZE_trm ctxt rtrm' qtrm')
|
|
1566 |
val inj_goal = Syntax.check_term ctxt (inj_REPABS ctxt (reg_goal, qtrm'))
|
|
1567 |
in
|
|
1568 |
Drule.instantiate' []
|
|
1569 |
[SOME (cterm_of thy rtrm'),
|
|
1570 |
SOME (cterm_of thy reg_goal),
|
|
1571 |
SOME (cterm_of thy inj_goal)]
|
|
1572 |
@{thm procedure}
|
|
1573 |
end
|
362
|
1574 |
*}
|
|
1575 |
|
|
1576 |
ML {*
|
|
1577 |
fun procedure_tac rthm ctxt =
|
|
1578 |
prepare_tac THEN'
|
360
|
1579 |
Subgoal.FOCUS (fn {context, concl, ...} =>
|
|
1580 |
let
|
|
1581 |
val rthm' = atomize_thm rthm
|
|
1582 |
val rule = procedure_inst context (prop_of rthm') (term_of concl)
|
|
1583 |
in
|
|
1584 |
EVERY1 [rtac rule, rtac rthm']
|
362
|
1585 |
end) ctxt
|
360
|
1586 |
*}
|
|
1587 |
|
|
1588 |
|
|
1589 |
ML {*
|
362
|
1590 |
(* FIXME: allex_prs and lambda_prs can be one function *)
|
360
|
1591 |
fun allex_prs_tac lthy quot =
|
|
1592 |
(EqSubst.eqsubst_tac lthy [0] @{thms FORALL_PRS[symmetric] EXISTS_PRS[symmetric]})
|
|
1593 |
THEN' (quotient_tac quot);
|
|
1594 |
*}
|
|
1595 |
|
|
1596 |
ML {*
|
|
1597 |
fun lambda_prs_tac lthy quot =
|
|
1598 |
(EqSubst.eqsubst_tac lthy [0] @{thms LAMBDA_PRS}
|
|
1599 |
THEN' (RANGE [quotient_tac quot, quotient_tac quot]));
|
|
1600 |
*}
|
|
1601 |
|
|
1602 |
ML {*
|
361
|
1603 |
fun TRY' tac = fn i => TRY (tac i)
|
|
1604 |
*}
|
|
1605 |
|
|
1606 |
ML {*
|
360
|
1607 |
fun clean_tac lthy quot defs reps_same =
|
|
1608 |
let
|
|
1609 |
val lower = flat (map (add_lower_defs lthy) defs)
|
|
1610 |
in
|
361
|
1611 |
TRY' (REPEAT_ALL_NEW (allex_prs_tac lthy quot)) THEN'
|
|
1612 |
TRY' (REPEAT_ALL_NEW (lambda_prs_tac lthy quot)) THEN'
|
|
1613 |
TRY' (REPEAT_ALL_NEW (EqSubst.eqsubst_tac lthy [0] lower)) THEN'
|
|
1614 |
simp_tac (HOL_ss addsimps [reps_same])
|
360
|
1615 |
end
|
|
1616 |
*}
|
|
1617 |
|
361
|
1618 |
ML {*
|
|
1619 |
fun lift_tac lthy thm rel_eqv rel_refl rty quot trans2 rsp_thms reps_same defs =
|
|
1620 |
(procedure_tac thm lthy) THEN'
|
|
1621 |
(regularize_tac lthy rel_eqv rel_refl) THEN'
|
|
1622 |
(REPEAT_ALL_NEW (r_mk_comb_tac lthy rty quot rel_refl trans2 rsp_thms)) THEN'
|
|
1623 |
(clean_tac lthy quot defs reps_same)
|
|
1624 |
*}
|
|
1625 |
|
|
1626 |
|
321
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1627 |
|
198
|
1628 |
end
|
239
|
1629 |
|
347
|
1630 |
|