| author | wenzelm | 
| Mon, 31 Dec 2018 12:02:31 +0100 | |
| changeset 69556 | 0a38f23ca4c5 | 
| parent 67710 | cc2db3239932 | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 47308 | 1  | 
(* Title: HOL/Tools/Lifting/lifting_def.ML  | 
2  | 
Author: Ondrej Kuncar  | 
|
3  | 
||
4  | 
Definitions for constants on quotient types.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature LIFTING_DEF =  | 
|
8  | 
sig  | 
|
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
9  | 
datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ  | 
| 60219 | 10  | 
type lift_def  | 
11  | 
val rty_of_lift_def: lift_def -> typ  | 
|
12  | 
val qty_of_lift_def: lift_def -> typ  | 
|
13  | 
val rhs_of_lift_def: lift_def -> term  | 
|
14  | 
val lift_const_of_lift_def: lift_def -> term  | 
|
15  | 
val def_thm_of_lift_def: lift_def -> thm  | 
|
16  | 
val rsp_thm_of_lift_def: lift_def -> thm  | 
|
17  | 
val abs_eq_of_lift_def: lift_def -> thm  | 
|
18  | 
val rep_eq_of_lift_def: lift_def -> thm option  | 
|
| 60227 | 19  | 
val code_eq_of_lift_def: lift_def -> code_eq  | 
| 60219 | 20  | 
val transfer_rules_of_lift_def: lift_def -> thm list  | 
| 60224 | 21  | 
val morph_lift_def: morphism -> lift_def -> lift_def  | 
22  | 
val inst_of_lift_def: Proof.context -> typ -> lift_def -> lift_def  | 
|
23  | 
val mk_lift_const_of_lift_def: typ -> lift_def -> term  | 
|
| 60219 | 24  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
25  | 
  type config = { notes: bool }
 | 
| 
60229
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
60227 
diff
changeset
 | 
26  | 
val map_config: (bool -> bool) -> config -> config  | 
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
27  | 
val default_config: config  | 
| 
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
28  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
29  | 
val generate_parametric_transfer_rule:  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
30  | 
Proof.context -> thm -> thm -> thm  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
31  | 
|
| 63352 | 32  | 
val add_lift_def:  | 
33  | 
config -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory ->  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
34  | 
lift_def * local_theory  | 
| 63352 | 35  | 
|
| 60227 | 36  | 
val prepare_lift_def:  | 
| 63352 | 37  | 
(binding * mixfix -> typ -> term -> thm -> thm list -> Proof.context ->  | 
38  | 
lift_def * local_theory) ->  | 
|
39  | 
binding * mixfix -> typ -> term -> thm list -> local_theory ->  | 
|
| 
60231
 
0daab758e087
don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
 
kuncar 
parents: 
60230 
diff
changeset
 | 
40  | 
term option * (thm -> Proof.context -> lift_def * local_theory)  | 
| 60227 | 41  | 
|
42  | 
val gen_lift_def:  | 
|
| 63352 | 43  | 
(binding * mixfix -> typ -> term -> thm -> thm list -> local_theory ->  | 
44  | 
lift_def * local_theory) ->  | 
|
45  | 
binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list ->  | 
|
| 60227 | 46  | 
local_theory -> lift_def * local_theory  | 
47  | 
||
| 63352 | 48  | 
val lift_def:  | 
49  | 
config -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list ->  | 
|
| 60219 | 50  | 
local_theory -> lift_def * local_theory  | 
| 47308 | 51  | 
|
52  | 
val can_generate_code_cert: thm -> bool  | 
|
| 53651 | 53  | 
end  | 
| 47308 | 54  | 
|
55  | 
structure Lifting_Def: LIFTING_DEF =  | 
|
56  | 
struct  | 
|
57  | 
||
| 47698 | 58  | 
open Lifting_Util  | 
| 47308 | 59  | 
|
60  | 
infix 0 MRSL  | 
|
61  | 
||
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
62  | 
datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ  | 
| 60227 | 63  | 
|
| 60219 | 64  | 
datatype lift_def = LIFT_DEF of {
 | 
65  | 
rty: typ,  | 
|
66  | 
qty: typ,  | 
|
67  | 
rhs: term,  | 
|
68  | 
lift_const: term,  | 
|
69  | 
def_thm: thm,  | 
|
70  | 
rsp_thm: thm,  | 
|
71  | 
abs_eq: thm,  | 
|
72  | 
rep_eq: thm option,  | 
|
| 60227 | 73  | 
code_eq: code_eq,  | 
| 60219 | 74  | 
transfer_rules: thm list  | 
75  | 
};  | 
|
76  | 
||
77  | 
fun rep_lift_def (LIFT_DEF lift_def) = lift_def;  | 
|
78  | 
val rty_of_lift_def = #rty o rep_lift_def;  | 
|
79  | 
val qty_of_lift_def = #qty o rep_lift_def;  | 
|
80  | 
val rhs_of_lift_def = #rhs o rep_lift_def;  | 
|
81  | 
val lift_const_of_lift_def = #lift_const o rep_lift_def;  | 
|
82  | 
val def_thm_of_lift_def = #def_thm o rep_lift_def;  | 
|
83  | 
val rsp_thm_of_lift_def = #rsp_thm o rep_lift_def;  | 
|
84  | 
val abs_eq_of_lift_def = #abs_eq o rep_lift_def;  | 
|
85  | 
val rep_eq_of_lift_def = #rep_eq o rep_lift_def;  | 
|
| 60227 | 86  | 
val code_eq_of_lift_def = #code_eq o rep_lift_def;  | 
| 60219 | 87  | 
val transfer_rules_of_lift_def = #transfer_rules o rep_lift_def;  | 
88  | 
||
| 60227 | 89  | 
fun mk_lift_def rty qty rhs lift_const def_thm rsp_thm abs_eq rep_eq code_eq transfer_rules =  | 
| 60219 | 90  | 
  LIFT_DEF {rty = rty, qty = qty,
 | 
91  | 
rhs = rhs, lift_const = lift_const,  | 
|
| 63352 | 92  | 
def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq,  | 
| 60227 | 93  | 
code_eq = code_eq, transfer_rules = transfer_rules };  | 
| 60219 | 94  | 
|
| 60227 | 95  | 
fun map_lift_def f1 f2 f3 f4 f5 f6 f7 f8 f9 f10  | 
| 60224 | 96  | 
  (LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const,
 | 
| 60227 | 97  | 
def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq,  | 
| 60224 | 98  | 
transfer_rules = transfer_rules }) =  | 
99  | 
  LIFT_DEF {rty = f1 rty, qty = f2 qty, rhs = f3 rhs, lift_const = f4 lift_const,
 | 
|
100  | 
def_thm = f5 def_thm, rsp_thm = f6 rsp_thm, abs_eq = f7 abs_eq, rep_eq = f8 rep_eq,  | 
|
| 60227 | 101  | 
code_eq = f9 code_eq, transfer_rules = f10 transfer_rules }  | 
| 60224 | 102  | 
|
103  | 
fun morph_lift_def phi =  | 
|
104  | 
let  | 
|
105  | 
val mtyp = Morphism.typ phi  | 
|
106  | 
val mterm = Morphism.term phi  | 
|
107  | 
val mthm = Morphism.thm phi  | 
|
108  | 
in  | 
|
| 60227 | 109  | 
map_lift_def mtyp mtyp mterm mterm mthm mthm mthm (Option.map mthm) I (map mthm)  | 
| 60224 | 110  | 
end  | 
111  | 
||
112  | 
fun mk_inst_of_lift_def qty lift_def = Vartab.empty |> Type.raw_match (qty_of_lift_def lift_def, qty)  | 
|
113  | 
||
114  | 
fun mk_lift_const_of_lift_def qty lift_def = Envir.subst_term_types (mk_inst_of_lift_def qty lift_def)  | 
|
115  | 
(lift_const_of_lift_def lift_def)  | 
|
116  | 
||
| 
67709
 
