1274
|
1 |
theory Lift
|
1280
|
2 |
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove"
|
1274
|
3 |
begin
|
|
4 |
|
|
5 |
atom_decl name
|
|
6 |
atom_decl ident
|
|
7 |
|
|
8 |
datatype rtrm2 =
|
|
9 |
rVr2 "name"
|
1280
|
10 |
| rAp2 "rtrm2" "rtrm2"
|
1274
|
11 |
| rLt2 "ras" "rtrm2" --"bind (bv2 l) in (r)"
|
|
12 |
and ras =
|
|
13 |
rAs "name" "rtrm2"
|
|
14 |
|
|
15 |
primrec rbv2 where "rbv2 (rAs x t) = {atom x}"
|
|
16 |
|
1280
|
17 |
ML {*
|
|
18 |
val thy1 = @{theory};
|
|
19 |
val info = Datatype.the_info @{theory} "Lift.rtrm2"
|
|
20 |
val number = 2; (* Number of defined types, rest are unfoldings *)
|
|
21 |
val binds = [[[[]], [[], []], [[], [(SOME @{term rbv2}, 0)]]],
|
|
22 |
[[[], []]] (*, [[], [[], []]] *) ];
|
|
23 |
val bvs = [(@{term rbv2}, 1)] (* Which type it operates on *)
|
|
24 |
val bv_simps = @{thms rbv2.simps}
|
1274
|
25 |
|
1280
|
26 |
val ntnames = [@{binding trm2}, @{binding as}]
|
|
27 |
val ncnames = ["Vr2", "Ap2", "Lt2", "As"]
|
|
28 |
|
|
29 |
|
1274
|
30 |
|
|
31 |
|
|
32 |
val descr = #descr info;
|
1280
|
33 |
val sorts = #sorts info;
|
|
34 |
val nos = map fst descr
|
|
35 |
val all_typs = map (fn i => typ_of_dtyp descr sorts (DtRec i)) nos
|
|
36 |
val typs = List.take (all_typs, number)
|
1274
|
37 |
val all_full_tnames = map (fn (_, (n, _, _)) => n) descr;
|
1280
|
38 |
val full_tnames = List.take (all_full_tnames, number);
|
1274
|
39 |
val induct = #induct info;
|
|
40 |
val inducts = #inducts info;
|
1280
|
41 |
val infos = map (Datatype.the_info thy1) all_full_tnames;
|
1274
|
42 |
val inject = flat (map #inject infos);
|
|
43 |
val distinct = flat (map #distinct infos);
|
1280
|
44 |
val ((raw_perm_def, raw_perm_simps, perms), thy2) = define_raw_perms info number thy1;
|
1274
|
45 |
val lthy1 = Theory_Target.init NONE thy2
|
1280
|
46 |
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>
diff
changeset
|
47 |
val alpha_ts_loc = #preds alpha
|
1274
|
48 |
val alpha_intros = #intrs alpha
|
|
49 |
val alpha_cases = #elims alpha
|
1280
|
50 |
val alpha_induct_loc = #induct alpha
|
|
51 |
val [alpha_induct] = ProofContext.export lthy2 lthy1 [alpha_induct_loc]
|
|
52 |
(* TODO replace when inducts is provided by the 2 lines below: *)
|
|
53 |
val alpha_inducts = Project_Rule.projects lthy2 (1 upto number) alpha_induct
|
|
54 |
(*val alpha_inducts_loc = #inducts alpha
|
|
55 |
val alpha_inducts = ProofContext.export lthy2 lthy1 alpha_inducts_loc*)
|
|
56 |
val alpha_inj_loc = build_alpha_inj alpha_intros (inject @ distinct) alpha_cases lthy2
|
|
57 |
val alpha_inj = ProofContext.export lthy2 lthy1 alpha_inj_loc
|
1274
|
58 |
val fv_def = ProofContext.export lthy2 lthy1 fv_def_loc
|
|
59 |
val morphism = ProofContext.export_morphism lthy2 lthy1
|
|
60 |
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>
diff
changeset
|
61 |
val alpha_ts = map (Morphism.term morphism) alpha_ts_loc
|
1280
|
62 |
fun build_bv_eqvt (t, n) =
|
|
63 |
build_eqvts Binding.empty [t] [nth perms n]
|
|
64 |
(bv_simps @ raw_perm_def) (nth inducts n)
|
|
65 |
val (bv_eqvts, lthy3) = fold_map build_bv_eqvt bvs lthy2;
|
|
66 |
val (fv_eqvts, lthy4) = build_eqvts Binding.empty fv_ts_loc perms (fv_def_loc @ raw_perm_def) induct lthy3;
|
|
67 |
val alpha_eqvt_loc = build_alpha_eqvts alpha_ts_loc perms (raw_perm_def @ alpha_inj_loc) alpha_induct_loc lthy4;
|
|
68 |
val alpha_eqvt = ProofContext.export lthy4 lthy1 alpha_eqvt_loc
|
|
69 |
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>
diff
changeset
|
70 |
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>
diff
changeset
|
71 |
val lthy5 = define_quotient_type
|
1280
|
72 |
(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>
diff
changeset
|
73 |
(ALLGOALS (resolve_tac alpha_equivp)) lthy4;
|
1280
|
74 |
val consts =
|
|
75 |
flat (map (fn (i, (_, _, l)) =>
|
|
76 |
map (fn (cname, dts) =>
|
|
77 |
Const (cname, map (typ_of_dtyp descr sorts) dts --->
|
|
78 |
typ_of_dtyp descr sorts (DtRec i))) l) descr);
|
|
79 |
val (csdefl, lthy6) = fold_map Quotient_Def.quotient_lift_const (ncnames ~~ consts) lthy5;
|
|
80 |
val (cs, def) = split_list csdefl;
|
|
81 |
val ((_, fv_rsp), lthy7) = prove_const_rsp Binding.empty fv_ts
|
|
82 |
(fn _ => fvbv_rsp_tac alpha_induct fv_def 1) lthy6
|
|
83 |
val (bvs_rsp', lthy8) = fold_map (
|
|
84 |
fn (bv_t, i) => prove_const_rsp Binding.empty [bv_t]
|
|
85 |
(fn _ => fvbv_rsp_tac (nth alpha_inducts i) bv_simps 1)) bvs lthy7
|
|
86 |
val bvs_rsp = flat (map snd bvs_rsp')
|
|
87 |
val (const_rsps, lthy9) = fold_map (fn cnst => prove_const_rsp Binding.empty [cnst]
|
|
88 |
(fn _ => constr_rsp_tac alpha_inj (fv_rsp @ bvs_rsp) alpha_equivp 1)) consts lthy8
|
|
89 |
val (perms_rsp, lthy10) = prove_const_rsp Binding.empty perms
|
|
90 |
(fn _ => asm_simp_tac (HOL_ss addsimps alpha_eqvt) 1) lthy9;
|
|
91 |
val lift_induct = snd (Quotient_Tacs.lifted_attrib (Context.Proof lthy1, induct));
|
|
92 |
val lthy11 = snd (Local_Theory.note ((@{binding lift_induct}, []), [lift_induct]) lthy10)
|
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>
diff
changeset
|
93 |
*}
|
1274
|
94 |
|
1280
|
95 |
setup {* fn _ => Local_Theory.exit_global lthy11 *}
|
|
96 |
thm lift_induct
|
1274
|
97 |
|
1280
|
98 |
end
|
1274
|
99 |
|