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