2337
|
1 |
(* Title: nominal_dt_alpha.ML
|
|
2 |
Author: Christian Urban
|
|
3 |
Author: Cezary Kaliszyk
|
|
4 |
|
2426
|
5 |
Performing quotient constructions and lifting theorems
|
2337
|
6 |
*)
|
|
7 |
|
|
8 |
signature NOMINAL_DT_QUOT =
|
|
9 |
sig
|
2400
|
10 |
val define_qtypes: (string list * binding * mixfix) list -> typ list -> term list ->
|
2337
|
11 |
thm list -> local_theory -> Quotient_Info.quotdata_info list * local_theory
|
|
12 |
|
2400
|
13 |
val define_qconsts: typ list -> (string * term * mixfix) list -> local_theory ->
|
2337
|
14 |
Quotient_Info.qconsts_info list * local_theory
|
2346
|
15 |
|
2400
|
16 |
val define_qperms: typ list -> string list -> (string * term * mixfix) list ->
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
17 |
thm list -> local_theory -> local_theory
|
2400
|
18 |
|
|
19 |
val define_qsizes: typ list -> string list -> (string * term * mixfix) list ->
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
20 |
local_theory -> local_theory
|
2426
|
21 |
|
2430
|
22 |
val lift_thm: Proof.context -> typ list -> thm list -> thm -> thm
|
2337
|
23 |
end
|
|
24 |
|
|
25 |
structure Nominal_Dt_Quot: NOMINAL_DT_QUOT =
|
|
26 |
struct
|
|
27 |
|
|
28 |
(* defines the quotient types *)
|
2400
|
29 |
fun define_qtypes qtys_descr alpha_tys alpha_trms alpha_equivp_thms lthy =
|
2337
|
30 |
let
|
|
31 |
val qty_args1 = map2 (fn ty => fn trm => (ty, trm, false)) alpha_tys alpha_trms
|
|
32 |
val qty_args2 = (qtys_descr ~~ qty_args1) ~~ alpha_equivp_thms
|
|
33 |
in
|
|
34 |
fold_map Quotient_Type.add_quotient_type qty_args2 lthy
|
|
35 |
end
|
|
36 |
|
2338
|
37 |
|
2337
|
38 |
(* defines quotient constants *)
|
2400
|
39 |
fun define_qconsts qtys consts_specs lthy =
|
2337
|
40 |
let
|
2338
|
41 |
val (qconst_infos, lthy') =
|
|
42 |
fold_map (Quotient_Def.lift_raw_const qtys) consts_specs lthy
|
2337
|
43 |
val phi = ProofContext.export_morphism lthy' lthy
|
|
44 |
in
|
|
45 |
(map (Quotient_Info.transform_qconsts phi) qconst_infos, lthy')
|
|
46 |
end
|
|
47 |
|
|
48 |
|
2400
|
49 |
(* defines the quotient permutations and proves pt-class *)
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
50 |
fun define_qperms qtys qfull_ty_names perm_specs raw_perm_laws lthy =
|
2346
|
51 |
let
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
52 |
val lthy' =
|
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
53 |
lthy
|
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
54 |
|> Local_Theory.exit_global
|
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
55 |
|> Class.instantiation (qfull_ty_names, [], @{sort pt})
|
2398
|
56 |
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
57 |
val (_, lthy'') = define_qconsts qtys perm_specs lthy'
|
2398
|
58 |
|
2430
|
59 |
val lifted_perm_laws = map (Quotient_Tacs.lifted lthy'' qtys []) raw_perm_laws
|
2398
|
60 |
|
2346
|
61 |
fun tac _ =
|
|
62 |
Class.intro_classes_tac [] THEN
|
2400
|
63 |
(ALLGOALS (resolve_tac lifted_perm_laws))
|
2346
|
64 |
in
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
65 |
lthy''
|
2398
|
66 |
|> Class.prove_instantiation_exit tac
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
67 |
|> Named_Target.theory_init
|
2346
|
68 |
end
|
|
69 |
|
2337
|
70 |
|
2400
|
71 |
(* defines the size functions and proves size-class *)
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
72 |
fun define_qsizes qtys qfull_ty_names size_specs lthy =
|
2400
|
73 |
let
|
|
74 |
fun tac _ = Class.intro_classes_tac []
|
|
75 |
in
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
76 |
lthy
|
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
77 |
|> Local_Theory.exit_global
|
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
78 |
|> Class.instantiation (qfull_ty_names, [], @{sort size})
|
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
79 |
|> snd o (define_qconsts qtys size_specs)
|
2400
|
80 |
|> Class.prove_instantiation_exit tac
|
2401
7645e18e8b19
modified the code for class instantiations (with help from Florian)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
81 |
|> Named_Target.theory_init
|
2400
|
82 |
end
|
|
83 |
|
2426
|
84 |
|
|
85 |
(* lifts a theorem and deletes all "_raw" suffixes *)
|
|
86 |
|
|
87 |
fun unraw_str s =
|
|
88 |
unsuffix "_raw" s
|
|
89 |
handle Fail _ => s
|
|
90 |
|
|
91 |
fun unraw_vars_thm thm =
|
|
92 |
let
|
|
93 |
fun unraw_var_str ((s, i), T) = ((unraw_str s, i), T)
|
|
94 |
|
|
95 |
val vars = Term.add_vars (prop_of thm) []
|
|
96 |
val vars' = map (Var o unraw_var_str) vars
|
|
97 |
in
|
|
98 |
Thm.certify_instantiate ([], (vars ~~ vars')) thm
|
|
99 |
end
|
|
100 |
|
|
101 |
fun unraw_bounds_thm th =
|
|
102 |
let
|
|
103 |
val trm = Thm.prop_of th
|
|
104 |
val trm' = Term.map_abs_vars unraw_str trm
|
|
105 |
in
|
|
106 |
Thm.rename_boundvars trm trm' th
|
|
107 |
end
|
|
108 |
|
2430
|
109 |
fun lift_thm ctxt qtys simps thm =
|
|
110 |
thm
|
|
111 |
|> Quotient_Tacs.lifted ctxt qtys simps
|
2426
|
112 |
|> unraw_bounds_thm
|
|
113 |
|> unraw_vars_thm
|
|
114 |
|
|
115 |
|
2337
|
116 |
end (* structure *)
|
|
117 |
|