| author | blanchet | 
| Tue, 19 Aug 2014 09:34:57 +0200 | |
| changeset 57993 | c52255a71114 | 
| parent 57961 | 10b2f60b70f0 | 
| child 58821 | 11e226e8a095 | 
| permissions | -rw-r--r-- | 
| 47308 | 1 | (* Title: HOL/Tools/Lifting/lifting_info.ML | 
| 2 | Author: Ondrej Kuncar | |
| 3 | ||
| 4 | Context data for the lifting package. | |
| 5 | *) | |
| 6 | ||
| 7 | signature LIFTING_INFO = | |
| 8 | sig | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 9 |   type quot_map = {rel_quot_thm: thm}
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 10 | val lookup_quot_maps: Proof.context -> string -> quot_map option | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 11 | val print_quot_maps: Proof.context -> unit | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 12 | |
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 13 |   type pcr = {pcrel_def: thm, pcr_cr_eq: thm}
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 14 |   type quotient = {quot_thm: thm, pcr_info: pcr option}
 | 
| 53651 | 15 | val pcr_eq: pcr * pcr -> bool | 
| 16 | val quotient_eq: quotient * quotient -> bool | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 17 | val transform_quotient: morphism -> quotient -> quotient | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 18 | val lookup_quotients: Proof.context -> string -> quotient option | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 19 | val update_quotients: string -> quotient -> Context.generic -> Context.generic | 
| 53650 | 20 | val delete_quotients: thm -> Context.generic -> Context.generic | 
| 47308 | 21 | val print_quotients: Proof.context -> unit | 
| 22 | ||
| 53651 | 23 |   type restore_data = {quotient : quotient, transfer_rules: thm Item_Net.T}
 | 
| 24 | val lookup_restore_data: Proof.context -> string -> restore_data option | |
| 25 | val init_restore_data: string -> quotient -> Context.generic -> Context.generic | |
| 26 | val add_transfer_rules_in_restore_data: string -> thm Item_Net.T -> Context.generic -> Context.generic | |
| 27 | ||
| 56519 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 kuncar parents: 
56518diff
changeset | 28 | val get_relator_eq_onp_rules: Proof.context -> thm list | 
| 47936 
756f30eac792
infrastructure that makes possible to prove that a relation is reflexive
 kuncar parents: 
47784diff
changeset | 29 | |
| 47982 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 kuncar parents: 
47951diff
changeset | 30 | val get_reflexivity_rules: Proof.context -> thm list | 
| 
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
 kuncar parents: 
47951diff
changeset | 31 | val add_reflexivity_rule_attribute: attribute | 
| 47634 
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
 kuncar parents: 
47308diff
changeset | 32 | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 33 |   type relator_distr_data = {pos_mono_rule: thm, neg_mono_rule: thm, 
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 34 | pos_distr_rules: thm list, neg_distr_rules: thm list} | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 35 | val lookup_relator_distr_data: Proof.context -> string -> relator_distr_data option | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 36 | |
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 37 | val add_no_code_type: string -> Context.generic -> Context.generic | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 38 | val is_no_code_type: Proof.context -> string -> bool | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 39 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 40 | val get_quot_maps : Proof.context -> quot_map Symtab.table | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 41 | val get_quotients : Proof.context -> quotient Symtab.table | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 42 | val get_relator_distr_data : Proof.context -> relator_distr_data Symtab.table | 
| 53651 | 43 | val get_restore_data : Proof.context -> restore_data Symtab.table | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 44 | val get_no_code_types : Proof.context -> unit Symtab.table | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 45 | |
| 47308 | 46 | val setup: theory -> theory | 
| 53651 | 47 | end | 
| 47308 | 48 | |
| 49 | structure Lifting_Info: LIFTING_INFO = | |
| 50 | struct | |
| 51 | ||
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 52 | open Lifting_Util | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 53 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 54 | (** data container **) | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 55 | |
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 56 | type quot_map = {rel_quot_thm: thm}
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 57 | type pcr = {pcrel_def: thm, pcr_cr_eq: thm}
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 58 | type quotient = {quot_thm: thm, pcr_info: pcr option}
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 59 | type relator_distr_data = {pos_mono_rule: thm, neg_mono_rule: thm, 
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 60 | pos_distr_rules: thm list, neg_distr_rules: thm list} | 
| 53651 | 61 | type restore_data = {quotient : quotient, transfer_rules: thm Item_Net.T}
 | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 62 | |