3c9e0b4709e7
tuned -- use existing Morphism.instantiate_morphism;
 
wenzelm 
parents: 
67703 
diff
changeset
 | 
117  | 
fun inst_of_lift_def ctxt qty lift_def =  | 
| 
 
3c9e0b4709e7
tuned -- use existing Morphism.instantiate_morphism;
 
wenzelm 
parents: 
67703 
diff
changeset
 | 
118  | 
let  | 
| 
 
3c9e0b4709e7
tuned -- use existing Morphism.instantiate_morphism;
 
wenzelm 
parents: 
67703 
diff
changeset
 | 
119  | 
val instT =  | 
| 
 
3c9e0b4709e7
tuned -- use existing Morphism.instantiate_morphism;
 
wenzelm 
parents: 
67703 
diff
changeset
 | 
120  | 
Vartab.fold (fn (a, (S, T)) => cons ((a, S), Thm.ctyp_of ctxt T))  | 
| 
 
3c9e0b4709e7
tuned -- use existing Morphism.instantiate_morphism;
 
wenzelm 
parents: 
67703 
diff
changeset
 | 
121  | 
(mk_inst_of_lift_def qty lift_def) []  | 
| 
 
3c9e0b4709e7
tuned -- use existing Morphism.instantiate_morphism;
 
wenzelm 
parents: 
67703 
diff
changeset
 | 
122  | 
val phi = Morphism.instantiate_morphism (instT, [])  | 
| 
 
3c9e0b4709e7
tuned -- use existing Morphism.instantiate_morphism;
 
wenzelm 
parents: 
67703 
diff
changeset
 | 
123  | 
in morph_lift_def phi lift_def end  | 
| 
 
3c9e0b4709e7
tuned -- use existing Morphism.instantiate_morphism;
 
wenzelm 
parents: 
67703 
diff
changeset
 | 
124  | 
|
| 60224 | 125  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
126  | 
(* Config *)  | 
| 
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
127  | 
|
| 
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
128  | 
type config = { notes: bool };
 | 
| 
60229
 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 
kuncar 
parents: 
60227 
diff
changeset
 | 
129  | 
fun map_config f1 { notes = notes } = { notes = f1 notes }
 | 
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
130  | 
val default_config = { notes = true };
 | 
| 
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
131  | 
|
| 51994 | 132  | 
(* Reflexivity prover *)  | 
133  | 
||
| 55609 | 134  | 
fun mono_eq_prover ctxt prop =  | 
| 51994 | 135  | 
let  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56245 
diff
changeset
 | 
136  | 
val refl_rules = Lifting_Info.get_reflexivity_rules ctxt  | 
| 
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56245 
diff
changeset
 | 
137  | 
val transfer_rules = Transfer.get_transfer_raw ctxt  | 
| 63352 | 138  | 
|
139  | 
fun main_tac i = (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt refl_rules) THEN_ALL_NEW  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59458 
diff
changeset
 | 
140  | 
(REPEAT_ALL_NEW (DETERM o resolve_tac ctxt transfer_rules))) i  | 
| 51994 | 141  | 
in  | 
| 56540 | 142  | 
SOME (Goal.prove ctxt [] [] prop (K (main_tac 1)))  | 
| 55609 | 143  | 
handle ERROR _ => NONE  | 
| 51994 | 144  | 
end  | 
| 
56518
 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 
kuncar 
parents: 
56245 
diff
changeset
 | 
145  | 
|
| 
57642
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
146  | 
fun try_prove_refl_rel ctxt rel =  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
147  | 
let  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
148  | 
fun mk_ge_eq x =  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
149  | 
let  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
150  | 
val T = fastype_of x  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
151  | 
in  | 
| 63352 | 152  | 
        Const (@{const_name "less_eq"}, T --> T --> HOLogic.boolT) $
 | 
| 
57642
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
153  | 
          (Const (@{const_name HOL.eq}, T)) $ x
 | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
154  | 
end;  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
155  | 
val goal = HOLogic.mk_Trueprop (mk_ge_eq rel);  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
156  | 
in  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
157  | 
mono_eq_prover ctxt goal  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
158  | 
end;  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
159  | 
|
| 51994 | 160  | 
fun try_prove_reflexivity ctxt prop =  | 
| 55609 | 161  | 
let  | 
| 59630 | 162  | 
val cprop = Thm.cterm_of ctxt prop  | 
| 55609 | 163  | 
    val rule = @{thm ge_eq_refl}
 | 
