author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 02 Mar 2010 14:24:57 +0100 | |
changeset 1303 | c28403308b34 |
parent 1291 | 24889782da92 |
child 1309 | b395b902cf0d |
permissions | -rw-r--r-- |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory Lift |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
2 |
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove" |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
|
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
atom_decl name |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
atom_decl ident |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
|
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
8 |
(* datatype rtrm2 = |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
rVr2 "name" |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
10 |
| rAp2 "rtrm2" "rtrm2" |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
11 |
| rLt2 "ras" "rtrm2" --"bind (bv2 l) in (r)" |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
and ras = |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
13 |
rAs "name" "rtrm2" |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
14 |
|
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
15 |
primrec rbv2 where "rbv2 (rAs x t) = {atom x}" |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
16 |
ML {* |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
17 |
val thy1 = @{theory}; |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
18 |
val info = Datatype.the_info @{theory} "Lift.rtrm2" |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
19 |
val number = 2; (* Number of defined types, rest are unfoldings *) |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
20 |
val binds = [[[[]], [[], []], [[], [(SOME @{term rbv2}, 0)]]], |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
21 |
[[[], []]] (*, [[], [[], []]] *) ]; |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
22 |
val bvs = [(@{term rbv2}, 1)] (* Which type it operates on *) |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
23 |
val bv_simps = @{thms rbv2.simps} |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
24 |
|
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
25 |
val ntnames = [@{binding trm2}, @{binding as}] |
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
26 |
val ncnames = ["Vr2", "Ap2", "Lt2", "As"] *} *) |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
27 |
|
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
28 |
datatype rkind = |
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
29 |
Type |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
30 |
| KPi "rty" "name" "rkind" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
31 |
and rty = |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
32 |
TConst "ident" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
33 |
| TApp "rty" "rtrm" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
34 |
| TPi "rty" "name" "rty" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
35 |
and rtrm = |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
36 |
Const "ident" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
37 |
| Var "name" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
38 |
| App "rtrm" "rtrm" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
39 |
| Lam "rty" "name" "rtrm" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
40 |
ML {* |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
41 |
val thy1 = @{theory}; |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
42 |
val info = Datatype.the_info @{theory} "Lift.rkind" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
43 |
val number = 3; |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
44 |
val binds = [[ [], [(NONE, 1, 2)]], |
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
45 |
[ [], [], [(NONE, 1, 2)] ], |
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
46 |
[ [], [], [], [(NONE, 1, 2)]]]; |
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
47 |
val bvs = [] |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
48 |
val bv_simps = [] |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
49 |
|
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
50 |
val ntnames = [@{binding kind}, @{binding ty}, @{binding trm}] |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
51 |
val ncnames = ["TYP", "KPI", "TCONST", "TAPP", "TPI", "CONST", "VAR", "APP", "LAM"] |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
52 |
|
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
53 |
(*datatype rtrm5 = |
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
54 |
rVr5 "name" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
55 |
| rAp5 "rtrm5" "rtrm5" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
56 |
| rLt5 "rlts" "rtrm5" --"bind (bv5 lts) in (rtrm5)" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
57 |
and rlts = |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
58 |
rLnil |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
59 |
| rLcons "name" "rtrm5" "rlts" |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
60 |
|
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
61 |
primrec |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
62 |
rbv5 |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
63 |
where |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
64 |
"rbv5 rLnil = {}" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
65 |
| "rbv5 (rLcons n t ltl) = {atom n} \<union> (rbv5 ltl)" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
66 |
|
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
67 |
ML |
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
68 |
val thy1 = @{theory}; |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
69 |
val info = Datatype.the_info @{theory} "Lift.rtrm5" |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
70 |
val number = 2; |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
71 |
val binds = [ [[], [], [(SOME @{term rbv5}, 0, 1)]], [[], []] ] |
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
72 |
val bvs = [(@{term rbv5}, 1)] |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
73 |
val bv_simps = @{thms rbv5.simps} |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
74 |
|
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
75 |
val ntnames = [@{binding trm5}, @{binding lts}] |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
76 |
val ncnames = ["Vr5", "Ap5", "Lt5", "Lnil", "Lcons"]*) |
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
77 |
|
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
78 |
|
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
79 |
val descr = #descr info; |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
80 |
val sorts = #sorts info; |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
81 |
val nos = map fst descr |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
82 |
val all_typs = map (fn i => typ_of_dtyp descr sorts (DtRec i)) nos |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
83 |
val typs = List.take (all_typs, number) |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
84 |
val all_full_tnames = map (fn (_, (n, _, _)) => n) descr; |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
85 |
val full_tnames = List.take (all_full_tnames, number); |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
86 |
val induct = #induct info; |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
87 |
val inducts = #inducts info; |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
88 |
val infos = map (Datatype.the_info thy1) all_full_tnames; |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
89 |
val inject = flat (map #inject infos); |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
90 |
val distinct = flat (map #distinct infos); |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
91 |
val ((raw_perm_def, raw_perm_simps, perms), thy2) = define_raw_perms info number thy1; |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
92 |
val lthy1 = Theory_Target.init NONE thy2 |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
93 |
val (((fv_ts_loc, fv_def_loc), alpha), lthy2) = define_fv_alpha info binds lthy1; |
1276
3365fce80f0f
To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1274
diff
changeset
|
94 |
val alpha_ts_loc = #preds alpha |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
95 |
val alpha_intros = #intrs alpha |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
96 |
val alpha_cases = #elims alpha |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
97 |
val alpha_induct_loc = #induct alpha |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
98 |
val [alpha_induct] = ProofContext.export lthy2 lthy1 [alpha_induct_loc] |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
99 |
(* TODO replace when inducts is provided by the 2 lines below: *) |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
100 |
val alpha_inducts = Project_Rule.projects lthy2 (1 upto number) alpha_induct |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
101 |
(*val alpha_inducts_loc = #inducts alpha |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
102 |
val alpha_inducts = ProofContext.export lthy2 lthy1 alpha_inducts_loc*) |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
103 |
val alpha_inj_loc = build_alpha_inj alpha_intros (inject @ distinct) alpha_cases lthy2 |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
104 |
val alpha_inj = ProofContext.export lthy2 lthy1 alpha_inj_loc |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
105 |
val fv_def = ProofContext.export lthy2 lthy1 fv_def_loc |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
106 |
val morphism = ProofContext.export_morphism lthy2 lthy1 |
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
107 |
val fv_ts = map (Morphism.term morphism) fv_ts_loc |
1276
3365fce80f0f
To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1274
diff
changeset
|
108 |
val alpha_ts = map (Morphism.term morphism) alpha_ts_loc |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
109 |
fun build_bv_eqvt (t, n) = |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
110 |
build_eqvts Binding.empty [t] [nth perms n] |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
111 |
(bv_simps @ raw_perm_def) (nth inducts n) |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
112 |
val (bv_eqvts, lthy3) = fold_map build_bv_eqvt bvs lthy2; |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
113 |
val (fv_eqvts, lthy4) = build_eqvts Binding.empty fv_ts_loc perms (fv_def_loc @ raw_perm_def) induct lthy3; |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
114 |
val alpha_eqvt_loc = build_alpha_eqvts alpha_ts_loc perms (raw_perm_def @ alpha_inj_loc) alpha_induct_loc lthy4; |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
115 |
val alpha_eqvt = ProofContext.export lthy4 lthy1 alpha_eqvt_loc; |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
116 |
val alpha_equivp_loc = build_equivps alpha_ts_loc induct alpha_induct_loc inject alpha_inj_loc distinct alpha_cases alpha_eqvt_loc lthy4; |
1276
3365fce80f0f
To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1274
diff
changeset
|
117 |
val alpha_equivp = ProofContext.export lthy4 lthy1 alpha_equivp_loc |
3365fce80f0f
To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1274
diff
changeset
|
118 |
val lthy5 = define_quotient_type |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
119 |
(map (fn ((b, t), alpha) => (([], b, NoSyn), (t, alpha))) ((ntnames ~~ typs) ~~ alpha_ts)) |
1276
3365fce80f0f
To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1274
diff
changeset
|
120 |
(ALLGOALS (resolve_tac alpha_equivp)) lthy4; |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
121 |
val consts = |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
122 |
flat (map (fn (i, (_, _, l)) => |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
123 |
map (fn (cname, dts) => |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
124 |
Const (cname, map (typ_of_dtyp descr sorts) dts ---> |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
125 |
typ_of_dtyp descr sorts (DtRec i))) l) descr); |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
126 |
val (csdefl, lthy6) = fold_map Quotient_Def.quotient_lift_const (ncnames ~~ consts) lthy5; |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
127 |
val (cs, def) = split_list csdefl; |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
128 |
val (bvs_rsp', lthy7) = fold_map ( |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
129 |
fn (bv_t, i) => prove_const_rsp Binding.empty [bv_t] |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
130 |
(fn _ => fvbv_rsp_tac (nth alpha_inducts i) bv_simps 1)) bvs lthy6; |
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
131 |
val ((_, fv_rsp), lthy8) = prove_const_rsp Binding.empty fv_ts |
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
132 |
(fn _ => fvbv_rsp_tac alpha_induct fv_def 1) lthy7; |
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
133 |
val bvs_rsp = flat (map snd bvs_rsp'); |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
134 |
val (const_rsps, lthy9) = fold_map (fn cnst => prove_const_rsp Binding.empty [cnst] |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
135 |
(fn _ => constr_rsp_tac alpha_inj (fv_rsp @ bvs_rsp) alpha_equivp 1)) consts lthy8 |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
136 |
val (perms_rsp, lthy10) = prove_const_rsp Binding.empty perms |
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
137 |
(fn _ => asm_simp_tac (HOL_ss addsimps alpha_eqvt) 1) lthy9; |
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
138 |
val thy3 = Local_Theory.exit_global lthy10; |
1291
24889782da92
Trying to prove equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1282
diff
changeset
|
139 |
(* TODO: fix this hack... *) |
1303
c28403308b34
More fixes for new alpha, the whole lift script should now work again.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1291
diff
changeset
|
140 |
(*val tinfo = #abs_type (Typedef.the_info thy3 "Lift.trm5");*) |
1291
24889782da92
Trying to prove equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1282
diff
changeset
|
141 |
(*val thy4 = define_lifted_perms ["Term1.trm1"] [("permute_trm1", @{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"})] |
24889782da92
Trying to prove equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1282
diff
changeset
|
142 |
@{thms permute_rtrm1_permute_bp_zero permute_rtrm1_permute_bp_append}*) |
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
143 |
val lthy11 = Theory_Target.init NONE thy3; |
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
144 |
val lift_induct = snd (Quotient_Tacs.lifted_attrib (Context.Proof lthy11, induct)); |
1291
24889782da92
Trying to prove equivariance.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1282
diff
changeset
|
145 |
val lthy12 = snd (Local_Theory.note ((@{binding lift_induct}, []), [lift_induct]) lthy11); |
1276
3365fce80f0f
To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1274
diff
changeset
|
146 |
*} |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
147 |
|
1282
ea46a354f382
More about the general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1280
diff
changeset
|
148 |
setup {* fn _ => Local_Theory.exit_global lthy12 *} |
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
149 |
thm lift_induct |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
150 |
|
1280
1f057f8da8aa
Progress with general lifting procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
151 |
end |
1274
d867021d8ac1
Preparing the generalized lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
152 |