author | wenzelm |
Sat, 05 Oct 2024 14:58:36 +0200 | |
changeset 81116 | 0fb1e2dd4122 |
parent 81106 | 849efff7de15 |
permissions | -rw-r--r-- |
47308 | 1 |
(* Title: HOL/Tools/Lifting/lifting_setup.ML |
2 |
Author: Ondrej Kuncar |
|
3 |
||
47352 | 4 |
Setting up the lifting infrastructure. |
47308 | 5 |
*) |
6 |
||
7 |
signature LIFTING_SETUP = |
|
8 |
sig |
|
9 |
exception SETUP_LIFTING_INFR of string |
|
10 |
||
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
11 |
type config = { notes: bool }; |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
12 |
val default_config: config; |
47308 | 13 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
14 |
val setup_by_quotient: config -> thm -> thm option -> thm option -> local_theory -> |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
15 |
binding * local_theory |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
16 |
|
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
17 |
val setup_by_typedef_thm: config -> thm -> local_theory -> binding * local_theory |
53651 | 18 |
|
19 |
val lifting_restore: Lifting_Info.quotient -> Context.generic -> Context.generic |
|
60226
ec23f2a97ba4
publish lifting_forget and lifting_udpate interface
kuncar
parents:
60225
diff
changeset
|
20 |
|
ec23f2a97ba4
publish lifting_forget and lifting_udpate interface
kuncar
parents:
60225
diff
changeset
|
21 |
val lifting_forget: string -> local_theory -> local_theory |
ec23f2a97ba4
publish lifting_forget and lifting_udpate interface
kuncar
parents:
60225
diff
changeset
|
22 |
val update_transfer_rules: string -> local_theory -> local_theory |
ec23f2a97ba4
publish lifting_forget and lifting_udpate interface
kuncar
parents:
60225
diff
changeset
|
23 |
val pointer_of_bundle_binding: Proof.context -> binding -> string |
53651 | 24 |
end |
47308 | 25 |
|
47334 | 26 |
structure Lifting_Setup: LIFTING_SETUP = |
47308 | 27 |
struct |
28 |
||
47698 | 29 |
open Lifting_Util |
47308 | 30 |
|
47698 | 31 |
infix 0 MRSL |
47308 | 32 |
|
33 |
exception SETUP_LIFTING_INFR of string |
|
34 |
||
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
35 |
(* Config *) |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
36 |
|
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
37 |
type config = { notes: bool }; |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
38 |
val default_config = { notes = true }; |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
39 |
|
60239 | 40 |
fun define_crel (config: config) rep_fun lthy = |
47308 | 41 |
let |
47361
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47352
diff
changeset
|
42 |
val (qty, rty) = (dest_funT o fastype_of) rep_fun |
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
kuncar
parents:
47352
diff
changeset
|
43 |
val rep_fun_graph = (HOLogic.eq_const rty) $ Bound 1 $ (rep_fun $ Bound 0) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
44 |
val def_term = Abs ("x", rty, Abs ("y", qty, rep_fun_graph)) |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
45 |
val qty_name = Binding.name (Long_Name.base_name (dest_Type_name qty)) |
50227 | 46 |
val crel_name = Binding.prefix_name "cr_" qty_name |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
47 |
val (fixed_def_term, lthy1) = lthy |> yield_singleton (Variable.importT_terms) def_term |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
48 |
val ((_, (_ , def_thm)), lthy2) = |
63352 | 49 |
if #notes config then |
50 |
Local_Theory.define |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
51 |
((crel_name, NoSyn), ((Thm.def_binding crel_name, []), fixed_def_term)) lthy1 |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
52 |
else |
63352 | 53 |
Local_Theory.define |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
54 |
((Binding.concealed crel_name, NoSyn), (Binding.empty_atts, fixed_def_term)) lthy1 |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
55 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
56 |
(def_thm, lthy2) |
47308 | 57 |
end |
58 |
||
50227 | 59 |
fun print_define_pcrel_warning msg = |
60 |
let |
|
61 |
val warning_msg = cat_lines |
|
62 |
["Generation of a parametrized correspondence relation failed.", |
|
63 |
(Pretty.string_of (Pretty.block |
|
64 |
[Pretty.str "Reason:", Pretty.brk 2, msg]))] |
|
65 |
in |
|
66 |
warning warning_msg |
|
67 |
end |
|
68 |
||
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
69 |
fun define_pcrel (config: config) crel lthy0 = |
50227 | 70 |
let |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
71 |
val (fixed_crel, lthy1) = yield_singleton Variable.importT_terms crel lthy0 |
50288
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents:
50231
diff
changeset
|
72 |
val [rty', qty] = (binder_types o fastype_of) fixed_crel |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
73 |
val (param_rel, args) = Lifting_Term.generate_parametrized_relator lthy1 rty' |
50288
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents:
50231
diff
changeset
|
74 |
val rty_raw = (domain_type o range_type o fastype_of) param_rel |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
75 |
val tyenv_match = Sign.typ_match (Proof_Context.theory_of lthy1) (rty_raw, rty') Vartab.empty |
50288
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents:
50231
diff
changeset
|
76 |
val param_rel_subst = Envir.subst_term (tyenv_match,Vartab.empty) param_rel |
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents:
50231
diff
changeset
|
77 |
val args_subst = map (Envir.subst_term (tyenv_match,Vartab.empty)) args |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
78 |
val (instT, lthy2) = lthy1 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
79 |
|> Variable.declare_names fixed_crel |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
80 |
|> Variable.importT_inst (param_rel_subst :: args_subst) |
74266 | 81 |
val args_fixed = (map (Term_Subst.instantiate (instT, Vars.empty))) args_subst |
82 |
val param_rel_fixed = Term_Subst.instantiate (instT, Vars.empty) param_rel_subst |
|
50288
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents:
50231
diff
changeset
|
83 |
val rty = (domain_type o fastype_of) param_rel_fixed |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
84 |
val qty_name = dest_Type_name qty |
50227 | 85 |
val pcrel_name = Binding.prefix_name "pcr_" ((Binding.name o Long_Name.base_name) qty_name) |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
86 |
val relator_type = foldr1 (op -->) ((map type_of args_fixed) @ [rty, qty, HOLogic.boolT]) |
50288
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
kuncar
parents:
50231
diff
changeset
|
87 |
val lhs = Library.foldl (op $) ((Free (Binding.name_of pcrel_name, relator_type)), args_fixed) |
80673 | 88 |
val rhs = \<^Const>\<open>relcompp rty rty' qty for param_rel_fixed fixed_crel\<close> |
50227 | 89 |
val definition_term = Logic.mk_equals (lhs, rhs) |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
90 |
fun note_def lthy = |
63180 | 91 |
Specification.definition (SOME (pcrel_name, SOME relator_type, NoSyn)) [] [] |
63352 | 92 |
(Binding.empty_atts, definition_term) lthy |>> (snd #> snd); |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
93 |
fun raw_def lthy = |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
94 |
let |
63395 | 95 |
val ((_, rhs), prove) = |
96 |
Local_Defs.derived_def lthy (K []) {conditional = true} definition_term; |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
97 |
val ((_, (_, raw_th)), lthy') = |
63352 | 98 |
Local_Theory.define |
99 |
((Binding.concealed pcrel_name, NoSyn), (Binding.empty_atts, rhs)) lthy; |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
100 |
val th = prove lthy' raw_th; |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
101 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
102 |
(th, lthy') |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
103 |
end |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
104 |
val (def_thm, lthy3) = if #notes config then note_def lthy2 else raw_def lthy2 |
50227 | 105 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
106 |
(SOME def_thm, lthy3) |
50227 | 107 |
end |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
108 |
handle Lifting_Term.PARAM_QUOT_THM (_, msg) => (print_define_pcrel_warning msg; (NONE, lthy0)) |
50227 | 109 |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
110 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
111 |
local |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
112 |
val eq_OO_meta = mk_meta_eq @{thm eq_OO} |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
113 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
114 |
fun print_generate_pcr_cr_eq_error ctxt term = |
56257 | 115 |
let |
69593 | 116 |
val goal = Const (\<^const_name>\<open>HOL.eq\<close>, dummyT) $ term $ Const (\<^const_name>\<open>HOL.eq\<close>, dummyT) |
56257 | 117 |
val error_msg = cat_lines |
118 |
["Generation of a pcr_cr_eq failed.", |
|
119 |
(Pretty.string_of (Pretty.block |
|
120 |
[Pretty.str "Reason: Cannot prove this: ", Pretty.brk 2, Syntax.pretty_term ctxt goal])), |
|
121 |
"Most probably a relator_eq rule for one of the involved types is missing."] |
|
122 |
in |
|
123 |
error error_msg |
|
124 |
end |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
125 |
in |
60239 | 126 |
fun define_pcr_cr_eq (config: config) lthy pcr_rel_def = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
127 |
let |
59582 | 128 |
val lhs = (Thm.term_of o Thm.lhs_of) pcr_rel_def |
60239 | 129 |
val qty_name = |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
130 |
Binding.name |
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
131 |
(Long_Name.base_name (dest_Type_name (List.last (binder_types (fastype_of lhs))))) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
132 |
val args = (snd o strip_comb) lhs |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
133 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
134 |
fun make_inst var ctxt = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
135 |
let |
60784 | 136 |
val typ = snd (relation_types (#2 (dest_Var var))) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
137 |
val sort = Type.sort_of_atyp typ |
59630 | 138 |
val (fresh_var, ctxt') = yield_singleton Variable.invent_types sort ctxt |
60784 | 139 |
val inst = (#1 (dest_Var var), Thm.cterm_of ctxt' (HOLogic.eq_const (TFree fresh_var))) |
140 |
in (inst, ctxt') end |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
141 |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
142 |
val (args_inst, args_ctxt) = fold_map make_inst args lthy |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
143 |
val pcr_cr_eq = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
144 |
pcr_rel_def |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
145 |
|> infer_instantiate args_ctxt args_inst |
52883
0a7c97c76f46
expand equalities in the transfer relation in transfer_prover if the relation doesn't follow the functional structure
kuncar
parents:
51994
diff
changeset
|
146 |
|> Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv |
74545 | 147 |
(Conv.bottom_rewrs_conv (Transfer.get_relator_eq args_ctxt) args_ctxt))) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
148 |
in |
59582 | 149 |
case (Thm.term_of o Thm.rhs_of) pcr_cr_eq of |
80673 | 150 |
\<^Const_>\<open>relcompp _ _ _ for \<^Const_>\<open>HOL.eq _\<close> _\<close> => |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
151 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
152 |
val thm = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
153 |
pcr_cr_eq |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
154 |
|> Conv.fconv_rule (Conv.arg_conv (Conv.rewr_conv eq_OO_meta)) |
67710
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents:
66251
diff
changeset
|
155 |
|> HOLogic.mk_obj_eq |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
156 |
|> singleton (Variable.export args_ctxt lthy) |
70321 | 157 |
val lthy' = lthy |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
158 |
|> #notes config ? |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
159 |
(Local_Theory.note ((Binding.qualify_name true qty_name "pcr_cr_eq", []), [thm]) #> #2) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
160 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
161 |
(thm, lthy') |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
162 |
end |
80673 | 163 |
| \<^Const_>\<open>relcompp _ _ _ for t _\<close> => print_generate_pcr_cr_eq_error args_ctxt t |
164 |
| _ => error "generate_pcr_cr_eq: implementation error" |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
165 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
166 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
167 |
|
59487
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
59458
diff
changeset
|
168 |
fun define_code_constr quot_thm lthy = |
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
169 |
let |
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47943
diff
changeset
|
170 |
val abs = quot_thm_abs quot_thm |
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
171 |
in |
59487
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
59458
diff
changeset
|
172 |
if is_Const abs then |
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
173 |
let |
59487
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
59458
diff
changeset
|
174 |
val (fixed_abs, lthy') = yield_singleton Variable.importT_terms abs lthy |
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
175 |
in |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
63395
diff
changeset
|
176 |
Local_Theory.background_theory |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
63395
diff
changeset
|
177 |
(Code.declare_datatype_global [dest_Const fixed_abs]) lthy' |
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
178 |
end |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
179 |
else |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
180 |
lthy |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
181 |
end |
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
182 |
|
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
63395
diff
changeset
|
183 |
fun define_abs_type quot_thm = |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
63395
diff
changeset
|
184 |
Lifting_Def.can_generate_code_cert quot_thm ? |
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
63395
diff
changeset
|
185 |
Code.declare_abstype (quot_thm RS @{thm Quotient_abs_rep}); |
47308 | 186 |
|
53651 | 187 |
local |
188 |
exception QUOT_ERROR of Pretty.T list |
|
189 |
in |
|
47379
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
190 |
fun quot_thm_sanity_check ctxt quot_thm = |
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
191 |
let |
53651 | 192 |
val _ = |
59582 | 193 |
if (Thm.nprems_of quot_thm > 0) then |
53651 | 194 |
raise QUOT_ERROR [Pretty.block |
195 |
[Pretty.str "The Quotient theorem has extra assumptions:", |
|
196 |
Pretty.brk 1, |
|
61268 | 197 |
Thm.pretty_thm ctxt quot_thm]] |
53651 | 198 |
else () |
59582 | 199 |
val _ = quot_thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_Quotient |
53651 | 200 |
handle TERM _ => raise QUOT_ERROR |
201 |
[Pretty.block |
|
202 |
[Pretty.str "The Quotient theorem is not of the right form:", |
|
203 |
Pretty.brk 1, |
|
61268 | 204 |
Thm.pretty_thm ctxt quot_thm]] |
47379
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
205 |
val ((_, [quot_thm_fixed]), ctxt') = Variable.importT [quot_thm] ctxt |
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47943
diff
changeset
|
206 |
val (rty, qty) = quot_thm_rty_qty quot_thm_fixed |
47379
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
207 |
val rty_tfreesT = Term.add_tfree_namesT rty [] |
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
208 |
val qty_tfreesT = Term.add_tfree_namesT qty [] |
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
209 |
val extra_rty_tfrees = |
47545 | 210 |
case subtract (op =) qty_tfreesT rty_tfreesT of |
47379
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
211 |
[] => [] |
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
212 |
| extras => [Pretty.block ([Pretty.str "Extra variables in the raw type:", |
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
213 |
Pretty.brk 1] @ |
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
214 |
((Pretty.commas o map (Pretty.str o quote)) extras) @ |
47545 | 215 |
[Pretty.str "."])] |
80673 | 216 |
val not_type_constr = |
217 |
if is_Type qty then [] |
|
218 |
else [Pretty.block [Pretty.str "The quotient type ", |
|
219 |
Pretty.quote (Syntax.pretty_typ ctxt' qty), |
|
220 |
Pretty.brk 1, |
|
221 |
Pretty.str "is not a type constructor."]] |
|
47379
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
222 |
val errs = extra_rty_tfrees @ not_type_constr |
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
223 |
in |
53651 | 224 |
if null errs then () else raise QUOT_ERROR errs |
225 |
end |
|
226 |
handle QUOT_ERROR errs => error (cat_lines (["Sanity check of the quotient theorem failed:"] |
|
227 |
@ (map (Pretty.string_of o Pretty.item o single) errs))) |
|
228 |
end |
|
229 |
||
230 |
fun lifting_bundle qty_full_name qinfo lthy = |
|
231 |
let |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
232 |
val thy = Proof_Context.theory_of lthy |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
233 |
|
63003 | 234 |
val binding = |
235 |
Binding.qualify_name true (qty_full_name |> Long_Name.base_name |> Binding.name) "lifting" |
|
53651 | 236 |
val morphed_binding = Morphism.binding (Local_Theory.target_morphism lthy) binding |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
237 |
val bundle_name = |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
238 |
Name_Space.full_name (Name_Space.naming_of (Context.Theory thy)) morphed_binding |
53651 | 239 |
fun phi_qinfo phi = Lifting_Info.transform_quotient phi qinfo |
240 |
||
241 |
val dummy_thm = Thm.transfer thy Drule.dummy_thm |
|
242 |
val restore_lifting_att = |
|
78090 | 243 |
([dummy_thm], [map Token.make_string0 ["Lifting.lifting_restore_internal", bundle_name]]) |
53651 | 244 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
245 |
lthy |
78095 | 246 |
|> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
247 |
(fn phi => Lifting_Info.init_restore_data bundle_name (phi_qinfo phi)) |
81106 | 248 |
|> Bundle.bundle {open_bundle = false} ((binding, [restore_lifting_att])) [] |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
249 |
|> pair binding |
47379
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
250 |
end |
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
251 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
252 |
fun setup_lifting_infr config quot_thm opt_reflp_thm lthy = |
47308 | 253 |
let |
47379
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
kuncar
parents:
47361
diff
changeset
|
254 |
val _ = quot_thm_sanity_check lthy quot_thm |
57663
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
kuncar
parents:
56519
diff
changeset
|
255 |
val (_, qty) = quot_thm_rty_qty quot_thm |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
256 |
val (pcrel_def, lthy1) = define_pcrel config (quot_thm_crel quot_thm) lthy |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
257 |
(**) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
258 |
val pcrel_def = Option.map (Morphism.thm (Local_Theory.target_morphism lthy1)) pcrel_def |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
259 |
(**) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
260 |
val (pcr_cr_eq, lthy2) = |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
261 |
case pcrel_def of |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
262 |
SOME pcrel_def => apfst SOME (define_pcr_cr_eq config lthy1 pcrel_def) |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
263 |
| NONE => (NONE, lthy1) |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
52883
diff
changeset
|
264 |
val pcr_info = case pcrel_def of |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
265 |
SOME pcrel_def => SOME { pcrel_def = pcrel_def, pcr_cr_eq = the pcr_cr_eq } |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
266 |
| NONE => NONE |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
52883
diff
changeset
|
267 |
val quotients = { quot_thm = quot_thm, pcr_info = pcr_info } |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
268 |
val qty_full_name = dest_Type_name qty |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
52883
diff
changeset
|
269 |
fun quot_info phi = Lifting_Info.transform_quotient phi quotients |
78095 | 270 |
val reflexivity_rule_attr = Attrib.internal \<^here> (K Lifting_Info.add_reflexivity_rule_attribute) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
271 |
val lthy3 = |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
272 |
case opt_reflp_thm of |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
273 |
SOME reflp_thm => |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
274 |
lthy2 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
275 |
|> (#2 oo Local_Theory.note) |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
276 |
((Binding.empty, [reflexivity_rule_attr]), [reflp_thm RS @{thm reflp_ge_eq}]) |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
277 |
|> define_code_constr quot_thm |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
278 |
| NONE => lthy2 |> define_abs_type quot_thm |
47308 | 279 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
280 |
lthy3 |
78095 | 281 |
|> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} |
47308 | 282 |
(fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi)) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
283 |
|> lifting_bundle qty_full_name quotients |
47308 | 284 |
end |
285 |
||
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
286 |
local |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
287 |
fun importT_inst_exclude exclude ts ctxt = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
288 |
let |
53651 | 289 |
val tvars = rev (subtract op= exclude (fold Term.add_tvars ts [])) |
290 |
val (tfrees, ctxt') = Variable.invent_types (map #2 tvars) ctxt |
|
74266 | 291 |
in (TVars.make (tvars ~~ map TFree tfrees), ctxt') end |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
292 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
293 |
fun import_inst_exclude exclude ts ctxt = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
294 |
let |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
295 |
val excludeT = fold (Term.add_tvarsT o snd) exclude [] |
53651 | 296 |
val (instT, ctxt') = importT_inst_exclude excludeT ts ctxt |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
297 |
val vars = map (apsnd (Term_Subst.instantiateT instT)) |
53651 | 298 |
(rev (subtract op= exclude (fold Term.add_vars ts []))) |
299 |
val (xs, ctxt'') = Variable.variant_fixes (map (#1 o #1) vars) ctxt' |
|
74266 | 300 |
val inst = Vars.make (vars ~~ map Free (xs ~~ map #2 vars)) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
301 |
in ((instT, inst), ctxt'') end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
302 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
303 |
fun import_terms_exclude exclude ts ctxt = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
304 |
let val (inst, ctxt') = import_inst_exclude exclude ts ctxt |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
305 |
in (map (Term_Subst.instantiate inst) ts, ctxt') end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
306 |
in |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
307 |
fun reduce_goal not_fix goal tac ctxt = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
308 |
let |
59630 | 309 |
val (fixed_goal, ctxt') = yield_singleton (import_terms_exclude not_fix) goal ctxt |
310 |
val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal) |
|
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
311 |
in |
59630 | 312 |
(singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE tac init_goal)) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
313 |
end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
314 |
end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
315 |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
316 |
local |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56257
diff
changeset
|
317 |
val OO_rules = @{thms left_total_OO left_unique_OO right_total_OO right_unique_OO bi_total_OO |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56257
diff
changeset
|
318 |
bi_unique_OO} |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
319 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
320 |
fun parametrize_class_constraint ctxt0 pcr_def constraint = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
321 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
322 |
fun generate_transfer_rule pcr_def constraint goal ctxt = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
323 |
let |
59630 | 324 |
val (fixed_goal, ctxt') = yield_singleton (Variable.import_terms true) goal ctxt |
325 |
val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal) |
|
326 |
val rules = Transfer.get_transfer_raw ctxt' |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
327 |
val rules = constraint :: OO_rules @ rules |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59487
diff
changeset
|
328 |
val tac = |
63170 | 329 |
K (Local_Defs.unfold0_tac ctxt' [pcr_def]) THEN' REPEAT_ALL_NEW (resolve_tac ctxt' rules) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
330 |
in |
59630 | 331 |
(singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE (tac 1) init_goal)) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
332 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
333 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
334 |
fun make_goal pcr_def constr = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
335 |
let |
80636
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
336 |
val pred_name = dest_Const_name (strip_args 1 (HOLogic.dest_Trueprop (Thm.prop_of constr))) |
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
337 |
val arg = fst (Logic.dest_equals (Thm.prop_of pcr_def)) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
338 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
339 |
HOLogic.mk_Trueprop ((Const (pred_name, (fastype_of arg) --> HOLogic.boolT)) $ arg) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
340 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
341 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
342 |
val check_assms = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
343 |
let |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56257
diff
changeset
|
344 |
val right_names = ["right_total", "right_unique", "left_total", "left_unique", "bi_total", |
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56257
diff
changeset
|
345 |
"bi_unique"] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
346 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
347 |
fun is_right_name name = member op= right_names (Long_Name.base_name name) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
348 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
349 |
fun is_trivial_assm (Const (name, _) $ Var (_, _)) = is_right_name name |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
350 |
| is_trivial_assm (Const (name, _) $ Free (_, _)) = is_right_name name |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
351 |
| is_trivial_assm _ = false |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
352 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
353 |
fn thm => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
354 |
let |
59582 | 355 |
val prems = map HOLogic.dest_Trueprop (Thm.prems_of thm) |
356 |
val thm_name = |
|
80636
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
357 |
Long_Name.base_name |
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
358 |
(dest_Const_name (strip_args 1 (HOLogic.dest_Trueprop (Thm.concl_of thm)))) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
359 |
val non_trivial_assms = filter_out is_trivial_assm prems |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
360 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
361 |
if null non_trivial_assms then () |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
362 |
else |
61876 | 363 |
Pretty.block ([Pretty.str "Non-trivial assumptions in ", |
364 |
Pretty.str thm_name, |
|
365 |
Pretty.str " transfer rule found:", |
|
366 |
Pretty.brk 1] @ |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
367 |
Pretty.commas (map (Syntax.pretty_term ctxt0) non_trivial_assms)) |
61876 | 368 |
|> Pretty.string_of |
369 |
|> warning |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
370 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
371 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
372 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
373 |
val goal = make_goal pcr_def constraint |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
374 |
val thm = generate_transfer_rule pcr_def constraint goal ctxt0 |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
375 |
val _ = check_assms thm |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
376 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
377 |
thm |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
378 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
379 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
380 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
381 |
local |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
382 |
val id_unfold = (Conv.rewr_conv (mk_meta_eq @{thm id_def})) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
383 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
384 |
fun generate_parametric_id lthy rty id_transfer_rule = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
385 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
386 |
(* it doesn't raise an exception because it would have already raised it in define_pcrel *) |
59630 | 387 |
val (quot_thm, _, ctxt') = Lifting_Term.prove_param_quot_thm lthy rty |
388 |
val parametrized_relator = |
|
389 |
singleton (Variable.export_terms ctxt' lthy) (quot_thm_crel quot_thm) |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
390 |
val id_transfer = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
391 |
@{thm id_transfer} |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
392 |
|> Thm.incr_indexes (Term.maxidx_of_term parametrized_relator + 1) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
393 |
|> Conv.fconv_rule(HOLogic.Trueprop_conv (Conv.arg_conv id_unfold then_conv Conv.arg1_conv id_unfold)) |
60784 | 394 |
val var = hd (Term.add_vars (Thm.prop_of id_transfer) []) |
395 |
val inst = [(#1 var, Thm.cterm_of lthy parametrized_relator)] |
|
396 |
val id_par_thm = infer_instantiate lthy inst id_transfer |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
397 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
398 |
Lifting_Def.generate_parametric_transfer_rule lthy id_transfer_rule id_par_thm |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
399 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
400 |
handle Lifting_Term.MERGE_TRANSFER_REL msg => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
401 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
402 |
val error_msg = cat_lines |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
403 |
["Generation of a parametric transfer rule for the abs. or the rep. function failed.", |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
404 |
"A non-parametric version will be used.", |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
405 |
(Pretty.string_of (Pretty.block |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
406 |
[Pretty.str "Reason:", Pretty.brk 2, msg]))] |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
407 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
408 |
(warning error_msg; id_transfer_rule) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
409 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
410 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
411 |
|
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
412 |
local |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
413 |
fun rewrite_first_Domainp_arg rewr_thm thm = Conv.fconv_rule (Conv.concl_conv ~1 (HOLogic.Trueprop_conv |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
414 |
(Conv.arg1_conv (Conv.arg_conv (Conv.rewr_conv rewr_thm))))) thm |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
415 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
416 |
fun fold_Domainp_pcrel pcrel_def thm = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
417 |
let |
59582 | 418 |
val ct = |
419 |
thm |> Thm.cprop_of |> Drule.strip_imp_concl |
|
420 |
|> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg |
|
59586 | 421 |
val pcrel_def = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) pcrel_def |
70473 | 422 |
val thm' = Thm.instantiate (Thm.match (ct, Thm.rhs_of pcrel_def)) thm |
53651 | 423 |
handle Pattern.MATCH => raise CTERM ("fold_Domainp_pcrel", [ct, Thm.rhs_of pcrel_def]) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
424 |
in |
70473 | 425 |
rewrite_first_Domainp_arg (Thm.symmetric pcrel_def) thm' |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
426 |
end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
427 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
428 |
fun reduce_Domainp ctxt rules thm = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
429 |
let |
59582 | 430 |
val goal = thm |> Thm.prems_of |> hd |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
431 |
val var = goal |> HOLogic.dest_Trueprop |> dest_comb |> snd |> dest_Var |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59487
diff
changeset
|
432 |
val reduced_assm = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59487
diff
changeset
|
433 |
reduce_goal [var] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
434 |
in |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
435 |
reduced_assm RS thm |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
436 |
end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
437 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
438 |
fun parametrize_domain dom_thm (pcr_info : Lifting_Info.pcr) ctxt0 = |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
439 |
let |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
440 |
fun reduce_first_assm ctxt rules thm = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
441 |
let |
59582 | 442 |
val goal = thm |> Thm.prems_of |> hd |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59487
diff
changeset
|
443 |
val reduced_assm = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59487
diff
changeset
|
444 |
reduce_goal [] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
445 |
in |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
446 |
reduced_assm RS thm |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
447 |
end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
448 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
52883
diff
changeset
|
449 |
val pcr_cr_met_eq = #pcr_cr_eq pcr_info RS @{thm eq_reflection} |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
450 |
val pcr_Domainp_eq = rewrite_first_Domainp_arg (Thm.symmetric pcr_cr_met_eq) dom_thm |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
52883
diff
changeset
|
451 |
val pcrel_def = #pcrel_def pcr_info |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
452 |
val pcr_Domainp_par_left_total = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
453 |
(dom_thm RS @{thm pcr_Domainp_par_left_total}) |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
454 |
|> fold_Domainp_pcrel pcrel_def |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
455 |
|> reduce_first_assm ctxt0 (Lifting_Info.get_reflexivity_rules ctxt0) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
456 |
val pcr_Domainp_par = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
457 |
(dom_thm RS @{thm pcr_Domainp_par}) |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
458 |
|> fold_Domainp_pcrel pcrel_def |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
459 |
|> reduce_Domainp ctxt0 (Transfer.get_relator_domain ctxt0) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
460 |
val pcr_Domainp = |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
461 |
(dom_thm RS @{thm pcr_Domainp}) |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
462 |
|> fold_Domainp_pcrel pcrel_def |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
463 |
val thms = |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
464 |
[("domain", [pcr_Domainp], @{attributes [transfer_domain_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
465 |
("domain_par", [pcr_Domainp_par], @{attributes [transfer_domain_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
466 |
("domain_par_left_total", [pcr_Domainp_par_left_total], @{attributes [transfer_domain_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
467 |
("domain_eq", [pcr_Domainp_eq], @{attributes [transfer_domain_rule]})] |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
468 |
in |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
469 |
thms |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
470 |
end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
471 |
|
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56257
diff
changeset
|
472 |
fun parametrize_total_domain left_total pcrel_def ctxt = |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
473 |
let |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
474 |
val thm = |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56257
diff
changeset
|
475 |
(left_total RS @{thm pcr_Domainp_total}) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
476 |
|> fold_Domainp_pcrel pcrel_def |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
477 |
|> reduce_Domainp ctxt (Transfer.get_relator_domain ctxt) |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
478 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
479 |
[("domain", [thm], @{attributes [transfer_domain_rule]})] |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
480 |
end |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
481 |
|
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
482 |
end |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
483 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
484 |
fun get_pcrel_info ctxt qty_full_name = |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
52883
diff
changeset
|
485 |
#pcr_info (the (Lifting_Info.lookup_quotients ctxt qty_full_name)) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
486 |
|
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
487 |
fun get_Domainp_thm quot_thm = |
56519
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
kuncar
parents:
56518
diff
changeset
|
488 |
the (get_first (try(curry op RS quot_thm)) [@{thm eq_onp_to_Domainp}, @{thm Quotient_to_Domainp}]) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
489 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
490 |
fun notes names thms = |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
491 |
let |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
492 |
val notes = |
63352 | 493 |
if names then map (fn (name, thms, attrs) => ((name, []), [(thms, attrs)])) thms |
494 |
else map_filter (fn (_, thms, attrs) => if null attrs then NONE |
|
495 |
else SOME (Binding.empty_atts, [(thms, attrs)])) thms |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
496 |
in |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
497 |
Local_Theory.notes notes #> snd |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
498 |
end |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
499 |
|
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
500 |
fun map_thms map_name map_thm thms = |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
501 |
map (fn (name, thms, attr) => (map_name name, map map_thm thms, attr)) thms |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
502 |
|
47852
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
503 |
(* |
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
504 |
Sets up the Lifting package by a quotient theorem. |
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
505 |
|
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
506 |
quot_thm - a quotient theorem (Quotient R Abs Rep T) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
507 |
opt_reflp_thm - a theorem saying that a relation from quot_thm is reflexive |
47852
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
508 |
(in the form "reflp R") |
54335
03b10317ba78
update documentation of important public ML functions in Lifting
kuncar
parents:
54333
diff
changeset
|
509 |
opt_par_thm - a parametricity theorem for R |
47852
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
510 |
*) |
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
511 |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
512 |
fun setup_by_quotient (config: config) quot_thm opt_reflp_thm opt_par_thm lthy0 = |
47521
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
513 |
let |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
514 |
(**) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
515 |
val quot_thm = Morphism.thm (Local_Theory.target_morphism lthy0) quot_thm |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
516 |
(**) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
517 |
val (rty, qty) = quot_thm_rty_qty quot_thm |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
518 |
val induct_attr = Attrib.internal \<^here> (K (Induct.induct_type (dest_Type_name qty))) |
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
519 |
val qty_full_name = dest_Type_name qty |
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
520 |
val qty_name = Binding.name (Long_Name.base_name qty_full_name) |
63003 | 521 |
val qualify = Binding.qualify_name true qty_name |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
522 |
val notes1 = case opt_reflp_thm of |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
523 |
SOME reflp_thm => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
524 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
525 |
val thms = |
63003 | 526 |
[("abs_induct", @{thms Quotient_total_abs_induct}, [induct_attr]), |
527 |
("abs_eq_iff", @{thms Quotient_total_abs_eq_iff}, [])] |
|
528 |
in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
529 |
| NONE => |
63003 | 530 |
let val thms = [("abs_induct", @{thms Quotient_abs_induct}, [induct_attr])] |
531 |
in map_thms qualify (fn thm => quot_thm RS thm) thms end |
|
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
532 |
val dom_thm = get_Domainp_thm quot_thm |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
533 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
534 |
fun setup_transfer_rules_nonpar notes = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
535 |
let |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
536 |
val notes1 = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
537 |
case opt_reflp_thm of |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
538 |
SOME reflp_thm => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
539 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
540 |
val thms = |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
541 |
[("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
542 |
("left_total", @{thms Quotient_left_total}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
543 |
("bi_total", @{thms Quotient_bi_total}, @{attributes [transfer_rule]})] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
544 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
545 |
map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
546 |
end |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
547 |
| NONE => map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})] |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
548 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
549 |
val notes2 = map_thms qualify (fn thm => quot_thm RS thm) |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
550 |
[("rel_eq_transfer", @{thms Quotient_rel_eq_transfer}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
551 |
("right_unique", @{thms Quotient_right_unique}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
552 |
("right_total", @{thms Quotient_right_total}, @{attributes [transfer_rule]})] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
553 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
554 |
notes2 @ notes1 @ notes |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
555 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
556 |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
557 |
fun generate_parametric_rel_eq ctxt transfer_rule opt_param_thm = |
70322 | 558 |
(case opt_param_thm of |
559 |
NONE => transfer_rule |
|
560 |
| SOME param_thm => |
|
561 |
(Lifting_Def.generate_parametric_transfer_rule ctxt transfer_rule param_thm |
|
562 |
handle Lifting_Term.MERGE_TRANSFER_REL msg => |
|
563 |
error ("Generation of a parametric transfer rule for the quotient relation failed:\n" |
|
564 |
^ Pretty.string_of msg))) |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
565 |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
566 |
fun setup_transfer_rules_par ctxt notes = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
567 |
let |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
568 |
val pcrel_info = the (get_pcrel_info ctxt qty_full_name) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
569 |
val pcrel_def = #pcrel_def pcrel_info |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
570 |
val notes1 = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
571 |
case opt_reflp_thm of |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
572 |
SOME reflp_thm => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
573 |
let |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56257
diff
changeset
|
574 |
val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total}) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
575 |
val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
576 |
val domain_thms = parametrize_total_domain left_total pcrel_def ctxt |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
577 |
val id_abs_transfer = generate_parametric_id ctxt rty |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
578 |
(Lifting_Term.parametrize_transfer_rule ctxt |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
579 |
([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer})) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
580 |
val left_total = parametrize_class_constraint ctxt pcrel_def left_total |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
581 |
val bi_total = parametrize_class_constraint ctxt pcrel_def bi_total |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
582 |
val thms = |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
583 |
[("id_abs_transfer", [id_abs_transfer], @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
584 |
("left_total", [left_total], @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
585 |
("bi_total", [bi_total], @{attributes [transfer_rule]})] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
586 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
587 |
map_thms qualify I thms @ map_thms qualify I domain_thms |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
588 |
end |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
589 |
| NONE => |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
590 |
let |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
591 |
val thms = parametrize_domain dom_thm pcrel_info ctxt |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
592 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
593 |
map_thms qualify I thms |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
594 |
end |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
595 |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
596 |
val rel_eq_transfer = generate_parametric_rel_eq ctxt |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
597 |
(Lifting_Term.parametrize_transfer_rule ctxt (quot_thm RS @{thm Quotient_rel_eq_transfer})) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
598 |
opt_par_thm |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
599 |
val right_unique = parametrize_class_constraint ctxt pcrel_def |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
600 |
(quot_thm RS @{thm Quotient_right_unique}) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
601 |
val right_total = parametrize_class_constraint ctxt pcrel_def |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
602 |
(quot_thm RS @{thm Quotient_right_total}) |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
603 |
val notes2 = map_thms qualify I |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
604 |
[("rel_eq_transfer", [rel_eq_transfer], @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
605 |
("right_unique", [right_unique], @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
606 |
("right_total", [right_total], @{attributes [transfer_rule]})] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
607 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
608 |
notes2 @ notes1 @ notes |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
609 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
610 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
611 |
fun setup_rules lthy = |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
612 |
let |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
613 |
val thms = |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
614 |
if is_some (get_pcrel_info lthy qty_full_name) |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
615 |
then setup_transfer_rules_par lthy notes1 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
616 |
else setup_transfer_rules_nonpar notes1 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
617 |
in notes (#notes config) thms lthy end |
47521
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
618 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
619 |
lthy0 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
620 |
|> setup_lifting_infr config quot_thm opt_reflp_thm |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
621 |
||> setup_rules |
47521
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
622 |
end |
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
623 |
|
47852
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
624 |
(* |
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
625 |
Sets up the Lifting package by a typedef theorem. |
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
626 |
|
47937
70375fa2679d
generate abs_eq, use it as a code equation for total quotients; no_abs_code renamed to no_code; added no_code for quotient_type command
kuncar
parents:
47936
diff
changeset
|
627 |
gen_code - flag if an abstract type given by typedef_thm should be registred |
47852
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
628 |
as an abstract type in the code generator |
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
629 |
typedef_thm - a typedef theorem (type_definition Rep Abs S) |
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
630 |
*) |
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
kuncar
parents:
47779
diff
changeset
|
631 |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
632 |
fun setup_by_typedef_thm config typedef_thm lthy0 = |
47308 | 633 |
let |
59582 | 634 |
val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o Thm.prop_of) typedef_thm |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
635 |
val (T_def, lthy1) = define_crel config rep_fun lthy0 |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
636 |
(**) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
637 |
val T_def = Morphism.thm (Local_Theory.target_morphism lthy1) T_def |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
638 |
(**) |
80673 | 639 |
val quot_thm = |
640 |
case typedef_set of |
|
641 |
\<^Const_>\<open>top _\<close> => [typedef_thm, T_def] MRSL @{thm UNIV_typedef_to_Quotient} |
|
642 |
| \<^Const_>\<open>Collect _ for \<open>Abs _\<close>\<close> => [typedef_thm, T_def] MRSL @{thm open_typedef_to_Quotient} |
|
643 |
| _ => [typedef_thm, T_def] MRSL @{thm typedef_to_Quotient} |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
644 |
val (rty, qty) = quot_thm_rty_qty quot_thm |
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
645 |
val qty_full_name = dest_Type_name qty |
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
646 |
val qty_name = Binding.name (Long_Name.base_name qty_full_name) |
63003 | 647 |
val qualify = Binding.qualify_name true qty_name |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
648 |
val opt_reflp_thm = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
649 |
case typedef_set of |
80673 | 650 |
\<^Const_>\<open>top _\<close> => |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
651 |
SOME ((typedef_thm RS @{thm UNIV_typedef_to_equivp}) RS @{thm equivp_reflp2}) |
80673 | 652 |
| _ => NONE |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
653 |
val dom_thm = get_Domainp_thm quot_thm |
47308 | 654 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
655 |
fun setup_transfer_rules_nonpar notes = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
656 |
let |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
657 |
val notes1 = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
658 |
case opt_reflp_thm of |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
659 |
SOME reflp_thm => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
660 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
661 |
val thms = |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
662 |
[("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
663 |
("left_total", @{thms Quotient_left_total}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
664 |
("bi_total", @{thms Quotient_bi_total}, @{attributes [transfer_rule]})] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
665 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
666 |
map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
667 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
668 |
| NONE => |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
669 |
map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
670 |
val thms = |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
671 |
[("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
672 |
("left_unique", @{thms typedef_left_unique}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
673 |
("right_unique", @{thms typedef_right_unique}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
674 |
("right_total", @{thms typedef_right_total}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
675 |
("bi_unique", @{thms typedef_bi_unique}, @{attributes [transfer_rule]})] |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
676 |
in |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
677 |
map_thms qualify (fn thm => [typedef_thm, T_def] MRSL thm) thms @ notes1 @ notes |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
678 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
679 |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
680 |
fun setup_transfer_rules_par ctxt notes = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
681 |
let |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
682 |
val pcrel_info = (the (get_pcrel_info ctxt qty_full_name)) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
683 |
val pcrel_def = #pcrel_def pcrel_info |
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
684 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
685 |
val notes1 = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
686 |
case opt_reflp_thm of |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
687 |
SOME reflp_thm => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
688 |
let |
56518
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
kuncar
parents:
56257
diff
changeset
|
689 |
val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total}) |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
690 |
val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
691 |
val domain_thms = parametrize_total_domain left_total pcrel_def ctxt |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
692 |
val left_total = parametrize_class_constraint ctxt pcrel_def left_total |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
693 |
val bi_total = parametrize_class_constraint ctxt pcrel_def bi_total |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
694 |
val id_abs_transfer = generate_parametric_id ctxt rty |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
695 |
(Lifting_Term.parametrize_transfer_rule ctxt |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
696 |
([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer})) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
697 |
val thms = |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
698 |
[("left_total", [left_total], @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
699 |
("bi_total", [bi_total], @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
700 |
("id_abs_transfer",[id_abs_transfer], @{attributes [transfer_rule]})] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
701 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
702 |
map_thms qualify I thms @ map_thms qualify I domain_thms |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
703 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
704 |
| NONE => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
705 |
let |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
706 |
val thms = parametrize_domain dom_thm pcrel_info ctxt |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
707 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
708 |
map_thms qualify I thms |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
709 |
end |
51956
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
kuncar
parents:
51374
diff
changeset
|
710 |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
711 |
val notes2 = map_thms qualify (fn thm => generate_parametric_id ctxt rty |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
712 |
(Lifting_Term.parametrize_transfer_rule ctxt ([typedef_thm, T_def] MRSL thm))) |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
713 |
[("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]})]; |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
714 |
val notes3 = |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
715 |
map_thms qualify |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
716 |
(fn thm => parametrize_class_constraint ctxt pcrel_def ([typedef_thm, T_def] MRSL thm)) |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
717 |
[("left_unique", @{thms typedef_left_unique}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
718 |
("right_unique", @{thms typedef_right_unique},@{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
719 |
("bi_unique", @{thms typedef_bi_unique}, @{attributes [transfer_rule]}), |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
720 |
("right_total", @{thms typedef_right_total}, @{attributes [transfer_rule]})] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
721 |
in |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
722 |
notes3 @ notes2 @ notes1 @ notes |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
723 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
724 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
725 |
val notes1 = [(Binding.prefix_name "Quotient_" qty_name, [quot_thm], [])] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
726 |
|
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
727 |
fun setup_rules lthy = |
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
728 |
let |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
729 |
val thms = |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
730 |
if is_some (get_pcrel_info lthy qty_full_name) |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
731 |
then setup_transfer_rules_par lthy notes1 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
732 |
else setup_transfer_rules_nonpar notes1 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
733 |
in notes (#notes config) thms lthy end |
47308 | 734 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
735 |
lthy1 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
736 |
|> setup_lifting_infr config quot_thm opt_reflp_thm |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
737 |
||> setup_rules |
47308 | 738 |
end |
739 |
||
59487
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
59458
diff
changeset
|
740 |
fun setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm lthy = |
47521
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
741 |
let |
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
742 |
val input_thm = singleton (Attrib.eval_thms lthy) xthm |
59582 | 743 |
val input_term = (HOLogic.dest_Trueprop o Thm.prop_of) input_thm |
47566
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
744 |
handle TERM _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." |
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
745 |
|
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
746 |
fun sanity_check_reflp_thm reflp_thm = |
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
747 |
let |
59582 | 748 |
val reflp_tm = (HOLogic.dest_Trueprop o Thm.prop_of) reflp_thm |
47566
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
749 |
handle TERM _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." |
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
750 |
in |
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
751 |
case reflp_tm of |
75503
e5d88927e017
introduced predicate reflp_on and redefined reflp to be an abbreviation
desharna
parents:
74545
diff
changeset
|
752 |
\<^Const_>\<open>reflp_on _ for \<^Const>\<open>top_class.top _\<close> _\<close> => () |
47566
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
753 |
| _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." |
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
754 |
end |
55487
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
755 |
|
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
756 |
fun check_qty qty = if not (is_Type qty) |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
757 |
then error "The abstract type must be a type constructor." |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
758 |
else () |
60225
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
kuncar
parents:
59936
diff
changeset
|
759 |
|
47566
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
760 |
fun setup_quotient () = |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
761 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
762 |
val opt_reflp_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_reflp_xthm |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
763 |
val _ = if is_some opt_reflp_thm then sanity_check_reflp_thm (the opt_reflp_thm) else () |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
764 |
val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm |
55487
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
765 |
val _ = check_qty (snd (quot_thm_rty_qty input_thm)) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
766 |
in |
60231
0daab758e087
don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents:
60226
diff
changeset
|
767 |
setup_by_quotient default_config input_thm opt_reflp_thm opt_par_thm lthy |> snd |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50288
diff
changeset
|
768 |
end |
47566
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
769 |
|
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
kuncar
parents:
47545
diff
changeset
|
770 |
fun setup_typedef () = |
55487
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
771 |
let |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
772 |
val qty = (range_type o fastype_of o hd o get_args 2) input_term |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
773 |
val _ = check_qty qty |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
774 |
in |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
775 |
case opt_reflp_xthm of |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
776 |
SOME _ => error "The reflexivity theorem cannot be specified if the type_definition theorem is used." |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
777 |
| NONE => ( |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
778 |
case opt_par_xthm of |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
779 |
SOME _ => error "The parametricity theorem cannot be specified if the type_definition theorem is used." |
60231
0daab758e087
don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
kuncar
parents:
60226
diff
changeset
|
780 |
| NONE => setup_by_typedef_thm default_config input_thm lthy |> snd |
55487
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
781 |
) |
6380313b8ed5
abstract type must be a type constructor; check it
kuncar
parents:
54335
diff
changeset
|
782 |
end |
47521
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
783 |
in |
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
784 |
case input_term of |
80673 | 785 |
\<^Const_>\<open>Quotient _ _ for _ _ _ _\<close> => setup_quotient () |
786 |
| \<^Const_>\<open>type_definition _ _ for _ _ _\<close> => setup_typedef () |
|
787 |
| _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." |
|
47521
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
kuncar
parents:
47379
diff
changeset
|
788 |
end |
47308 | 789 |
|
790 |
val _ = |
|
69593 | 791 |
Outer_Syntax.local_theory \<^command_keyword>\<open>setup_lifting\<close> |
50214 | 792 |
"setup lifting infrastructure" |
62969 | 793 |
(Parse.thm -- Scan.option Parse.thm |
69593 | 794 |
-- Scan.option (\<^keyword>\<open>parametric\<close> |-- Parse.!!! Parse.thm) >> |
59487
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
59458
diff
changeset
|
795 |
(fn ((xthm, opt_reflp_xthm), opt_par_xthm) => |
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents:
59458
diff
changeset
|
796 |
setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm)) |
53651 | 797 |
|
798 |
(* restoring lifting infrastructure *) |
|
799 |
||
800 |
local |
|
801 |
exception PCR_ERROR of Pretty.T list |
|
802 |
in |
|
803 |
||
804 |
fun lifting_restore_sanity_check ctxt (qinfo:Lifting_Info.quotient) = |
|
805 |
let |
|
806 |
val quot_thm = (#quot_thm qinfo) |
|
807 |
val _ = quot_thm_sanity_check ctxt quot_thm |
|
808 |
val pcr_info_err = |
|
809 |
(case #pcr_info qinfo of |
|
810 |
SOME pcr => |
|
811 |
let |
|
812 |
val pcrel_def = #pcrel_def pcr |
|
813 |
val pcr_cr_eq = #pcr_cr_eq pcr |
|
59582 | 814 |
val (def_lhs, _) = Logic.dest_equals (Thm.prop_of pcrel_def) |
53651 | 815 |
handle TERM _ => raise PCR_ERROR [Pretty.block |
816 |
[Pretty.str "The pcr definiton theorem is not a plain meta equation:", |
|
817 |
Pretty.brk 1, |
|
61268 | 818 |
Thm.pretty_thm ctxt pcrel_def]] |
53651 | 819 |
val pcr_const_def = head_of def_lhs |
59582 | 820 |
val (eq_lhs, eq_rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of pcr_cr_eq)) |
53651 | 821 |
handle TERM _ => raise PCR_ERROR [Pretty.block |
822 |
[Pretty.str "The pcr_cr equation theorem is not a plain equation:", |
|
823 |
Pretty.brk 1, |
|
61268 | 824 |
Thm.pretty_thm ctxt pcr_cr_eq]] |
53651 | 825 |
val (pcr_const_eq, eqs) = strip_comb eq_lhs |
80673 | 826 |
fun is_eq \<^Const_>\<open>HOL.eq _\<close> = true |
53651 | 827 |
| is_eq _ = false |
828 |
fun eq_Const (Const (name1, _)) (Const (name2, _)) = (name1 = name2) |
|
829 |
| eq_Const _ _ = false |
|
830 |
val all_eqs = if not (forall is_eq eqs) then |
|
831 |
[Pretty.block |
|
832 |
[Pretty.str "Arguments of the lhs of the pcr_cr equation theorem are not only equalities:", |
|
833 |
Pretty.brk 1, |
|
61268 | 834 |
Thm.pretty_thm ctxt pcr_cr_eq]] |
53651 | 835 |
else [] |
836 |
val pcr_consts_not_equal = if not (eq_Const pcr_const_def pcr_const_eq) then |
|
837 |
[Pretty.block |
|
838 |
[Pretty.str "Parametrized correspondence relation constants in pcr_def and pcr_cr_eq are not equal:", |
|
839 |
Pretty.brk 1, |
|
840 |
Syntax.pretty_term ctxt pcr_const_def, |
|
841 |
Pretty.brk 1, |
|
842 |
Pretty.str "vs.", |
|
843 |
Pretty.brk 1, |
|
844 |
Syntax.pretty_term ctxt pcr_const_eq]] |
|
845 |
else [] |
|
846 |
val crel = quot_thm_crel quot_thm |
|
847 |
val cr_consts_not_equal = if not (eq_Const crel eq_rhs) then |
|
848 |
[Pretty.block |
|
849 |
[Pretty.str "Correspondence relation constants in the Quotient theorem and pcr_cr_eq are not equal:", |
|
850 |
Pretty.brk 1, |
|
851 |
Syntax.pretty_term ctxt crel, |
|
852 |
Pretty.brk 1, |
|
853 |
Pretty.str "vs.", |
|
854 |
Pretty.brk 1, |
|
855 |
Syntax.pretty_term ctxt eq_rhs]] |
|
856 |
else [] |
|
857 |
in |
|
858 |
all_eqs @ pcr_consts_not_equal @ cr_consts_not_equal |
|
859 |
end |
|
860 |
| NONE => []) |
|
861 |
val errs = pcr_info_err |
|
862 |
in |
|
863 |
if null errs then () else raise PCR_ERROR errs |
|
864 |
end |
|
865 |
handle PCR_ERROR errs => error (cat_lines (["Sanity check failed:"] |
|
866 |
@ (map (Pretty.string_of o Pretty.item o single) errs))) |
|
867 |
end |
|
868 |
||
54335
03b10317ba78
update documentation of important public ML functions in Lifting
kuncar
parents:
54333
diff
changeset
|
869 |
(* |
03b10317ba78
update documentation of important public ML functions in Lifting
kuncar
parents:
54333
diff
changeset
|
870 |
Registers the data in qinfo in the Lifting infrastructure. |
03b10317ba78
update documentation of important public ML functions in Lifting
kuncar
parents:
54333
diff
changeset
|
871 |
*) |
03b10317ba78
update documentation of important public ML functions in Lifting
kuncar
parents:
54333
diff
changeset
|
872 |
|
53651 | 873 |
fun lifting_restore qinfo ctxt = |
874 |
let |
|
875 |
val _ = lifting_restore_sanity_check (Context.proof_of ctxt) qinfo |
|
876 |
val (_, qty) = quot_thm_rty_qty (#quot_thm qinfo) |
|
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
78690
diff
changeset
|
877 |
val qty_full_name = dest_Type_name qty |
53651 | 878 |
val stored_qinfo = Lifting_Info.lookup_quotients (Context.proof_of ctxt) qty_full_name |
879 |
in |
|
880 |
if is_some (stored_qinfo) andalso not (Lifting_Info.quotient_eq (qinfo, (the stored_qinfo))) |
|
881 |
then error (Pretty.string_of |
|
882 |
(Pretty.block |
|
883 |
[Pretty.str "Lifting is already setup for the type", |
|
884 |
Pretty.brk 1, |
|
885 |
Pretty.quote (Syntax.pretty_typ (Context.proof_of ctxt) qty)])) |
|
886 |
else Lifting_Info.update_quotients qty_full_name qinfo ctxt |
|
887 |
end |
|
888 |
||
889 |
val parse_opt_pcr = |
|
890 |
Scan.optional (Attrib.thm -- Attrib.thm >> |
|
891 |
(fn (pcrel_def, pcr_cr_eq) => SOME {pcrel_def = pcrel_def, pcr_cr_eq = pcr_cr_eq})) NONE |
|
892 |
||
893 |
val lifting_restore_attribute_setup = |
|
69593 | 894 |
Attrib.setup \<^binding>\<open>lifting_restore\<close> |
53651 | 895 |
((Attrib.thm -- parse_opt_pcr) >> |
896 |
(fn (quot_thm, opt_pcr) => |
|
897 |
let val qinfo = { quot_thm = quot_thm, pcr_info = opt_pcr} |
|
898 |
in Thm.declaration_attribute (K (lifting_restore qinfo)) end)) |
|
899 |
"restoring lifting infrastructure" |
|
900 |
||
901 |
val _ = Theory.setup lifting_restore_attribute_setup |
|
902 |
||
903 |
fun lifting_restore_internal bundle_name ctxt = |
|
904 |
let |
|
905 |
val restore_info = Lifting_Info.lookup_restore_data (Context.proof_of ctxt) bundle_name |
|
906 |
in |
|
907 |
case restore_info of |
|
908 |
SOME restore_info => |
|
909 |
ctxt |
|
910 |
|> lifting_restore (#quotient restore_info) |
|
911 |
|> fold_rev Transfer.transfer_raw_add (Item_Net.content (#transfer_rules restore_info)) |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
912 |
| NONE => ctxt |
53651 | 913 |
end |
914 |
||
915 |
val lifting_restore_internal_attribute_setup = |
|
69593 | 916 |
Attrib.setup \<^binding>\<open>lifting_restore_internal\<close> |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
61268
diff
changeset
|
917 |
(Scan.lift Parse.string >> |
59084
f982f3072d79
more robust bundle_name: avoid assumptions about identifier, keywords etc.;
wenzelm
parents:
59083
diff
changeset
|
918 |
(fn name => Thm.declaration_attribute (K (lifting_restore_internal name)))) |
53651 | 919 |
"restoring lifting infrastructure; internal attribute; not meant to be used directly by regular users" |
920 |
||
921 |
val _ = Theory.setup lifting_restore_internal_attribute_setup |
|
922 |
||
923 |
(* lifting_forget *) |
|
924 |
||
80673 | 925 |
fun fold_transfer_rel f \<^Const_>\<open>Transfer.Rel _ _ for rel _ _\<close> = f rel |
926 |
| fold_transfer_rel f \<^Const_>\<open>HOL.eq _ for \<^Const_>\<open>Domainp _ _ for rel\<close> _\<close> = f rel |
|
927 |
| fold_transfer_rel f \<^Const_>\<open>right_unique _ _ for rel\<close> = f rel |
|
928 |
| fold_transfer_rel f \<^Const_>\<open>left_unique _ _ for rel\<close> = f rel |
|
929 |
| fold_transfer_rel f \<^Const_>\<open>right_total _ _ for rel\<close> = f rel |
|
930 |
| fold_transfer_rel f \<^Const_>\<open>left_total _ _ for rel\<close> = f rel |
|
931 |
| fold_transfer_rel f \<^Const_>\<open>bi_unique _ _ for rel\<close> = f rel |
|
932 |
| fold_transfer_rel f \<^Const_>\<open>bi_total _ _ for rel\<close> = f rel |
|
69593 | 933 |
| fold_transfer_rel f _ = f \<^term>\<open>undefined\<close> |
53651 | 934 |
|
935 |
fun filter_transfer_rules_by_rel transfer_rel transfer_rules = |
|
936 |
let |
|
80636
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
937 |
val transfer_rel_name = dest_Const_name transfer_rel; |
53651 | 938 |
fun has_transfer_rel thm = |
939 |
let |
|
59582 | 940 |
val concl = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |
53651 | 941 |
in |
942 |
member op= (fold_transfer_rel (fn tm => Term.add_const_names tm []) concl) transfer_rel_name |
|
943 |
end |
|
944 |
handle TERM _ => false |
|
945 |
in |
|
946 |
filter has_transfer_rel transfer_rules |
|
947 |
end |
|
948 |
||
949 |
type restore_data = {quotient : Lifting_Info.quotient, transfer_rules: thm Item_Net.T} |
|
950 |
||
53754 | 951 |
fun get_transfer_rel (qinfo : Lifting_Info.quotient) = |
53651 | 952 |
let |
59582 | 953 |
fun get_pcrel pcr_def = pcr_def |> Thm.concl_of |> Logic.dest_equals |> fst |> head_of |
53651 | 954 |
in |
955 |
if is_some (#pcr_info qinfo) |
|
956 |
then get_pcrel (#pcrel_def (the (#pcr_info qinfo))) |
|
957 |
else quot_thm_crel (#quot_thm qinfo) |
|
958 |
end |
|
959 |
||
960 |
fun pointer_of_bundle_name bundle_name ctxt = |
|
961 |
let |
|
81116
0fb1e2dd4122
first-class support for "unbundle no foobar_syntax" -- avoid redundant "bundle no_foobar_syntax" definitions;
wenzelm
parents:
81106
diff
changeset
|
962 |
val bundle = Bundle.get ctxt (Bundle.check ctxt bundle_name) |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
61268
diff
changeset
|
963 |
fun err () = error "The provided bundle is not a lifting bundle" |
53651 | 964 |
in |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
61268
diff
changeset
|
965 |
(case bundle of |
78690 | 966 |
[(_, [arg_src])] => (Token.read ctxt Parse.string arg_src handle ERROR _ => err ()) |
61814
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
wenzelm
parents:
61268
diff
changeset
|
967 |
| _ => err ()) |
53651 | 968 |
end |
969 |
||
60226
ec23f2a97ba4
publish lifting_forget and lifting_udpate interface
kuncar
parents:
60225
diff
changeset
|
970 |
fun pointer_of_bundle_binding ctxt binding = Name_Space.full_name (Name_Space.naming_of |
ec23f2a97ba4
publish lifting_forget and lifting_udpate interface
kuncar
parents:
60225
diff
changeset
|
971 |
(Context.Theory (Proof_Context.theory_of ctxt))) binding |
ec23f2a97ba4
publish lifting_forget and lifting_udpate interface
kuncar
parents:
60225
diff
changeset
|
972 |
|
53651 | 973 |
fun lifting_forget pointer lthy = |
974 |
let |
|
975 |
fun get_transfer_rules_to_delete qinfo ctxt = |
|
976 |
let |
|
53754 | 977 |
val transfer_rel = get_transfer_rel qinfo |
53651 | 978 |
in |
979 |
filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw ctxt) |
|
980 |
end |
|
981 |
in |
|
982 |
case Lifting_Info.lookup_restore_data lthy pointer of |
|
983 |
SOME restore_info => |
|
984 |
let |
|
985 |
val qinfo = #quotient restore_info |
|
986 |
val quot_thm = #quot_thm qinfo |
|
987 |
val transfer_rules = get_transfer_rules_to_delete qinfo lthy |
|
988 |
in |
|
78095 | 989 |
Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} |
53651 | 990 |
(K (fold (Transfer.transfer_raw_del) transfer_rules #> Lifting_Info.delete_quotients quot_thm)) |
991 |
lthy |
|
992 |
end |
|
993 |
| NONE => error "The lifting bundle refers to non-existent restore data." |
|
994 |
end |
|
995 |
||
996 |
||
997 |
fun lifting_forget_cmd bundle_name lthy = |
|
998 |
lifting_forget (pointer_of_bundle_name bundle_name lthy) lthy |
|
999 |
||
1000 |
||
1001 |
val _ = |
|
69593 | 1002 |
Outer_Syntax.local_theory \<^command_keyword>\<open>lifting_forget\<close> |
53651 | 1003 |
"unsetup Lifting and Transfer for the given lifting bundle" |
69349
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
wenzelm
parents:
67710
diff
changeset
|
1004 |
(Parse.name_position >> lifting_forget_cmd) |
53651 | 1005 |
|
1006 |
(* lifting_update *) |
|
1007 |
||
1008 |
fun update_transfer_rules pointer lthy = |
|
1009 |
let |
|
53754 | 1010 |
fun new_transfer_rules ({ quotient = qinfo, ... }:Lifting_Info.restore_data) lthy = |
53651 | 1011 |
let |
1012 |
val transfer_rel = get_transfer_rel qinfo |
|
1013 |
val transfer_rules = filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw lthy) |
|
1014 |
in |
|
1015 |
fn phi => fold_rev |
|
74152 | 1016 |
(Item_Net.update o Morphism.thm phi) transfer_rules Thm.item_net |
53651 | 1017 |
end |
1018 |
in |
|
1019 |
case Lifting_Info.lookup_restore_data lthy pointer of |
|
1020 |
SOME refresh_data => |
|
78095 | 1021 |
Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} |
53651 | 1022 |
(fn phi => Lifting_Info.add_transfer_rules_in_restore_data pointer |
1023 |
(new_transfer_rules refresh_data lthy phi)) lthy |
|
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
1024 |
| NONE => error "The lifting bundle refers to non-existent restore data." |
53651 | 1025 |
end |
1026 |
||
1027 |
fun lifting_update_cmd bundle_name lthy = |
|
1028 |
update_transfer_rules (pointer_of_bundle_name bundle_name lthy) lthy |
|
1029 |
||
1030 |
val _ = |
|
69593 | 1031 |
Outer_Syntax.local_theory \<^command_keyword>\<open>lifting_update\<close> |
53651 | 1032 |
"add newly introduced transfer rules to a bundle storing the state of Lifting and Transfer" |
69349
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
wenzelm
parents:
67710
diff
changeset
|
1033 |
(Parse.name_position >> lifting_update_cmd) |
53651 | 1034 |
|
1035 |
end |