| author | blanchet | 
| Mon, 23 Jul 2012 15:32:30 +0200 | |
| changeset 48442 | 3c9890c19e90 | 
| parent 47982 | 7aa35601ff65 | 
| child 50175 | b27cf0646080 | 
| 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 | ||
| 47566 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 11 | val setup_by_quotient: bool -> thm -> thm option -> local_theory -> local_theory | 
| 47308 | 12 | |
| 47566 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 13 | val setup_by_typedef_thm: bool -> thm -> local_theory -> local_theory | 
| 47308 | 14 | end; | 
| 15 | ||
| 47334 | 16 | structure Lifting_Setup: LIFTING_SETUP = | 
| 47308 | 17 | struct | 
| 18 | ||
| 47698 | 19 | open Lifting_Util | 
| 47308 | 20 | |
| 47698 | 21 | infix 0 MRSL | 
| 47308 | 22 | |
| 23 | exception SETUP_LIFTING_INFR of string | |
| 24 | ||
| 47361 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 kuncar parents: 
47352diff
changeset | 25 | fun define_cr_rel rep_fun lthy = | 
| 47308 | 26 | let | 
| 47361 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 kuncar parents: 
47352diff
changeset | 27 | 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: 
47352diff
changeset | 28 | val rep_fun_graph = (HOLogic.eq_const rty) $ Bound 1 $ (rep_fun $ Bound 0) | 
| 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 kuncar parents: 
47352diff
changeset | 29 |     val def_term = Abs ("x", rty, Abs ("y", qty, rep_fun_graph));
 | 