| 53684 | 63 | fun pcr_eq ({pcrel_def = pcrel_def1, pcr_cr_eq = pcr_cr_eq1},
 | 
| 64 |            {pcrel_def = pcrel_def2, pcr_cr_eq = pcr_cr_eq2}) = 
 | |
| 65 | Thm.eq_thm (pcrel_def1, pcrel_def2) andalso Thm.eq_thm (pcr_cr_eq1, pcr_cr_eq2) | |
| 66 | ||
| 67 | fun option_eq _ (NONE,NONE) = true | |
| 68 | | option_eq _ (NONE,_) = false | |
| 69 | | option_eq _ (_,NONE) = false | |
| 70 | | option_eq cmp (SOME x, SOME y) = cmp (x,y); | |
| 71 | ||
| 72 | fun quotient_eq ({quot_thm = quot_thm1, pcr_info = pcr_info1},
 | |
| 73 |                 {quot_thm = quot_thm2, pcr_info = pcr_info2}) =
 | |
| 74 | Thm.eq_thm (quot_thm1, quot_thm2) andalso option_eq pcr_eq (pcr_info1, pcr_info2) | |
| 75 | ||
| 53754 | 76 | fun join_restore_data key (rd1:restore_data, rd2) = | 
| 53684 | 77 | if pointer_eq (rd1, rd2) then raise Symtab.SAME else | 
| 78 | if not (quotient_eq (#quotient rd1, #quotient rd2)) then raise Symtab.DUP key else | |
| 79 |     { quotient = #quotient rd1, 
 | |
| 80 | transfer_rules = Item_Net.merge (#transfer_rules rd1, #transfer_rules rd2)} | |
| 81 | ||
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 82 | structure Data = Generic_Data | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 83 | ( | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 84 | type T = | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 85 |     { quot_maps : quot_map Symtab.table,
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 86 | quotients : quotient Symtab.table, | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 87 | reflexivity_rules : thm Item_Net.T, | 
| 53651 | 88 | relator_distr_data : relator_distr_data Symtab.table, | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 89 | restore_data : restore_data Symtab.table, | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 90 | no_code_types : unit Symtab.table | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 91 | } | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 92 | val empty = | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 93 |     { quot_maps = Symtab.empty,
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 94 | quotients = Symtab.empty, | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 95 | reflexivity_rules = Thm.full_rules, | 
| 53651 | 96 | relator_distr_data = Symtab.empty, | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 97 | restore_data = Symtab.empty, | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 98 | no_code_types = Symtab.empty | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 99 | } | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 100 | val extend = I | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 101 | fun merge | 
| 53651 | 102 |     ( { quot_maps = qm1, quotients = q1, reflexivity_rules = rr1, relator_distr_data = rdd1, 
 | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 103 | restore_data = rd1, no_code_types = nct1 }, | 
| 53651 | 104 |       { quot_maps = qm2, quotients = q2, reflexivity_rules = rr2, relator_distr_data = rdd2,
 | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 105 | restore_data = rd2, no_code_types = nct2 } ) = | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 106 |     { quot_maps = Symtab.merge (K true) (qm1, qm2),
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 107 | quotients = Symtab.merge (K true) (q1, q2), | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 108 | reflexivity_rules = Item_Net.merge (rr1, rr2), | 
| 53651 | 109 | relator_distr_data = Symtab.merge (K true) (rdd1, rdd2), | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 110 | restore_data = Symtab.join join_restore_data (rd1, rd2), | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 111 | no_code_types = Symtab.merge (K true) (nct1, nct2) | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 112 | } | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 113 | ) | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 114 | |
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 115 | fun map_data f1 f2 f3 f4 f5 f6 | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 116 |   { quot_maps, quotients, reflexivity_rules, relator_distr_data, restore_data, no_code_types } =
 | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 117 |   { quot_maps = f1 quot_maps,
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 118 | quotients = f2 quotients, | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 119 | reflexivity_rules = f3 reflexivity_rules, | 
| 53651 | 120 | relator_distr_data = f4 relator_distr_data, | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 121 | restore_data = f5 restore_data, | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 122 | no_code_types = f6 no_code_types | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 123 | } | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 124 | |
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 125 | fun map_quot_maps f = map_data f I I I I I | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 126 | fun map_quotients f = map_data I f I I I I | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 127 | fun map_reflexivity_rules f = map_data I I f I I I | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 128 | fun map_relator_distr_data f = map_data I I I f I I | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 129 | fun map_restore_data f = map_data I I I I f I | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 130 | fun map_no_code_types f = map_data I I I I I f | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 131 | |
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 132 | val get_quot_maps' = #quot_maps o Data.get | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 133 | val get_quotients' = #quotients o Data.get | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 134 | val get_reflexivity_rules' = #reflexivity_rules o Data.get | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 135 | val get_relator_distr_data' = #relator_distr_data o Data.get | 
| 53651 | 136 | val get_restore_data' = #restore_data o Data.get | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 137 | val get_no_code_types' = #no_code_types o Data.get | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 138 | |
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 139 | fun get_quot_maps ctxt = get_quot_maps' (Context.Proof ctxt) | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 140 | fun get_quotients ctxt = get_quotients' (Context.Proof ctxt) | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 141 | fun get_relator_distr_data ctxt = get_relator_distr_data' (Context.Proof ctxt) | 
| 53651 | 142 | fun get_restore_data ctxt = get_restore_data' (Context.Proof ctxt) | 
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 143 | fun get_no_code_types ctxt = get_no_code_types' (Context.Proof ctxt) | 
| 47308 | 144 | |
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 145 | (* info about Quotient map theorems *) | 
| 47308 | 146 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 147 | val lookup_quot_maps = Symtab.lookup o get_quot_maps | 
| 47308 | 148 | |
| 47784 | 149 | fun quot_map_thm_sanity_check rel_quot_thm ctxt = | 
| 150 | let | |
| 151 | fun quot_term_absT ctxt quot_term = | |
| 152 | let | |
| 153 | val (_, abs, _, _) = (dest_Quotient o HOLogic.dest_Trueprop) quot_term | |
| 154 | handle TERM (_, [t]) => error (Pretty.string_of (Pretty.block | |
| 155 | [Pretty.str "The Quotient map theorem is not in the right form.", | |
| 156 | Pretty.brk 1, | |
| 157 | Pretty.str "The following term is not the Quotient predicate:", | |
| 158 | Pretty.brk 1, | |
| 159 | Syntax.pretty_term ctxt t])) | |
| 160 | in | |
| 161 | fastype_of abs | |
| 162 | end | |
| 163 | ||
| 164 | val ((_, [rel_quot_thm_fixed]), ctxt') = Variable.importT [rel_quot_thm] ctxt | |
| 165 | val rel_quot_thm_prop = prop_of rel_quot_thm_fixed | |
| 166 | val rel_quot_thm_concl = Logic.strip_imp_concl rel_quot_thm_prop | |
| 167 | val rel_quot_thm_prems = Logic.strip_imp_prems rel_quot_thm_prop; | |
| 168 | val concl_absT = quot_term_absT ctxt' rel_quot_thm_concl | |
| 169 | val concl_tfrees = Term.add_tfree_namesT (concl_absT) [] | |
| 170 | val prems_tfrees = fold (fn typ => fn list => Term.add_tfree_namesT (quot_term_absT ctxt' typ) list) | |
| 171 | rel_quot_thm_prems [] | |
| 172 | val extra_prem_tfrees = | |
| 173 | case subtract (op =) concl_tfrees prems_tfrees of | |
| 174 | [] => [] | |
| 175 | | extras => [Pretty.block ([Pretty.str "Extra type variables in the premises:", | |
| 176 | Pretty.brk 1] @ | |
| 177 | ((Pretty.commas o map (Pretty.str o quote)) extras) @ | |
| 178 | [Pretty.str "."])] | |
| 179 | val errs = extra_prem_tfrees | |
| 180 | in | |
| 181 | if null errs then () else error (cat_lines (["Sanity check of the quotient map theorem failed:",""] | |
| 182 | @ (map Pretty.string_of errs))) | |
| 183 | end | |
| 184 | ||
| 185 | ||
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 186 | fun add_quot_map rel_quot_thm ctxt = | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 187 | let | 
| 47784 | 188 | val _ = Context.cases (K ()) (quot_map_thm_sanity_check rel_quot_thm) ctxt | 
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 189 | val rel_quot_thm_concl = (Logic.strip_imp_concl o prop_of) rel_quot_thm | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 190 | val (_, abs, _, _) = (dest_Quotient o HOLogic.dest_Trueprop) rel_quot_thm_concl | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 191 | val relatorT_name = (fst o dest_Type o fst o dest_funT o fastype_of) abs | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 192 |     val minfo = {rel_quot_thm = rel_quot_thm}
 | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 193 | in | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 194 | Data.map (map_quot_maps (Symtab.update (relatorT_name, minfo))) ctxt | 
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 195 | end | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 196 | |
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 197 | val quot_map_attribute_setup = | 
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 198 |   Attrib.setup @{binding quot_map} (Scan.succeed (Thm.declaration_attribute add_quot_map))
 | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 199 | "declaration of the Quotient map theorem" | 
| 47308 | 200 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 201 | fun print_quot_maps ctxt = | 
| 47308 | 202 | let | 
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 203 |     fun prt_map (ty_name, {rel_quot_thm}) =
 | 
| 47308 | 204 | Pretty.block (separate (Pretty.brk 2) | 
| 205 | [Pretty.str "type:", | |
| 206 | Pretty.str ty_name, | |
| 207 | Pretty.str "quot. theorem:", | |
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47634diff
changeset | 208 | Syntax.pretty_term ctxt (prop_of rel_quot_thm)]) | 
| 47308 | 209 | in | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 210 | map prt_map (Symtab.dest (get_quot_maps ctxt)) | 
| 47308 | 211 | |> Pretty.big_list "maps for type constructors:" | 
| 212 | |> Pretty.writeln | |
| 213 | end | |
| 214 | ||
| 215 | (* info about quotient types *) | |
| 53651 | 216 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 217 | fun transform_pcr_info phi {pcrel_def, pcr_cr_eq} =
 | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 218 |   {pcrel_def = Morphism.thm phi pcrel_def, pcr_cr_eq = Morphism.thm phi pcr_cr_eq}
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 219 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 220 | fun transform_quotient phi {quot_thm, pcr_info} =
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 221 |   {quot_thm = Morphism.thm phi quot_thm, pcr_info = Option.map (transform_pcr_info phi) pcr_info}
 | 
| 47308 | 222 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 223 | fun lookup_quotients ctxt type_name = Symtab.lookup (get_quotients ctxt) type_name | 
| 47308 | 224 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 225 | fun update_quotients type_name qinfo ctxt = | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 226 | Data.map (map_quotients (Symtab.update (type_name, qinfo))) ctxt | 
| 47308 | 227 | |
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 228 | fun delete_quotients quot_thm ctxt = | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 229 | let | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 230 | val (_, qtyp) = quot_thm_rty_qty quot_thm | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 231 | val qty_full_name = (fst o dest_Type) qtyp | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 232 | val symtab = get_quotients' ctxt | 
| 53284 | 233 | fun compare_data (_, data:quotient) = Thm.eq_thm_prop (#quot_thm data, quot_thm) | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 234 | in | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 235 | if Symtab.member compare_data symtab (qty_full_name, quot_thm) | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 236 | then Data.map (map_quotients (Symtab.delete qty_full_name)) ctxt | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 237 | else ctxt | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 238 | end | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 239 | |
| 47308 | 240 | fun print_quotients ctxt = | 
| 241 | let | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 242 |     fun prt_quot (qty_name, {quot_thm, pcr_info}: quotient) =
 | 
| 47308 | 243 | Pretty.block (separate (Pretty.brk 2) | 
| 244 | [Pretty.str "type:", | |
| 245 | Pretty.str qty_name, | |
| 246 | Pretty.str "quot. thm:", | |
| 50227 | 247 | Syntax.pretty_term ctxt (prop_of quot_thm), | 
| 248 | Pretty.str "pcrel_def thm:", | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 249 | option_fold (Pretty.str "-") ((Syntax.pretty_term ctxt) o prop_of o #pcrel_def) pcr_info, | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 250 | Pretty.str "pcr_cr_eq thm:", | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 251 | option_fold (Pretty.str "-") ((Syntax.pretty_term ctxt) o prop_of o #pcr_cr_eq) pcr_info]) | 
| 47308 | 252 | in | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 253 | map prt_quot (Symtab.dest (get_quotients ctxt)) | 
| 47308 | 254 | |> Pretty.big_list "quotients:" | 
| 255 | |> Pretty.writeln | |
| 256 | end | |
| 257 | ||
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 258 | val quot_del_attribute_setup = | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 259 |   Attrib.setup @{binding quot_del} (Scan.succeed (Thm.declaration_attribute delete_quotients))
 | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 260 | "deletes the Quotient theorem" | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 261 | |
| 53651 | 262 | (* data for restoring Transfer/Lifting context *) | 
| 263 | ||
| 264 | fun lookup_restore_data ctxt bundle_name = Symtab.lookup (get_restore_data ctxt) bundle_name | |
| 265 | ||
| 266 | fun update_restore_data bundle_name restore_data ctxt = | |
| 267 | Data.map (map_restore_data (Symtab.update (bundle_name, restore_data))) ctxt | |
| 268 | ||
| 269 | fun init_restore_data bundle_name qinfo ctxt = | |
| 270 |   update_restore_data bundle_name { quotient = qinfo, transfer_rules = Thm.full_rules } ctxt
 | |
| 271 | ||
| 272 | fun add_transfer_rules_in_restore_data bundle_name transfer_rules ctxt = | |
| 273 | case Symtab.lookup (get_restore_data' ctxt) bundle_name of | |
| 274 |     SOME restore_data => update_restore_data bundle_name { quotient = #quotient restore_data, 
 | |
| 275 | transfer_rules = Item_Net.merge ((#transfer_rules restore_data), transfer_rules) } ctxt | |
| 276 |     | NONE => error ("The restore data " ^ quote bundle_name ^ " is not defined.")
 | |
| 277 | ||
| 56519 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 kuncar parents: 
56518diff
changeset | 278 | (* theorems that a relator of an eq_onp is an eq_onp of the corresponding predicate *) | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 279 | |
| 57961 | 280 | fun get_relator_eq_onp_rules ctxt = | 
| 281 |   map safe_mk_meta_eq (rev (Named_Theorems.get ctxt @{named_theorems relator_eq_onp}))
 | |
| 47634 
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
 kuncar parents: 
47308diff
changeset | 282 | |
| 51994 | 283 | (* info about reflexivity rules *) | 
| 284 | ||
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 285 | fun get_reflexivity_rules ctxt = Item_Net.content (get_reflexivity_rules' (Context.Proof ctxt)) | 
| 51994 | 286 | |
| 55563 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 287 | fun add_reflexivity_rule thm = Data.map (map_reflexivity_rules (Item_Net.update thm)) | 
| 51994 | 288 | val add_reflexivity_rule_attribute = Thm.declaration_attribute add_reflexivity_rule | 
| 289 | ||
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 290 | (* info about relator distributivity theorems *) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 291 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 292 | fun map_relator_distr_data' f1 f2 f3 f4 | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 293 |   {pos_mono_rule, neg_mono_rule, pos_distr_rules, neg_distr_rules} =
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 294 |   {pos_mono_rule   = f1 pos_mono_rule, 
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 295 | neg_mono_rule = f2 neg_mono_rule, | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 296 | pos_distr_rules = f3 pos_distr_rules, | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 297 | neg_distr_rules = f4 neg_distr_rules} | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 298 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 299 | fun map_pos_mono_rule f = map_relator_distr_data' f I I I | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 300 | fun map_neg_mono_rule f = map_relator_distr_data' I f I I | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 301 | fun map_pos_distr_rules f = map_relator_distr_data' I I f I | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 302 | fun map_neg_distr_rules f = map_relator_distr_data' I I I f | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 303 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 304 | fun introduce_polarities rule = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 305 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 306 |     val dest_less_eq = HOLogic.dest_bin @{const_name "less_eq"} dummyT
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 307 | val prems_pairs = map (dest_less_eq o HOLogic.dest_Trueprop) (prems_of rule) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 308 | val equal_prems = filter op= prems_pairs | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 309 | val _ = if null equal_prems then () | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 310 | else error "The rule contains reflexive assumptions." | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 311 | val concl_pairs = rule | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 312 | |> concl_of | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 313 | |> HOLogic.dest_Trueprop | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 314 | |> dest_less_eq | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 315 | |> pairself (snd o strip_comb) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 316 | |> op~~ | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 317 | |> filter_out op= | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 318 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 319 | val _ = if has_duplicates op= concl_pairs | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 320 | then error "The rule contains duplicated variables in the conlusion." else () | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 321 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 322 | fun rewrite_prem prem_pair = | 
| 51427 
08bb00239652
proper use of "member", without embarking on delicate questions about SML equality types;
 wenzelm parents: 
51426diff
changeset | 323 | if member op= concl_pairs prem_pair | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 324 |       then HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric @{thm POS_def}))
 | 
| 51427 
08bb00239652
proper use of "member", without embarking on delicate questions about SML equality types;
 wenzelm parents: 
51426diff
changeset | 325 | else if member op= concl_pairs (swap prem_pair) | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 326 |         then HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric @{thm NEG_def}))
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 327 | else error "The rule contains a non-relevant assumption." | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 328 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 329 | fun rewrite_prems [] = Conv.all_conv | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 330 | | rewrite_prems (x::xs) = Conv.implies_conv (rewrite_prem x) (rewrite_prems xs) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 331 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 332 | val rewrite_prems_conv = rewrite_prems prems_pairs | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 333 | val rewrite_concl_conv = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 334 |       Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric @{thm POS_def})))
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 335 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 336 | (Conv.fconv_rule (rewrite_prems_conv then_conv rewrite_concl_conv)) rule | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 337 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 338 | handle | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 339 | TERM _ => error "The rule has a wrong format." | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 340 | | CTERM _ => error "The rule has a wrong format." | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 341 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 342 | fun negate_mono_rule mono_rule = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 343 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 344 |     val rewr_conv = HOLogic.Trueprop_conv (Conv.rewrs_conv [@{thm POS_NEG}, @{thm NEG_POS}])
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 345 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 346 | Conv.fconv_rule (Conv.prems_conv ~1 rewr_conv then_conv Conv.concl_conv ~1 rewr_conv) mono_rule | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 347 | end; | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 348 | |
| 55563 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 349 | fun add_reflexivity_rules mono_rule ctxt = | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 350 | let | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 351 | fun find_eq_rule thm ctxt = | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 352 | let | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 353 | val concl_rhs = (hd o get_args 1 o HOLogic.dest_Trueprop o concl_of) thm; | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 354 | val rules = Item_Net.retrieve (Transfer.get_relator_eq_item_net ctxt) concl_rhs; | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 355 | in | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 356 | find_first (fn thm => Pattern.matches (Proof_Context.theory_of ctxt) (concl_rhs, | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 357 | (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.concl_of) thm)) rules | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 358 | end | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 359 | |
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 360 | val eq_rule = find_eq_rule mono_rule (Context.proof_of ctxt); | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 361 | val eq_rule = if is_some eq_rule then the eq_rule else error | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 362 | "No corresponding rule that the relator preserves equality was found." | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 363 | in | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 364 | ctxt | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 365 |     |> add_reflexivity_rule (Drule.zero_var_indexes (@{thm ord_le_eq_trans} OF [mono_rule, eq_rule]))
 | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 366 | |> add_reflexivity_rule | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 367 |       (Drule.zero_var_indexes (@{thm ord_eq_le_trans} OF [sym OF [eq_rule], mono_rule]))
 | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 368 | end | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 369 | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 370 | fun add_mono_rule mono_rule ctxt = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 371 | let | 
| 55563 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 372 | val pol_mono_rule = introduce_polarities mono_rule | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 373 | val mono_ruleT_name = (fst o dest_Type o fst o relation_types o fst o relation_types o snd o | 
| 55563 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 374 | dest_Const o head_of o HOLogic.dest_Trueprop o concl_of) pol_mono_rule | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 375 | val _ = if Symtab.defined (get_relator_distr_data' ctxt) mono_ruleT_name | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 376 |       then error ("Monotocity rule for type " ^ quote mono_ruleT_name ^ " is already_defined.")
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 377 | else () | 
| 55563 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 378 | val neg_mono_rule = negate_mono_rule pol_mono_rule | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 379 |     val relator_distr_data = {pos_mono_rule = pol_mono_rule, neg_mono_rule = neg_mono_rule, 
 | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 380 | pos_distr_rules = [], neg_distr_rules = []} | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 381 | in | 
| 55563 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 382 | ctxt | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 383 | |> Data.map (map_relator_distr_data (Symtab.update (mono_ruleT_name, relator_distr_data))) | 
| 
a64d49f49ca3
implement the reflexivity prover as a monotonicity prover that proves R >= op=; derive "reflexivity" rules for relators from mono rules and eq rules
 kuncar parents: 
53754diff
changeset | 384 | |> add_reflexivity_rules mono_rule | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 385 | end; | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 386 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 387 | local | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 388 | fun add_distr_rule update_entry distr_rule ctxt = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 389 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 390 | val distr_ruleT_name = (fst o dest_Type o fst o relation_types o fst o relation_types o snd o | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 391 | dest_Const o head_of o HOLogic.dest_Trueprop o concl_of) distr_rule | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 392 | in | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 393 | if Symtab.defined (get_relator_distr_data' ctxt) distr_ruleT_name then | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 394 | Data.map (map_relator_distr_data (Symtab.map_entry distr_ruleT_name (update_entry distr_rule))) | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 395 | ctxt | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 396 | else error "The monoticity rule is not defined." | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 397 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 398 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 399 | fun rewrite_concl_conv thm ctm = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 400 | Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric thm))) ctm | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 401 | handle CTERM _ => error "The rule has a wrong format." | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 402 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 403 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 404 | fun add_pos_distr_rule distr_rule ctxt = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 405 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 406 |       val distr_rule = Conv.fconv_rule (rewrite_concl_conv @{thm POS_def}) distr_rule
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 407 | fun update_entry distr_rule data = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 408 |         map_pos_distr_rules (cons (@{thm POS_trans} OF [distr_rule, #pos_mono_rule data])) data
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 409 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 410 | add_distr_rule update_entry distr_rule ctxt | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 411 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 412 | handle THM _ => error "Combining of the distr. rule and the monotonicity rule together has failed." | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 413 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 414 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 415 | fun add_neg_distr_rule distr_rule ctxt = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 416 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 417 |       val distr_rule = Conv.fconv_rule (rewrite_concl_conv @{thm NEG_def}) distr_rule
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 418 | fun update_entry distr_rule data = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 419 |         map_neg_distr_rules (cons (@{thm NEG_trans} OF [distr_rule, #neg_mono_rule data])) data
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 420 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 421 | add_distr_rule update_entry distr_rule ctxt | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 422 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 423 | handle THM _ => error "Combining of the distr. rule and the monotonicity rule together has failed." | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 424 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 425 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 426 | local | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 427 |   val eq_refl2 = sym RS @{thm eq_refl}
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 428 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 429 | fun add_eq_distr_rule distr_rule ctxt = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 430 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 431 |       val pos_distr_rule = @{thm eq_refl} OF [distr_rule]
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 432 | val neg_distr_rule = eq_refl2 OF [distr_rule] | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 433 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 434 | ctxt | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 435 | |> add_pos_distr_rule pos_distr_rule | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 436 | |> add_neg_distr_rule neg_distr_rule | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 437 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 438 | end; | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 439 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 440 | local | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 441 | fun sanity_check rule = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 442 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 443 | val assms = map (perhaps (try HOLogic.dest_Trueprop)) (prems_of rule) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 444 | val concl = (perhaps (try HOLogic.dest_Trueprop)) (concl_of rule); | 
| 56257 | 445 | val (lhs, rhs) = | 
| 446 | (case concl of | |
| 447 |           Const (@{const_name less_eq}, _) $ (lhs as Const (@{const_name relcompp},_) $ _ $ _) $ rhs =>
 | |
| 448 | (lhs, rhs) | |
| 449 |         | Const (@{const_name less_eq}, _) $ rhs $ (lhs as Const (@{const_name relcompp},_) $ _ $ _) =>
 | |
| 450 | (lhs, rhs) | |
| 451 |         | Const (@{const_name HOL.eq}, _) $ (lhs as Const (@{const_name relcompp},_) $ _ $ _) $ rhs =>
 | |
| 452 | (lhs, rhs) | |
| 453 | | _ => error "The rule has a wrong format.") | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 454 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 455 | val lhs_vars = Term.add_vars lhs [] | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 456 | val rhs_vars = Term.add_vars rhs [] | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 457 | val assms_vars = fold Term.add_vars assms []; | 
| 56257 | 458 | val _ = | 
| 459 | if has_duplicates op= lhs_vars | |
| 460 | then error "Left-hand side has variable duplicates" else () | |
| 461 | val _ = | |
| 462 | if subset op= (rhs_vars, lhs_vars) then () | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 463 | else error "Extra variables in the right-hand side of the rule" | 
| 56257 | 464 | val _ = | 
| 465 | if subset op= (assms_vars, lhs_vars) then () | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 466 | else error "Extra variables in the assumptions of the rule" | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 467 | val rhs_args = (snd o strip_comb) rhs; | 
| 56257 | 468 | fun check_comp t = | 
| 469 | (case t of | |
| 470 |           Const (@{const_name relcompp}, _) $ Var _ $ Var _ => ()
 | |
| 471 | | _ => error "There is an argument on the rhs that is not a composition.") | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 472 | val _ = map check_comp rhs_args | 
| 56257 | 473 | in () end | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 474 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 475 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 476 | fun add_distr_rule distr_rule ctxt = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 477 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 478 | val _ = sanity_check distr_rule | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 479 | val concl = (perhaps (try HOLogic.dest_Trueprop)) (concl_of distr_rule) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 480 | in | 
| 56257 | 481 | (case concl of | 
| 482 |         Const (@{const_name less_eq}, _) $ (Const (@{const_name relcompp},_) $ _ $ _) $ _ =>
 | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 483 | add_pos_distr_rule distr_rule ctxt | 
| 56257 | 484 |       | Const (@{const_name less_eq}, _) $ _ $ (Const (@{const_name relcompp},_) $ _ $ _) =>
 | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 485 | add_neg_distr_rule distr_rule ctxt | 
| 56257 | 486 |       | Const (@{const_name HOL.eq}, _) $ (Const (@{const_name relcompp},_) $ _ $ _) $ _ =>
 | 
| 487 | add_eq_distr_rule distr_rule ctxt) | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 488 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 489 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 490 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 491 | fun get_distr_rules_raw ctxt = Symtab.fold | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 492 |   (fn (_, {pos_distr_rules, neg_distr_rules, ...}) => fn rules => pos_distr_rules @ neg_distr_rules @ rules) 
 | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 493 | (get_relator_distr_data' ctxt) [] | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 494 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 495 | fun get_mono_rules_raw ctxt = Symtab.fold | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 496 |   (fn (_, {pos_mono_rule, neg_mono_rule, ...}) => fn rules => [pos_mono_rule, neg_mono_rule] @ rules) 
 | 
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 497 | (get_relator_distr_data' ctxt) [] | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 498 | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 499 | val lookup_relator_distr_data = Symtab.lookup o get_relator_distr_data | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 500 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 501 | val relator_distr_attribute_setup = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 502 |   Attrib.setup @{binding relator_mono} (Scan.succeed (Thm.declaration_attribute add_mono_rule))
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 503 | "declaration of relator's monoticity" | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 504 |   #> Attrib.setup @{binding relator_distr} (Scan.succeed (Thm.declaration_attribute add_distr_rule))
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 505 | "declaration of relator's distributivity over OO" | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 506 | #> Global_Theory.add_thms_dynamic | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 507 |      (@{binding relator_distr_raw}, get_distr_rules_raw)
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 508 | #> Global_Theory.add_thms_dynamic | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 509 |      (@{binding relator_mono_raw}, get_mono_rules_raw)
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 510 | |
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 511 | (* no_code types *) | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 512 | |
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 513 | fun add_no_code_type type_name ctxt = | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 514 | Data.map (map_no_code_types (Symtab.update (type_name, ()))) ctxt; | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 515 | |
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 516 | fun is_no_code_type ctxt type_name = (Symtab.defined o get_no_code_types) ctxt type_name | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
56519diff
changeset | 517 | |
| 47308 | 518 | (* theory setup *) | 
| 519 | ||
| 520 | val setup = | |
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 521 | quot_map_attribute_setup | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47936diff
changeset | 522 | #> quot_del_attribute_setup | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50227diff
changeset | 523 | #> relator_distr_attribute_setup | 
| 47308 | 524 | |
| 56519 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 kuncar parents: 
56518diff
changeset | 525 | (* setup fixed eq_onp rules *) | 
| 55731 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55563diff
changeset | 526 | |
| 57961 | 527 | val _ = Context.>> | 
| 528 |   (fold (Named_Theorems.add_thm @{named_theorems relator_eq_onp} o
 | |
| 529 |     Transfer.prep_transfer_domain_thm @{context})
 | |
| 56519 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 kuncar parents: 
56518diff
changeset | 530 |   @{thms composed_equiv_rel_eq_onp composed_equiv_rel_eq_eq_onp})
 | 
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56257diff
changeset | 531 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56257diff
changeset | 532 | (* setup fixed reflexivity rules *) | 
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56257diff
changeset | 533 | |
| 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56257diff
changeset | 534 | val _ = Context.>> (fold add_reflexivity_rule | 
| 56519 
c1048f5bbb45
more appropriate name (Lifting.invariant -> eq_onp)
 kuncar parents: 
56518diff
changeset | 535 |   @{thms order_refl[of "op="] eq_onp_le_eq Quotient_composition_le_eq Quotient_composition_ge_eq
 | 
| 56518 
beb3b6851665
left_total and left_unique rules are now transfer rules (cleaner solution, reflexvity_rule attribute not needed anymore)
 kuncar parents: 
56257diff
changeset | 536 | bi_unique_OO bi_total_OO right_unique_OO right_total_OO left_unique_OO left_total_OO}) | 
| 55731 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55563diff
changeset | 537 | |
| 47308 | 538 | (* outer syntax commands *) | 
| 539 | ||
| 540 | val _ = | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 541 |   Outer_Syntax.improper_command @{command_spec "print_quot_maps"} "print quotient map functions"
 | 
| 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51994diff
changeset | 542 | (Scan.succeed (Toplevel.keep (print_quot_maps o Toplevel.context_of))) | 
| 47308 | 543 | |
| 544 | val _ = | |
| 545 |   Outer_Syntax.improper_command @{command_spec "print_quotients"} "print quotients"
 | |
| 546 | (Scan.succeed (Toplevel.keep (print_quotients o Toplevel.context_of))) | |
| 547 | ||
| 53651 | 548 | end |