| 59582 | 164  | 
val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule)  | 
| 55609 | 165  | 
val insts = Thm.first_order_match (concl_pat, cprop)  | 
166  | 
val rule = Drule.instantiate_normalize insts rule  | 
|
| 59582 | 167  | 
val prop = hd (Thm.prems_of rule)  | 
| 55609 | 168  | 
in  | 
169  | 
case mono_eq_prover ctxt prop of  | 
|
| 
55610
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
170  | 
SOME thm => SOME (thm RS rule)  | 
| 55609 | 171  | 
| NONE => NONE  | 
172  | 
end  | 
|
| 51994 | 173  | 
|
| 63352 | 174  | 
(*  | 
| 
54335
 
03b10317ba78
update documentation of important public ML functions in Lifting
 
kuncar 
parents: 
53951 
diff
changeset
 | 
175  | 
Generates a parametrized transfer rule.  | 
| 
 
03b10317ba78
update documentation of important public ML functions in Lifting
 
kuncar 
parents: 
53951 
diff
changeset
 | 
176  | 
transfer_rule - of the form T t f  | 
| 
 
03b10317ba78
update documentation of important public ML functions in Lifting
 
kuncar 
parents: 
53951 
diff
changeset
 | 
177  | 
parametric_transfer_rule - of the form par_R t' t  | 
| 63352 | 178  | 
|
| 54947 | 179  | 
Result: par_T t' f, after substituing op= for relations in par_R that relate  | 
| 
54335
 
03b10317ba78
update documentation of important public ML functions in Lifting
 
kuncar 
parents: 
53951 
diff
changeset
 | 
180  | 
a type constructor to the same type constructor, it is a merge of (par_R' OO T) t' f  | 
| 
 
03b10317ba78
update documentation of important public ML functions in Lifting
 
kuncar 
parents: 
53951 
diff
changeset
 | 
181  | 
using Lifting_Term.merge_transfer_relations  | 
| 
 
03b10317ba78
update documentation of important public ML functions in Lifting
 
kuncar 
parents: 
53951 
diff
changeset
 | 
182  | 
*)  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
183  | 
|
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
184  | 
fun generate_parametric_transfer_rule ctxt transfer_rule parametric_transfer_rule =  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
185  | 
let  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
186  | 
fun preprocess ctxt thm =  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
187  | 
let  | 
| 59582 | 188  | 
val tm = (strip_args 2 o HOLogic.dest_Trueprop o Thm.concl_of) thm;  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
189  | 
val param_rel = (snd o dest_comb o fst o dest_comb) tm;  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
190  | 
val free_vars = Term.add_vars param_rel [];  | 
| 63352 | 191  | 
|
192  | 
fun make_subst (xi, typ) subst =  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
193  | 
let  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
194  | 
val [rty, rty'] = binder_types typ  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
195  | 
in  | 
| 67703 | 196  | 
if Term.is_TVar rty andalso Term.is_Type rty' then  | 
| 60784 | 197  | 
(xi, Thm.cterm_of ctxt (HOLogic.eq_const rty')) :: subst  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
198  | 
else  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
199  | 
subst  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
200  | 
end;  | 
| 60784 | 201  | 
|
202  | 
val inst_thm = infer_instantiate ctxt (fold make_subst free_vars []) thm;  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
203  | 
in  | 
| 63352 | 204  | 
Conv.fconv_rule  | 
| 59582 | 205  | 
((Conv.concl_conv (Thm.nprems_of inst_thm) o  | 
206  | 
HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv)  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54335 
diff
changeset
 | 
207  | 
(Raw_Simplifier.rewrite ctxt false (Transfer.get_sym_relator_eq ctxt))) inst_thm  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
208  | 
end  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
209  | 
|
| 59630 | 210  | 
fun inst_relcomppI ctxt ant1 ant2 =  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
211  | 
let  | 
| 59582 | 212  | 
val t1 = (HOLogic.dest_Trueprop o Thm.concl_of) ant1  | 
213  | 
val t2 = (HOLogic.dest_Trueprop o Thm.prop_of) ant2  | 
|
| 59630 | 214  | 
val fun1 = Thm.cterm_of ctxt (strip_args 2 t1)  | 
215  | 
val args1 = map (Thm.cterm_of ctxt) (get_args 2 t1)  | 
|
216  | 
val fun2 = Thm.cterm_of ctxt (strip_args 2 t2)  | 
|
217  | 
val args2 = map (Thm.cterm_of ctxt) (get_args 1 t2)  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
218  | 
        val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI}
 | 
| 60784 | 219  | 
val vars = map #1 (rev (Term.add_vars (Thm.prop_of relcomppI) []))  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
220  | 
in  | 
| 60784 | 221  | 
infer_instantiate ctxt (vars ~~ ([fun1] @ args1 @ [fun2] @ args2)) relcomppI  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
222  | 
end  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
223  | 
|
| 54995 | 224  | 
fun zip_transfer_rules ctxt thm =  | 
225  | 
let  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
226  | 
        fun mk_POS ty = Const (@{const_name POS}, ty --> ty --> HOLogic.boolT)
 | 
| 59582 | 227  | 
val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm  | 
| 59586 | 228  | 
val typ = Thm.typ_of_cterm rel  | 
| 59630 | 229  | 
val POS_const = Thm.cterm_of ctxt (mk_POS typ)  | 
230  | 
        val var = Thm.cterm_of ctxt (Var (("X", Thm.maxidx_of_cterm rel + 1), typ))
 | 
|
| 59582 | 231  | 
val goal =  | 
| 59630 | 232  | 
Thm.apply (Thm.cterm_of ctxt HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var)  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
233  | 
in  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
234  | 
        [Lifting_Term.merge_transfer_relations ctxt goal, thm] MRSL @{thm POS_apply}
 | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
235  | 
end  | 
| 63352 | 236  | 
|
| 59630 | 237  | 
val thm =  | 
238  | 
inst_relcomppI ctxt parametric_transfer_rule transfer_rule  | 
|
239  | 
OF [parametric_transfer_rule, transfer_rule]  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
240  | 
val preprocessed_thm = preprocess ctxt thm  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
241  | 
val orig_ctxt = ctxt  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
242  | 
val (fixed_thm, ctxt) = yield_singleton (apfst snd oo Variable.import true) preprocessed_thm ctxt  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
243  | 
val assms = cprems_of fixed_thm  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
244  | 
val add_transfer_rule = Thm.attribute_declaration Transfer.transfer_add  | 
| 54995 | 245  | 
val (prems, ctxt) = fold_map Thm.assume_hyps assms ctxt  | 
246  | 
val ctxt = Context.proof_map (fold add_transfer_rule prems) ctxt  | 
|
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
247  | 
val zipped_thm =  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
248  | 
fixed_thm  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
249  | 
|> undisch_all  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
250  | 
|> zip_transfer_rules ctxt  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
251  | 
|> implies_intr_list assms  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
252  | 
|> singleton (Variable.export ctxt orig_ctxt)  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
253  | 
in  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
254  | 
zipped_thm  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
255  | 
end  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
256  | 
|
| 63352 | 257  | 
fun print_generate_transfer_info msg =  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
258  | 
let  | 
| 63352 | 259  | 
val error_msg = cat_lines  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
260  | 
["Generation of a parametric transfer rule failed.",  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
261  | 
(Pretty.string_of (Pretty.block  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
262  | 
[Pretty.str "Reason:", Pretty.brk 2, msg]))]  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
263  | 
in  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
264  | 
error error_msg  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
265  | 
end  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
266  | 
|
| 
53951
 
03b74ef6d7c6
allow to specify multiple parametricity transfer rules in lift_definition
 
kuncar 
parents: 
53651 
diff
changeset
 | 
267  | 
fun map_ter _ x [] = x  | 
| 
 
03b74ef6d7c6
allow to specify multiple parametricity transfer rules in lift_definition
 
kuncar 
parents: 
53651 
diff
changeset
 | 
268  | 
| map_ter f _ xs = map f xs  | 
| 
 
03b74ef6d7c6
allow to specify multiple parametricity transfer rules in lift_definition
 
kuncar 
parents: 
53651 
diff
changeset
 | 
269  | 
|
| 
 
03b74ef6d7c6
allow to specify multiple parametricity transfer rules in lift_definition
 
kuncar 
parents: 
53651 
diff
changeset
 | 
270  | 
fun generate_transfer_rules lthy quot_thm rsp_thm def_thm par_thms =  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
271  | 
let  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
272  | 
val transfer_rule =  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
273  | 
      ([quot_thm, rsp_thm, def_thm] MRSL @{thm Quotient_to_transfer})
 | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
274  | 
|> Lifting_Term.parametrize_transfer_rule lthy  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
275  | 
in  | 
| 
53951
 
03b74ef6d7c6
allow to specify multiple parametricity transfer rules in lift_definition
 
kuncar 
parents: 
53651 
diff
changeset
 | 
276  | 
(map_ter (generate_parametric_transfer_rule lthy transfer_rule) [transfer_rule] par_thms  | 
| 
 
03b74ef6d7c6
allow to specify multiple parametricity transfer rules in lift_definition
 
kuncar 
parents: 
53651 
diff
changeset
 | 
277  | 
handle Lifting_Term.MERGE_TRANSFER_REL msg => (print_generate_transfer_info msg; [transfer_rule]))  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
278  | 
end  | 
| 
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
279  | 
|
| 47698 | 280  | 
(* Generation of the code certificate from the rsp theorem *)  | 
| 47308 | 281  | 
|
282  | 
fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V)
 | 
|
283  | 
| get_body_types (U, V) = (U, V)  | 
|
284  | 
||
285  | 
fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W)
 | 
|
286  | 
| get_binder_types _ = []  | 
|
287  | 
||
| 63352 | 288  | 
fun get_binder_types_by_rel (Const (@{const_name "rel_fun"}, _) $ _ $ S) (Type ("fun", [T, U]), Type ("fun", [V, W])) =
 | 
| 
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: 
47852 
diff
changeset
 | 
289  | 
(T, V) :: get_binder_types_by_rel S (U, W)  | 
| 
 
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: 
47852 
diff
changeset
 | 
290  | 
| get_binder_types_by_rel _ _ = []  | 
| 
 
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: 
47852 
diff
changeset
 | 
291  | 
|
| 63352 | 292  | 
fun get_body_type_by_rel (Const (@{const_name "rel_fun"}, _) $ _ $ S) (Type ("fun", [_, U]), Type ("fun", [_, V])) =
 | 
| 
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: 
47852 
diff
changeset
 | 
293  | 
get_body_type_by_rel S (U, V)  | 
| 
 
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: 
47852 
diff
changeset
 | 
294  | 
| get_body_type_by_rel _ (U, V) = (U, V)  | 
| 
 
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: 
47852 
diff
changeset
 | 
295  | 
|
| 
57642
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
296  | 
fun get_binder_rels (Const (@{const_name "rel_fun"}, _) $ R $ S) = R :: get_binder_rels S
 | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
297  | 
| get_binder_rels _ = []  | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
298  | 
|
| 63352 | 299  | 
fun force_rty_type ctxt rty rhs =  | 
| 47308 | 300  | 
let  | 
301  | 
val thy = Proof_Context.theory_of ctxt  | 
|
302  | 
val rhs_schematic = singleton (Variable.polymorphic ctxt) rhs  | 
|
303  | 
val rty_schematic = fastype_of rhs_schematic  | 
|
304  | 
val match = Sign.typ_match thy (rty_schematic, rty) Vartab.empty  | 
|
305  | 
in  | 
|
306  | 
Envir.subst_term_types match rhs_schematic  | 
|
307  | 
end  | 
|
308  | 
||
| 63352 | 309  | 
fun unabs_def ctxt def =  | 
| 47308 | 310  | 
let  | 
| 59582 | 311  | 
val (_, rhs) = Thm.dest_equals (Thm.cprop_of def)  | 
| 47308 | 312  | 
fun dest_abs (Abs (var_name, T, _)) = (var_name, T)  | 
313  | 
      | dest_abs tm = raise TERM("get_abs_var",[tm])
 | 
|
| 59582 | 314  | 
val (var_name, T) = dest_abs (Thm.term_of rhs)  | 
| 47308 | 315  | 
val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt  | 
| 59630 | 316  | 
val refl_thm = Thm.reflexive (Thm.cterm_of ctxt' (Free (hd new_var_names, T)))  | 
| 47308 | 317  | 
in  | 
318  | 
Thm.combination def refl_thm |>  | 
|
319  | 
singleton (Proof_Context.export ctxt' ctxt)  | 
|
320  | 
end  | 
|
321  | 
||
| 63352 | 322  | 
fun unabs_all_def ctxt def =  | 
| 47308 | 323  | 
let  | 
| 59582 | 324  | 
val (_, rhs) = Thm.dest_equals (Thm.cprop_of def)  | 
325  | 
val xs = strip_abs_vars (Thm.term_of rhs)  | 
|
| 63352 | 326  | 
in  | 
| 47308 | 327  | 
fold (K (unabs_def ctxt)) xs def  | 
328  | 
end  | 
|
329  | 
||
| 63352 | 330  | 
val map_fun_unfolded =  | 
| 47308 | 331  | 
  @{thm map_fun_def[abs_def]} |>
 | 
332  | 
  unabs_def @{context} |>
 | 
|
333  | 
  unabs_def @{context} |>
 | 
|
| 63170 | 334  | 
  Local_Defs.unfold0 @{context} [@{thm comp_def}]
 | 
| 47308 | 335  | 
|
336  | 
fun unfold_fun_maps ctm =  | 
|
337  | 
let  | 
|
338  | 
fun unfold_conv ctm =  | 
|
339  | 
case (Thm.term_of ctm) of  | 
|
| 63352 | 340  | 
        Const (@{const_name "map_fun"}, _) $ _ $ _ =>
 | 
| 47308 | 341  | 
(Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm  | 
342  | 
| _ => Conv.all_conv ctm  | 
|
343  | 
in  | 
|
| 
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: 
47852 
diff
changeset
 | 
344  | 
(Conv.fun_conv unfold_conv) ctm  | 
| 
 
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: 
47852 
diff
changeset
 | 
345  | 
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: 
47852 
diff
changeset
 | 
346  | 
|
| 
 
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: 
47852 
diff
changeset
 | 
347  | 
fun unfold_fun_maps_beta ctm =  | 
| 
 
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: 
47852 
diff
changeset
 | 
348  | 
let val try_beta_conv = Conv.try_conv (Thm.beta_conversion false)  | 
| 63352 | 349  | 
in  | 
350  | 
(unfold_fun_maps then_conv try_beta_conv) ctm  | 
|
| 47308 | 351  | 
end  | 
352  | 
||
353  | 
fun prove_rel ctxt rsp_thm (rty, qty) =  | 
|
354  | 
let  | 
|
355  | 
val ty_args = get_binder_types (rty, qty)  | 
|
| 63352 | 356  | 
fun disch_arg args_ty thm =  | 
| 47308 | 357  | 
let  | 
| 
47504
 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 
kuncar 
parents: 
47503 
diff
changeset
 | 
358  | 
val quot_thm = Lifting_Term.prove_quot_thm ctxt args_ty  | 
| 47308 | 359  | 
in  | 
360  | 
        [quot_thm, thm] MRSL @{thm apply_rsp''}
 | 
|
361  | 
end  | 
|
362  | 
in  | 
|
363  | 
fold disch_arg ty_args rsp_thm  | 
|
364  | 
end  | 
|
365  | 
||
366  | 
exception CODE_CERT_GEN of string  | 
|
367  | 
||
| 63352 | 368  | 
fun simplify_code_eq ctxt def_thm =  | 
| 63170 | 369  | 
  Local_Defs.unfold0 ctxt [@{thm o_apply}, @{thm map_fun_def}, @{thm id_apply}] def_thm
 | 
| 47308 | 370  | 
|
| 
47852
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
371  | 
(*  | 
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
372  | 
quot_thm - quotient theorem (Quotient R Abs Rep T).  | 
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
373  | 
returns: whether the Lifting package is capable to generate code for the abstract type  | 
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
374  | 
represented by quot_thm  | 
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
375  | 
*)  | 
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
376  | 
|
| 47308 | 377  | 
fun can_generate_code_cert quot_thm =  | 
| 
47951
 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 
kuncar 
parents: 
47937 
diff
changeset
 | 
378  | 
case quot_thm_rel quot_thm of  | 
| 47308 | 379  | 
    Const (@{const_name HOL.eq}, _) => true
 | 
| 
56519
 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 
kuncar 
parents: 
56518 
diff
changeset
 | 
380  | 
    | Const (@{const_name eq_onp}, _) $ _  => true
 | 
| 47308 | 381  | 
| _ => false  | 
382  | 
||
| 
55610
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
383  | 
fun generate_rep_eq ctxt def_thm rsp_thm (rty, qty) =  | 
| 47308 | 384  | 
let  | 
| 
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: 
47852 
diff
changeset
 | 
385  | 
val unfolded_def = Conv.fconv_rule (Conv.arg_conv unfold_fun_maps_beta) def_thm  | 
| 47308 | 386  | 
val unabs_def = unabs_all_def ctxt unfolded_def  | 
| 63352 | 387  | 
in  | 
388  | 
if body_type rty = body_type qty then  | 
|
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67709 
diff
changeset
 | 
389  | 
SOME (simplify_code_eq ctxt (HOLogic.mk_obj_eq unabs_def))  | 
| 63352 | 390  | 
else  | 
| 
55610
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
391  | 
let  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
392  | 
val quot_thm = Lifting_Term.prove_quot_thm ctxt (get_body_types (rty, qty))  | 
| 55945 | 393  | 
val rel_fun = prove_rel ctxt rsp_thm (rty, qty)  | 
394  | 
        val rep_abs_thm = [quot_thm, rel_fun] MRSL @{thm Quotient_rep_abs_eq}
 | 
|
| 
55610
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
395  | 
in  | 
| 59582 | 396  | 
case mono_eq_prover ctxt (hd (Thm.prems_of rep_abs_thm)) of  | 
| 
55610
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
397  | 
SOME mono_eq_thm =>  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
398  | 
let  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
399  | 
val rep_abs_eq = mono_eq_thm RS rep_abs_thm  | 
| 59630 | 400  | 
val rep = Thm.cterm_of ctxt (quot_thm_rep quot_thm)  | 
| 
67710
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67709 
diff
changeset
 | 
401  | 
val rep_refl = HOLogic.mk_obj_eq (Thm.reflexive rep)  | 
| 
 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 
wenzelm 
parents: 
67709 
diff
changeset
 | 
402  | 
              val repped_eq = [rep_refl, HOLogic.mk_obj_eq unabs_def] MRSL @{thm cong}
 | 
| 
55610
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
403  | 
val code_cert = [repped_eq, rep_abs_eq] MRSL trans  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
404  | 
in  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
405  | 
SOME (simplify_code_eq ctxt code_cert)  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
406  | 
end  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
407  | 
| NONE => NONE  | 
| 
 
9066b603dff6
refactoring; generate rep_eq always, not only when it would be accepted by the code generator
 
kuncar 
parents: 
55609 
diff
changeset
 | 
408  | 
end  | 
| 47308 | 409  | 
end  | 
410  | 
||
| 
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: 
47852 
diff
changeset
 | 
411  | 
fun generate_abs_eq ctxt def_thm rsp_thm quot_thm =  | 
| 47308 | 412  | 
let  | 
| 
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: 
47852 
diff
changeset
 | 
413  | 
val abs_eq_with_assms =  | 
| 
 
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: 
47852 
diff
changeset
 | 
414  | 
let  | 
| 
47951
 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 
kuncar 
parents: 
47937 
diff
changeset
 | 
415  | 
val (rty, qty) = quot_thm_rty_qty quot_thm  | 
| 
 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 
kuncar 
parents: 
47937 
diff
changeset
 | 
416  | 
val rel = quot_thm_rel 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: 
47852 
diff
changeset
 | 
417  | 
val ty_args = get_binder_types_by_rel rel (rty, qty)  | 
| 
 
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: 
47852 
diff
changeset
 | 
418  | 
val body_type = get_body_type_by_rel rel (rty, qty)  | 
| 
 
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: 
47852 
diff
changeset
 | 
419  | 
val quot_ret_thm = Lifting_Term.prove_quot_thm ctxt body_type  | 
| 63352 | 420  | 
|
421  | 
val rep_abs_folded_unmapped_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: 
47852 
diff
changeset
 | 
422  | 
let  | 
| 
 
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: 
47852 
diff
changeset
 | 
423  | 
            val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq}
 | 
| 59582 | 424  | 
val ctm = Thm.dest_equals_lhs (Thm.cprop_of rep_id)  | 
| 
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: 
47852 
diff
changeset
 | 
425  | 
val unfolded_maps_eq = unfold_fun_maps ctm  | 
| 
 
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: 
47852 
diff
changeset
 | 
426  | 
            val t1 = [quot_thm, def_thm, rsp_thm] MRSL @{thm Quotient_rep_abs_fold_unmap}
 | 
| 
 
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: 
47852 
diff
changeset
 | 
427  | 
val prems_pat = (hd o Drule.cprems_of) t1  | 
| 59582 | 428  | 
val insts = Thm.first_order_match (prems_pat, Thm.cprop_of unfolded_maps_eq)  | 
| 
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: 
47852 
diff
changeset
 | 
429  | 
in  | 
| 
 
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: 
47852 
diff
changeset
 | 
430  | 
unfolded_maps_eq RS (Drule.instantiate_normalize insts t1)  | 
| 
 
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: 
47852 
diff
changeset
 | 
431  | 
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: 
47852 
diff
changeset
 | 
432  | 
in  | 
| 
 
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: 
47852 
diff
changeset
 | 
433  | 
rep_abs_folded_unmapped_thm  | 
| 55945 | 434  | 
        |> fold (fn _ => fn thm => thm RS @{thm rel_funD2}) ty_args
 | 
| 
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: 
47852 
diff
changeset
 | 
435  | 
        |> (fn x => x RS (@{thm Quotient_rel_abs2} OF [quot_ret_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: 
47852 
diff
changeset
 | 
436  | 
end  | 
| 63352 | 437  | 
|
| 
57642
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
438  | 
val prem_rels = get_binder_rels (quot_thm_rel quot_thm);  | 
| 63352 | 439  | 
val proved_assms = prem_rels |> map (try_prove_refl_rel ctxt)  | 
440  | 
|> map_index (apfst (fn x => x + 1)) |> filter (is_some o snd) |> map (apsnd the)  | 
|
| 
57642
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
441  | 
      |> map (apsnd (fn assm => assm RS @{thm ge_eq_refl}))
 | 
| 
 
5bc43a73d768
prevent beta-contraction in proving extra assumptions for abs_eq
 
kuncar 
parents: 
56540 
diff
changeset
 | 
442  | 
val abs_eq = fold_rev (fn (i, assm) => fn thm => assm RSN (i, thm)) proved_assms abs_eq_with_assms  | 
| 47308 | 443  | 
in  | 
| 
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: 
47852 
diff
changeset
 | 
444  | 
simplify_code_eq ctxt abs_eq  | 
| 
 
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: 
47852 
diff
changeset
 | 
445  | 
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: 
47852 
diff
changeset
 | 
446  | 
|
| 
 
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: 
47852 
diff
changeset
 | 
447  | 
|
| 55724 | 448  | 
fun register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy =  | 
| 
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: 
47852 
diff
changeset
 | 
449  | 
let  | 
| 55724 | 450  | 
fun no_abstr (t $ u) = no_abstr t andalso no_abstr u  | 
451  | 
| no_abstr (Abs (_, _, t)) = no_abstr t  | 
|
452  | 
| no_abstr (Const (name, _)) = not (Code.is_abstr thy name)  | 
|
453  | 
| no_abstr _ = true  | 
|
| 63352 | 454  | 
fun is_valid_eq eqn = can (Code.assert_eqn thy) (mk_meta_eq eqn, true)  | 
| 59582 | 455  | 
andalso no_abstr (Thm.prop_of eqn)  | 
| 55724 | 456  | 
fun is_valid_abs_eq abs_eq = can (Code.assert_abs_eqn thy NONE) (mk_meta_eq abs_eq)  | 
| 47308 | 457  | 
|
458  | 
in  | 
|
| 55724 | 459  | 
if is_valid_eq abs_eq_thm then  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66017 
diff
changeset
 | 
460  | 
(ABS_EQ, Code.declare_default_eqns_global [(abs_eq_thm, true)] thy)  | 
| 
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: 
47852 
diff
changeset
 | 
461  | 
else  | 
| 55724 | 462  | 
let  | 
463  | 
val (rty_body, qty_body) = get_body_types (rty, qty)  | 
|
| 
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: 
47852 
diff
changeset
 | 
464  | 
in  | 
| 55724 | 465  | 
if rty_body = qty_body then  | 
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66017 
diff
changeset
 | 
466  | 
(REP_EQ, Code.declare_default_eqns_global [(the opt_rep_eq_thm, true)] thy)  | 
| 
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: 
47852 
diff
changeset
 | 
467  | 
else  | 
| 55724 | 468  | 
if is_some opt_rep_eq_thm andalso is_valid_abs_eq (the opt_rep_eq_thm)  | 
469  | 
then  | 
|
| 
66251
 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 
haftmann 
parents: 
66017 
diff
changeset
 | 
470  | 
(REP_EQ, Code.declare_abstract_eqn_global (the opt_rep_eq_thm) thy)  | 
| 55724 | 471  | 
else  | 
| 60227 | 472  | 
(NONE_EQ, thy)  | 
| 
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: 
47852 
diff
changeset
 | 
473  | 
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: 
47852 
diff
changeset
 | 
474  | 
end  | 
| 47308 | 475  | 
|
| 55724 | 476  | 
local  | 
| 
57663
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
477  | 
fun no_no_code ctxt (rty, qty) =  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
478  | 
if same_type_constrs (rty, qty) then  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
479  | 
forall (no_no_code ctxt) (Targs rty ~~ Targs qty)  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
480  | 
else  | 
| 67703 | 481  | 
if Term.is_Type qty then  | 
| 
57663
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
482  | 
if Lifting_Info.is_no_code_type ctxt (Tname qty) then false  | 
| 63352 | 483  | 
else  | 
484  | 
let  | 
|
| 
57663
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
485  | 
val (rty', rtyq) = Lifting_Term.instantiate_rtys ctxt (rty, qty)  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
486  | 
val (rty's, rtyqs) = (Targs rty', Targs rtyq)  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
487  | 
in  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
488  | 
forall (no_no_code ctxt) (rty's ~~ rtyqs)  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
489  | 
end  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
490  | 
else  | 
| 
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
491  | 
true  | 
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
492  | 
|
| 63352 | 493  | 
fun encode_code_eq ctxt abs_eq opt_rep_eq (rty, qty) =  | 
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
494  | 
let  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
495  | 
fun mk_type typ = typ |> Logic.mk_type |> Thm.cterm_of ctxt |> Drule.mk_term  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
496  | 
in  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
497  | 
Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty]  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
498  | 
end  | 
| 63352 | 499  | 
|
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
500  | 
exception DECODE  | 
| 63352 | 501  | 
|
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
502  | 
fun decode_code_eq thm =  | 
| 63352 | 503  | 
if Thm.nprems_of thm > 0 then raise DECODE  | 
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
504  | 
else  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
505  | 
let  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
506  | 
val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
507  | 
val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
508  | 
fun dest_type typ = typ |> Drule.dest_term |> Thm.term_of |> Logic.dest_type  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
509  | 
in  | 
| 63352 | 510  | 
(abs_eq, opt_rep_eq, (dest_type rty, dest_type qty))  | 
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
511  | 
end  | 
| 63352 | 512  | 
|
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
513  | 
structure Data = Generic_Data  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
514  | 
(  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
515  | 
type T = code_eq option  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
516  | 
val empty = NONE  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
517  | 
val extend = I  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
518  | 
fun merge _ = NONE  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
519  | 
);  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
520  | 
|
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
521  | 
fun register_encoded_code_eq thm thy =  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
522  | 
let  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
523  | 
val (abs_eq_thm, opt_rep_eq_thm, (rty, qty)) = decode_code_eq thm  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
524  | 
val (code_eq, thy) = register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
525  | 
in  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
526  | 
Context.theory_map (Data.put (SOME code_eq)) thy  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
527  | 
end  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
528  | 
handle DECODE => thy  | 
| 63352 | 529  | 
|
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
530  | 
val register_code_eq_attribute = Thm.declaration_attribute  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
531  | 
(fn thm => Context.mapping (register_encoded_code_eq thm) I)  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
532  | 
val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute)  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
533  | 
|
| 55724 | 534  | 
in  | 
535  | 
||
536  | 
fun register_code_eq abs_eq_thm opt_rep_eq_thm (rty, qty) lthy =  | 
|
537  | 
let  | 
|
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
538  | 
val encoded_code_eq = encode_code_eq lthy abs_eq_thm opt_rep_eq_thm (rty, qty)  | 
| 55724 | 539  | 
in  | 
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
540  | 
if no_no_code lthy (rty, qty) then  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
541  | 
let  | 
| 63352 | 542  | 
val lthy = (snd oo Local_Theory.note)  | 
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
543  | 
((Binding.empty, [register_code_eq_attrib]), [encoded_code_eq]) lthy  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
544  | 
val opt_code_eq = Data.get (Context.Theory (Proof_Context.theory_of lthy))  | 
| 63352 | 545  | 
val code_eq =  | 
546  | 
if is_some opt_code_eq then the opt_code_eq  | 
|
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
547  | 
else UNKNOWN_EQ (* UNKNOWN_EQ means that we are in a locale and we do not know  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
548  | 
which code equation is going to be used. This is going to be resolved at the  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
549  | 
point when an interpretation of the locale is executed. *)  | 
| 63352 | 550  | 
        val lthy = Local_Theory.declaration {syntax = false, pervasive = true}
 | 
| 
60230
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
551  | 
(K (Data.put NONE)) lthy  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
552  | 
in  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
553  | 
(code_eq, lthy)  | 
| 
 
4857d553c52c
go back to the complicated code equation registration (because of type classes) that was lost in 922586b1bc87; make it even more hackish to get which code equation was used
 
kuncar 
parents: 
60229 
diff
changeset
 | 
554  | 
end  | 
| 
57663
 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 
kuncar 
parents: 
57642 
diff
changeset
 | 
555  | 
else  | 
| 60227 | 556  | 
(NONE_EQ, lthy)  | 
| 55724 | 557  | 
end  | 
558  | 
end  | 
|
| 63352 | 559  | 
|
| 
47852
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
560  | 
(*  | 
| 63352 | 561  | 
Defines an operation on an abstract type in terms of a corresponding operation  | 
| 
47852
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
562  | 
on a representation type.  | 
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
563  | 
|
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
564  | 
var - a binding and a mixfix of the new constant being defined  | 
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
565  | 
qty - an abstract type of the new constant  | 
| 
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
566  | 
rhs - a term representing the new constant on the raw level  | 
| 
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: 
47852 
diff
changeset
 | 
567  | 
rsp_thm - a respectfulness theorem in the internal tagged form (like '(R ===> R ===> R) f f'),  | 
| 
47852
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
568  | 
i.e. "(Lifting_Term.equiv_relation (fastype_of rhs, qty)) $ rhs $ rhs"  | 
| 
54335
 
03b10317ba78
update documentation of important public ML functions in Lifting
 
kuncar 
parents: 
53951 
diff
changeset
 | 
569  | 
par_thms - a parametricity theorem for rhs  | 
| 
47852
 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 
kuncar 
parents: 
47699 
diff
changeset
 | 
570  | 
*)  | 
| 47308 | 571  | 
|
| 63341 | 572  | 
fun add_lift_def (config: config) (binding, mx) qty rhs rsp_thm par_thms lthy =  | 
| 47308 | 573  | 
let  | 
574  | 
val rty = fastype_of rhs  | 
|
| 
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: 
47852 
diff
changeset
 | 
575  | 
val quot_thm = Lifting_Term.prove_quot_thm lthy (rty, qty)  | 
| 
47951
 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 
kuncar 
parents: 
47937 
diff
changeset
 | 
576  | 
val absrep_trm = quot_thm_abs quot_thm  | 
| 47308 | 577  | 
val rty_forced = (domain_type o fastype_of) absrep_trm  | 
578  | 
val forced_rhs = force_rty_type lthy rty_forced rhs  | 
|
| 63341 | 579  | 
val lhs = Free (Binding.name_of binding, qty)  | 
| 47308 | 580  | 
val prop = Logic.mk_equals (lhs, absrep_trm $ forced_rhs)  | 
| 63395 | 581  | 
val (_, prop') = Local_Defs.cert_def lthy (K []) prop  | 
| 47308 | 582  | 
val (_, newrhs) = Local_Defs.abs_def prop'  | 
| 63341 | 583  | 
val var = ((#notes config = false ? Binding.concealed) binding, mx)  | 
584  | 
val def_name = Thm.make_def_binding (#notes config) (#1 var)  | 
|
| 63352 | 585  | 
|
586  | 
val ((lift_const, (_ , def_thm)), lthy) =  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
587  | 
Local_Theory.define (var, ((def_name, []), newrhs)) lthy  | 
| 47308 | 588  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
589  | 
val transfer_rules = generate_transfer_rules lthy quot_thm rsp_thm def_thm par_thms  | 
| 
49975
 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 
huffman 
parents: 
49885 
diff
changeset
 | 
590  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
591  | 
val abs_eq_thm = generate_abs_eq lthy def_thm rsp_thm quot_thm  | 
| 
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
592  | 
val opt_rep_eq_thm = generate_rep_eq lthy def_thm rsp_thm (rty_forced, qty)  | 
| 47351 | 593  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
594  | 
fun notes names =  | 
| 
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
595  | 
let  | 
| 63341 | 596  | 
val lhs_name = Binding.reset_pos (#1 var)  | 
| 63003 | 597  | 
val rsp_thmN = Binding.qualify_name true lhs_name "rsp"  | 
598  | 
val abs_eq_thmN = Binding.qualify_name true lhs_name "abs_eq"  | 
|
599  | 
val rep_eq_thmN = Binding.qualify_name true lhs_name "rep_eq"  | 
|
600  | 
val transfer_ruleN = Binding.qualify_name true lhs_name "transfer"  | 
|
| 63352 | 601  | 
val notes =  | 
602  | 
[(rsp_thmN, [], [rsp_thm]),  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
603  | 
          (transfer_ruleN, @{attributes [transfer_rule]}, transfer_rules),
 | 
| 63352 | 604  | 
(abs_eq_thmN, [], [abs_eq_thm])]  | 
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
605  | 
@ (case opt_rep_eq_thm of SOME rep_eq_thm => [(rep_eq_thmN, [], [rep_eq_thm])] | NONE => [])  | 
| 
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
606  | 
in  | 
| 
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
607  | 
if names then map (fn (name, attrs, thms) => ((name, []), [(thms, attrs)])) notes  | 
| 63352 | 608  | 
else map_filter (fn (_, attrs, thms) => if null attrs then NONE  | 
609  | 
else SOME (Binding.empty_atts, [(thms, attrs)])) notes  | 
|
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
610  | 
end  | 
| 60227 | 611  | 
val (code_eq, lthy) = register_code_eq abs_eq_thm opt_rep_eq_thm (rty_forced, qty) lthy  | 
| 63352 | 612  | 
val lift_def = mk_lift_def rty_forced qty newrhs lift_const def_thm rsp_thm abs_eq_thm  | 
| 60227 | 613  | 
opt_rep_eq_thm code_eq transfer_rules  | 
| 47308 | 614  | 
in  | 
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
615  | 
lthy  | 
| 66017 | 616  | 
|> Local_Theory.open_target |> snd  | 
| 
60225
 
eb4e322734bf
note theorems by Local_Theory.notes (it is faster); make note of the generated theorems optional
 
kuncar 
parents: 
60224 
diff
changeset
 | 
617  | 
|> Local_Theory.notes (notes (#notes config)) |> snd  | 
| 66017 | 618  | 
|> `(fn lthy => morph_lift_def (Local_Theory.target_morphism lthy) lift_def)  | 
619  | 
||> Local_Theory.close_target  | 
|
| 47308 | 620  | 
end  | 
621  | 
||
| 56540 | 622  | 
(* This is not very cheap way of getting the rules but we have only few active  | 
623  | 
liftings in the current setting *)  | 
|
624  | 
fun get_cr_pcr_eqs ctxt =  | 
|
625  | 
let  | 
|
626  | 
fun collect (data : Lifting_Info.quotient) l =  | 
|
| 63352 | 627  | 
if is_some (#pcr_info data)  | 
628  | 
then ((Thm.symmetric o safe_mk_meta_eq o #pcr_cr_eq o the o #pcr_info) data :: l)  | 
|
| 56540 | 629  | 
else l  | 
630  | 
val table = Lifting_Info.get_quotients ctxt  | 
|
631  | 
in  | 
|
632  | 
Symtab.fold (fn (_, data) => fn l => collect data l) table []  | 
|
633  | 
end  | 
|
634  | 
||
| 60227 | 635  | 
fun prepare_lift_def add_lift_def var qty rhs par_thms lthy =  | 
| 47308 | 636  | 
let  | 
| 51994 | 637  | 
val rsp_rel = Lifting_Term.equiv_relation lthy (fastype_of rhs, qty)  | 
| 47308 | 638  | 
val rty_forced = (domain_type o fastype_of) rsp_rel;  | 
| 51994 | 639  | 
val forced_rhs = force_rty_type lthy rty_forced rhs;  | 
| 63352 | 640  | 
val cr_to_pcr_conv = HOLogic.Trueprop_conv (Conv.fun2_conv  | 
| 56540 | 641  | 
(Raw_Simplifier.rewrite lthy false (get_cr_pcr_eqs lthy)))  | 
642  | 
val (prsp_tm, rsp_prsp_eq) = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs)  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59586 
diff
changeset
 | 
643  | 
|> Thm.cterm_of lthy  | 
| 56540 | 644  | 
|> cr_to_pcr_conv  | 
| 59582 | 645  | 
|> `Thm.concl_of  | 
| 56540 | 646  | 
|>> Logic.dest_equals  | 
647  | 
|>> snd  | 
|
648  | 
val to_rsp = rsp_prsp_eq RS Drule.equal_elim_rule2  | 
|
649  | 
val opt_proven_rsp_thm = try_prove_reflexivity lthy prsp_tm  | 
|
| 63352 | 650  | 
|
651  | 
fun after_qed internal_rsp_thm lthy =  | 
|
| 60227 | 652  | 
add_lift_def var qty rhs (internal_rsp_thm RS to_rsp) par_thms lthy  | 
| 47308 | 653  | 
in  | 
| 
51374
 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 
kuncar 
parents: 
51314 
diff
changeset
 | 
654  | 
case opt_proven_rsp_thm of  | 
| 60218 | 655  | 
SOME thm => (NONE, K (after_qed thm))  | 
| 63352 | 656  | 
| NONE => (SOME prsp_tm, after_qed)  | 
| 47308 | 657  | 
end  | 
658  | 
||
| 60227 | 659  | 
fun gen_lift_def add_lift_def var qty rhs tac par_thms lthy =  | 
| 60218 | 660  | 
let  | 
| 60227 | 661  | 
val (goal, after_qed) = prepare_lift_def add_lift_def var qty rhs par_thms lthy  | 
| 60218 | 662  | 
in  | 
663  | 
case goal of  | 
|
| 63352 | 664  | 
SOME goal =>  | 
665  | 
let  | 
|
| 60218 | 666  | 
          val rsp_thm = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => tac ctxt)
 | 
667  | 
|> Thm.close_derivation  | 
|
668  | 
in  | 
|
| 
60231
 
0daab758e087
don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
 
kuncar 
parents: 
60230 
diff
changeset
 | 
669  | 
after_qed rsp_thm lthy  | 
| 60218 | 670  | 
end  | 
| 
60231
 
0daab758e087
don't use the human-readable version of the rsp thm as a goal in the ML interface (there is no formal definition of its statement); make tactics more robust wrt. predicates in predicators; tuned
 
kuncar 
parents: 
60230 
diff
changeset
 | 
671  | 
| NONE => after_qed Drule.dummy_thm lthy  | 
| 60218 | 672  | 
end  | 
673  | 
||
| 60227 | 674  | 
fun lift_def config var qty rhs tac par_thms lthy = gen_lift_def (add_lift_def config)  | 
675  | 
var qty rhs tac par_thms lthy  | 
|
676  | 
||
| 53651 | 677  | 
end (* structure *)  |