| 47308 | 30 | val qty_name = (fst o dest_Type) qty | 
| 31 | val cr_rel_name = Binding.prefix_name "cr_" (Binding.qualified_name qty_name) | |
| 32 | val (fixed_def_term, lthy') = yield_singleton (Variable.importT_terms) def_term lthy | |
| 33 | val ((_, (_ , def_thm)), lthy'') = | |
| 34 | Local_Theory.define ((cr_rel_name, NoSyn), ((Thm.def_binding cr_rel_name, []), fixed_def_term)) lthy' | |
| 35 | in | |
| 36 | (def_thm, lthy'') | |
| 37 | end | |
| 38 | ||
| 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: 
47936diff
changeset | 39 | 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: 
47936diff
changeset | 40 | let | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47943diff
changeset | 41 | 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: 
47936diff
changeset | 42 | val abs_background = Morphism.term (Local_Theory.target_morphism lthy) abs | 
| 
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: 
47936diff
changeset | 43 | 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: 
47936diff
changeset | 44 | if gen_code andalso is_Const abs_background then | 
| 
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: 
47936diff
changeset | 45 | 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: 
47936diff
changeset | 46 | val (const_name, typ) = dest_Const abs_background | 
| 
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: 
47936diff
changeset | 47 | val fake_term = Logic.mk_type typ | 
| 
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: 
47936diff
changeset | 48 | val (fixed_fake_term, lthy') = yield_singleton(Variable.importT_terms) fake_term 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: 
47936diff
changeset | 49 | val fixed_type = Logic.dest_type fixed_fake_term | 
| 
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: 
47936diff
changeset | 50 | 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: 
47936diff
changeset | 51 | Local_Theory.background_theory(Code.add_datatype [(const_name, fixed_type)]) 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: 
47936diff
changeset | 52 | 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: 
47936diff
changeset | 53 | 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: 
47936diff
changeset | 54 | 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: 
47936diff
changeset | 55 | 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: 
47936diff
changeset | 56 | |
| 
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: 
47936diff
changeset | 57 | 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: 
47936diff
changeset | 58 | if gen_code andalso Lifting_Def.can_generate_code_cert quot_thm then | 
| 47308 | 59 | let | 
| 60 |       val abs_type_thm = quot_thm RS @{thm Quotient_abs_rep}
 | |
| 61 | val add_abstype_attribute = | |
| 62 | Thm.declaration_attribute (fn thm => Context.mapping (Code.add_abstype thm) I) | |
| 63 | val add_abstype_attrib = Attrib.internal (K add_abstype_attribute); | |
| 64 | in | |
| 65 | lthy | |
| 66 | |> (snd oo Local_Theory.note) ((Binding.empty, [add_abstype_attrib]), [abs_type_thm]) | |
| 67 | end | |
| 68 | else | |
| 69 | lthy | |
| 70 | ||
| 47379 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 71 | 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: 
47361diff
changeset | 72 | let | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 73 | val ((_, [quot_thm_fixed]), ctxt') = Variable.importT [quot_thm] ctxt | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47943diff
changeset | 74 | 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: 
47361diff
changeset | 75 | 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: 
47361diff
changeset | 76 | 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: 
47361diff
changeset | 77 | val extra_rty_tfrees = | 
| 47545 | 78 | 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: 
47361diff
changeset | 79 | [] => [] | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 80 | | 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: 
47361diff
changeset | 81 | Pretty.brk 1] @ | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 82 | ((Pretty.commas o map (Pretty.str o quote)) extras) @ | 
| 47545 | 83 | [Pretty.str "."])] | 
| 47379 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 84 | val not_type_constr = | 
| 47545 | 85 | case qty of | 
| 47379 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 86 | Type _ => [] | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 87 | | _ => [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: 
47361diff
changeset | 88 | 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: 
47361diff
changeset | 89 | Pretty.brk 1, | 
| 47545 | 90 | 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: 
47361diff
changeset | 91 | 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: 
47361diff
changeset | 92 | in | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 93 | if null errs then () else error (cat_lines (["Sanity check of the quotient theorem failed:",""] | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 94 | @ (map Pretty.string_of errs))) | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 95 | end | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 96 | |
| 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: 
47936diff
changeset | 97 | fun setup_lifting_infr gen_code quot_thm maybe_reflp_thm lthy = | 
| 47308 | 98 | let | 
| 47379 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47361diff
changeset | 99 | val _ = quot_thm_sanity_check lthy quot_thm | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47943diff
changeset | 100 | val (_, qtyp) = quot_thm_rty_qty quot_thm | 
| 47308 | 101 | val qty_full_name = (fst o dest_Type) qtyp | 
| 102 |     val quotients = { quot_thm = quot_thm }
 | |
| 103 | fun quot_info phi = Lifting_Info.transform_quotients phi quotients | |
| 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: 
47936diff
changeset | 104 | val lthy' = case maybe_reflp_thm of | 
| 
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: 
47936diff
changeset | 105 | SOME reflp_thm => lthy | 
| 47982 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 kuncar parents: 
47951diff
changeset | 106 | |> (snd oo Local_Theory.note) ((Binding.empty, [Lifting_Info.add_reflexivity_rule_attrib]), | 
| 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: 
47936diff
changeset | 107 | [reflp_thm]) | 
| 47982 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 kuncar parents: 
47951diff
changeset | 108 | |> (snd oo Local_Theory.note) ((Binding.empty, [Lifting_Info.add_reflexivity_rule_attrib]), | 
| 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 kuncar parents: 
47951diff
changeset | 109 |               [[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: 
47936diff
changeset | 110 | |> 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: 
47936diff
changeset | 111 | | 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: 
47936diff
changeset | 112 | |> define_abs_type gen_code quot_thm | 
| 47308 | 113 | 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: 
47936diff
changeset | 114 | lthy' | 
| 47308 | 115 |       |> Local_Theory.declaration {syntax = false, pervasive = true}
 | 
| 116 | (fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi)) | |
| 117 | end | |
| 118 | ||
| 47852 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 119 | (* | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 120 | Sets up the Lifting package by a quotient theorem. | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 121 | |
| 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: 
47936diff
changeset | 122 | 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: 
47779diff
changeset | 123 | as an abstract type in the code generator | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 124 | quot_thm - a quotient theorem (Quotient R Abs Rep T) | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 125 | maybe_reflp_thm - a theorem saying that a relation from quot_thm is reflexive | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 126 | (in the form "reflp R") | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 127 | *) | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 128 | |
| 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: 
47936diff
changeset | 129 | fun setup_by_quotient gen_code quot_thm maybe_reflp_thm lthy = | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 130 | let | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 131 | val transfer_attr = Attrib.internal (K Transfer.transfer_add) | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47943diff
changeset | 132 | val (_, qty) = quot_thm_rty_qty quot_thm | 
| 47575 | 133 | val induct_attr = Attrib.internal (K (Induct.induct_type (fst (dest_Type qty)))) | 
| 47583 | 134 | val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty | 
| 47545 | 135 | fun qualify suffix = Binding.qualified true suffix qty_name | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 136 | val lthy' = case maybe_reflp_thm of | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 137 | SOME reflp_thm => lthy | 
| 47545 | 138 | |> (snd oo Local_Theory.note) ((qualify "bi_total", [transfer_attr]), | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 139 |           [[quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}])
 | 
| 47545 | 140 | |> (snd oo Local_Theory.note) ((qualify "id_abs_transfer", [transfer_attr]), | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 141 |           [[quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer}])
 | 
| 47575 | 142 | |> (snd oo Local_Theory.note) ((qualify "abs_induct", [induct_attr]), | 
| 143 |           [[quot_thm, reflp_thm] MRSL @{thm Quotient_total_abs_induct}])
 | |
| 47889 
29212a4bb866
lifting package produces abs_eq_iff rules for total quotients
 huffman parents: 
47852diff
changeset | 144 | |> (snd oo Local_Theory.note) ((qualify "abs_eq_iff", []), | 
| 
29212a4bb866
lifting package produces abs_eq_iff rules for total quotients
 huffman parents: 
47852diff
changeset | 145 |           [[quot_thm, reflp_thm] MRSL @{thm Quotient_total_abs_eq_iff}])
 | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 146 | | NONE => lthy | 
| 47545 | 147 | |> (snd oo Local_Theory.note) ((qualify "All_transfer", [transfer_attr]), | 
| 148 |           [quot_thm RS @{thm Quotient_All_transfer}])
 | |
| 149 | |> (snd oo Local_Theory.note) ((qualify "Ex_transfer", [transfer_attr]), | |
| 150 |           [quot_thm RS @{thm Quotient_Ex_transfer}])
 | |
| 151 | |> (snd oo Local_Theory.note) ((qualify "forall_transfer", [transfer_attr]), | |
| 152 |           [quot_thm RS @{thm Quotient_forall_transfer}])
 | |
| 47575 | 153 | |> (snd oo Local_Theory.note) ((qualify "abs_induct", [induct_attr]), | 
| 154 |           [quot_thm RS @{thm Quotient_abs_induct}])
 | |
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 155 | in | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 156 | lthy' | 
| 47545 | 157 | |> (snd oo Local_Theory.note) ((qualify "right_unique", [transfer_attr]), | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 158 |         [quot_thm RS @{thm Quotient_right_unique}])
 | 
| 47545 | 159 | |> (snd oo Local_Theory.note) ((qualify "right_total", [transfer_attr]), | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 160 |         [quot_thm RS @{thm Quotient_right_total}])
 | 
| 47545 | 161 | |> (snd oo Local_Theory.note) ((qualify "rel_eq_transfer", [transfer_attr]), | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 162 |         [quot_thm RS @{thm Quotient_rel_eq_transfer}])
 | 
| 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: 
47936diff
changeset | 163 | |> setup_lifting_infr gen_code quot_thm maybe_reflp_thm | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 164 | end | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 165 | |
| 47852 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 166 | (* | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 167 | Sets up the Lifting package by a typedef theorem. | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 168 | |
| 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: 
47936diff
changeset | 169 | 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: 
47779diff
changeset | 170 | as an abstract type in the code generator | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 171 | typedef_thm - a typedef theorem (type_definition Rep Abs S) | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 172 | *) | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47779diff
changeset | 173 | |
| 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: 
47936diff
changeset | 174 | fun setup_by_typedef_thm gen_code typedef_thm lthy = | 
| 47308 | 175 | let | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 176 | val transfer_attr = Attrib.internal (K Transfer.transfer_add) | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 177 | val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o prop_of) typedef_thm | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 178 | val (T_def, lthy') = define_cr_rel rep_fun lthy | 
| 47545 | 179 | |
| 180 | val quot_thm = case typedef_set of | |
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 181 |       Const ("Orderings.top_class.top", _) => 
 | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 182 |         [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: 
47379diff
changeset | 183 |       | Const (@{const_name "Collect"}, _) $ Abs (_, _, _) => 
 | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 184 |         [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: 
47379diff
changeset | 185 | | _ => | 
| 47545 | 186 |         [typedef_thm, T_def] MRSL @{thm typedef_to_Quotient}
 | 
| 187 | ||
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47943diff
changeset | 188 | val (_, qty) = quot_thm_rty_qty quot_thm | 
| 47583 | 189 | val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty | 
| 47545 | 190 | fun qualify suffix = Binding.qualified true suffix qty_name | 
| 47623 
01e4fdf9d748
setup_lifting preprocesses forall_transfer rule by unfolding mem_Collect_eq
 huffman parents: 
47608diff
changeset | 191 |     val simplify = Raw_Simplifier.rewrite_rule [mk_meta_eq @{thm mem_Collect_eq}]
 | 
| 47308 | 192 | |
| 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: 
47936diff
changeset | 193 | val (maybe_reflp_thm, lthy'') = case typedef_set of | 
| 47545 | 194 |       Const ("Orderings.top_class.top", _) => 
 | 
| 195 | let | |
| 196 |           val equivp_thm = typedef_thm RS @{thm UNIV_typedef_to_equivp}
 | |
| 197 |           val reflp_thm = equivp_thm RS @{thm equivp_reflp2}
 | |
| 198 | in | |
| 199 | lthy' | |
| 200 | |> (snd oo Local_Theory.note) ((qualify "bi_total", [transfer_attr]), | |
| 201 |               [[quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}])
 | |
| 202 | |> (snd oo Local_Theory.note) ((qualify "id_abs_transfer", [transfer_attr]), | |
| 203 |               [[quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer}])
 | |
| 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: 
47936diff
changeset | 204 | |> pair (SOME reflp_thm) | 
| 47545 | 205 | end | 
| 206 | | _ => lthy' | |
| 207 | |> (snd oo Local_Theory.note) ((qualify "All_transfer", [transfer_attr]), | |
| 208 |           [[typedef_thm, T_def] MRSL @{thm typedef_All_transfer}])
 | |
| 209 | |> (snd oo Local_Theory.note) ((qualify "Ex_transfer", [transfer_attr]), | |
| 210 |           [[typedef_thm, T_def] MRSL @{thm typedef_Ex_transfer}])
 | |
| 211 | |> (snd oo Local_Theory.note) ((qualify "forall_transfer", [transfer_attr]), | |
| 47623 
01e4fdf9d748
setup_lifting preprocesses forall_transfer rule by unfolding mem_Collect_eq
 huffman parents: 
47608diff
changeset | 212 |           [simplify ([typedef_thm, T_def] MRSL @{thm typedef_forall_transfer})])
 | 
| 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: 
47936diff
changeset | 213 | |> pair NONE | 
| 47308 | 214 | in | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 215 | lthy'' | 
| 47943 
c09326cedb41
note Quotient theorem for typedefs in setup_lifting
 kuncar parents: 
47937diff
changeset | 216 | |> (snd oo Local_Theory.note) ((Binding.prefix_name "Quotient_" qty_name, []), | 
| 
c09326cedb41
note Quotient theorem for typedefs in setup_lifting
 kuncar parents: 
47937diff
changeset | 217 | [quot_thm]) | 
| 47545 | 218 | |> (snd oo Local_Theory.note) ((qualify "bi_unique", [transfer_attr]), | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 219 |         [[typedef_thm, T_def] MRSL @{thm typedef_bi_unique}])
 | 
| 47545 | 220 | |> (snd oo Local_Theory.note) ((qualify "rep_transfer", [transfer_attr]), | 
| 47535 
0f94b02fda1c
lifting_setup generates transfer rule for rep of typedefs
 huffman parents: 
47534diff
changeset | 221 |         [[typedef_thm, T_def] MRSL @{thm typedef_rep_transfer}])
 | 
| 47545 | 222 | |> (snd oo Local_Theory.note) ((qualify "right_unique", [transfer_attr]), | 
| 223 |         [[quot_thm] MRSL @{thm Quotient_right_unique}])
 | |
| 224 | |> (snd oo Local_Theory.note) ((qualify "right_total", [transfer_attr]), | |
| 225 |         [[quot_thm] MRSL @{thm Quotient_right_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: 
47936diff
changeset | 226 | |> setup_lifting_infr gen_code quot_thm maybe_reflp_thm | 
| 47308 | 227 | end | 
| 228 | ||
| 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: 
47936diff
changeset | 229 | fun setup_lifting_cmd gen_code xthm opt_reflp_xthm lthy = | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 230 | let | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 231 | 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: 
47379diff
changeset | 232 | 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: 
47545diff
changeset | 233 | 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: 
47545diff
changeset | 234 | |
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 235 | fun sanity_check_reflp_thm reflp_thm = | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 236 | let | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 237 | val reflp_tm = (HOLogic.dest_Trueprop o prop_of) reflp_thm | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 238 | 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: 
47545diff
changeset | 239 | in | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 240 | case reflp_tm of | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 241 |           Const (@{const_name reflp}, _) $ _ => ()
 | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 242 | | _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 243 | end | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 244 | |
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 245 | fun setup_quotient () = | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 246 | case opt_reflp_xthm of | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 247 | SOME reflp_xthm => | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 248 | let | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 249 | val reflp_thm = singleton (Attrib.eval_thms lthy) reflp_xthm | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 250 | val _ = sanity_check_reflp_thm reflp_thm | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 251 | 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: 
47936diff
changeset | 252 | setup_by_quotient gen_code input_thm (SOME reflp_thm) lthy | 
| 47566 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 253 | end | 
| 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: 
47936diff
changeset | 254 | | NONE => setup_by_quotient gen_code input_thm NONE lthy | 
| 47566 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 255 | |
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 256 | fun setup_typedef () = | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 257 | case opt_reflp_xthm of | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 258 | SOME _ => error "The reflexivity theorem cannot be specified if the type_definition theorem is used." | 
| 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: 
47936diff
changeset | 259 | | NONE => setup_by_typedef_thm gen_code input_thm lthy | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 260 | in | 
| 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 261 | case input_term of | 
| 47566 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 262 |       (Const (@{const_name Quotient}, _) $ _ $ _ $ _ $ _) => setup_quotient ()
 | 
| 
c201a1fe0a81
setup_lifting: no_code switch and supoport for quotient theorems
 kuncar parents: 
47545diff
changeset | 263 |       | (Const (@{const_name type_definition}, _) $ _ $ _ $ _) => setup_typedef ()
 | 
| 47521 
69f95ac85c3d
tuned the setup of lifting; generate transfer rules for typedef and Quotient thms
 kuncar parents: 
47379diff
changeset | 264 | | _ => 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: 
47379diff
changeset | 265 | end | 
| 47308 | 266 | |
| 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: 
47936diff
changeset | 267 | 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: 
47936diff
changeset | 268 |   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: 
47545diff
changeset | 269 | |
| 47308 | 270 | val _ = | 
| 271 |   Outer_Syntax.local_theory @{command_spec "setup_lifting"}
 | |
| 47352 | 272 | "Setup lifting infrastructure" | 
| 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: 
47936diff
changeset | 273 | (opt_gen_code -- Parse_Spec.xthm -- Scan.option Parse_Spec.xthm >> | 
| 
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: 
47936diff
changeset | 274 | (fn ((gen_code, xthm), opt_reflp_xthm) => setup_lifting_cmd gen_code xthm opt_reflp_xthm)) | 
| 47308 | 275 | end; |