| author | wenzelm | 
| Thu, 11 Aug 2016 18:26:16 +0200 | |
| changeset 63668 | 5efaa884ac6c | 
| parent 60642 | 48dd1cefb4ae | 
| child 67703 | 8c4806fe827f | 
| permissions | -rw-r--r-- | 
| 47698 | 1 | (* Title: HOL/Tools/Lifting/lifting_util.ML | 
| 2 | Author: Ondrej Kuncar | |
| 3 | ||
| 4 | General-purpose functions used by the Lifting package. | |
| 5 | *) | |
| 6 | ||
| 7 | signature LIFTING_UTIL = | |
| 8 | sig | |
| 9 | val MRSL: thm list * thm -> thm | |
| 50225 | 10 |   val option_fold: 'b -> ('a -> 'b) -> 'a option -> 'b
 | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 11 |   val map_snd: ('b -> 'c) -> ('a * 'b) list -> ('a * 'c) list
 | 
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47699diff
changeset | 12 | val dest_Quotient: term -> term * term * term * term | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 13 | |
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 14 | val quot_thm_rel: thm -> term | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 15 | val quot_thm_abs: thm -> term | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 16 | val quot_thm_rep: thm -> term | 
| 50226 | 17 | val quot_thm_crel: thm -> term | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 18 | val quot_thm_rty_qty: thm -> typ * typ | 
| 60229 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 19 | val Quotient_conv: conv -> conv -> conv -> conv -> conv | 
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 20 | val Quotient_R_conv: conv -> conv | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 21 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 22 | val undisch: thm -> thm | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 23 | val undisch_all: thm -> thm | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 24 | val is_fun_type: typ -> bool | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 25 | val get_args: int -> term -> term list | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 26 | val strip_args: int -> term -> term | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 27 | val all_args_conv: conv -> conv | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 28 | val is_Type: typ -> bool | 
| 55454 | 29 | val same_type_constrs: typ * typ -> bool | 
| 30 | val Targs: typ -> typ list | |
| 31 | val Tname: typ -> string | |
| 55945 | 32 | val is_rel_fun: term -> bool | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 33 | val relation_types: typ -> typ * typ | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 34 | val mk_HOL_eq: thm -> thm | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 35 | val safe_HOL_meta_eq: thm -> thm | 
| 55731 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 36 |   val map_interrupt: ('a -> 'b option) -> 'a list -> 'b list option
 | 
| 60223 | 37 | val instT_thm: Proof.context -> Type.tyenv -> thm -> thm | 
| 38 | val instT_morphism: Proof.context -> Type.tyenv -> morphism | |
| 60229 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 39 | val conceal_naming_result: (local_theory -> 'a * local_theory) -> local_theory -> 'a * local_theory | 
| 53651 | 40 | end | 
| 47698 | 41 | |
| 42 | ||
| 43 | structure Lifting_Util: LIFTING_UTIL = | |
| 44 | struct | |
| 45 | ||
| 46 | infix 0 MRSL | |
| 47 | ||
| 48 | fun ants MRSL thm = fold (fn rl => fn thm => rl RS thm) ants thm | |
| 49 | ||
| 50225 | 50 | fun option_fold a _ NONE = a | 
| 51 | | option_fold _ f (SOME x) = f x | |
| 52 | ||
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 53 | fun map_snd f xs = map (fn (a, b) => (a, f b)) xs | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 54 | |
| 47777 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47699diff
changeset | 55 | fun dest_Quotient (Const (@{const_name Quotient}, _) $ rel $ abs $ rep $ cr)
 | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47699diff
changeset | 56 | = (rel, abs, rep, cr) | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47699diff
changeset | 57 |   | dest_Quotient t = raise TERM ("dest_Quotient", [t])
 | 
| 
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
 kuncar parents: 
47699diff
changeset | 58 | |
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 59 | (* | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 60 | quot_thm_rel, quot_thm_abs, quot_thm_rep and quot_thm_rty_qty - simple functions | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 61 | for destructing quotient theorems (Quotient R Abs Rep T). | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 62 | *) | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 63 | |
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 64 | fun quot_thm_rel quot_thm = | 
| 59582 | 65 | case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 66 | (rel, _, _, _) => rel | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 67 | |
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 68 | fun quot_thm_abs quot_thm = | 
| 59582 | 69 | case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 70 | (_, abs, _, _) => abs | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 71 | |
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 72 | fun quot_thm_rep quot_thm = | 
| 59582 | 73 | case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of | 
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 74 | (_, _, rep, _) => rep | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 75 | |
| 50226 | 76 | fun quot_thm_crel quot_thm = | 
| 59582 | 77 | case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of | 
| 50226 | 78 | (_, _, _, crel) => crel | 
| 79 | ||
| 47951 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 80 | fun quot_thm_rty_qty quot_thm = | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 81 | let | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 82 | val abs = quot_thm_abs quot_thm | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 83 | val abs_type = fastype_of abs | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 84 | in | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 85 | (domain_type abs_type, range_type abs_type) | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 86 | end | 
| 
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
 kuncar parents: 
47777diff
changeset | 87 | |
| 60229 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 88 | fun Quotient_conv R_conv Abs_conv Rep_conv T_conv = Conv.combination_conv (Conv.combination_conv | 
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 89 | (Conv.combination_conv (Conv.arg_conv R_conv) Abs_conv) Rep_conv) T_conv; | 
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 90 | |
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 91 | fun Quotient_R_conv R_conv = Quotient_conv R_conv Conv.all_conv Conv.all_conv Conv.all_conv; | 
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 92 | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 93 | fun undisch thm = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 94 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 95 | val assm = Thm.cprem_of thm 1 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 96 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 97 | Thm.implies_elim thm (Thm.assume assm) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 98 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 99 | |
| 59582 | 100 | fun undisch_all thm = funpow (Thm.nprems_of thm) undisch thm | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 101 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 102 | fun is_fun_type (Type (@{type_name fun}, _)) = true
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 103 | | is_fun_type _ = false | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 104 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 105 | fun get_args n = rev o fst o funpow_yield n (swap o dest_comb) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 106 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 107 | fun strip_args n = funpow n (fst o dest_comb) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 108 | |
| 55456 | 109 | fun all_args_conv conv ctm = Conv.try_conv (Conv.combination_conv (all_args_conv conv) conv) ctm | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 110 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 111 | fun is_Type (Type _) = true | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 112 | | is_Type _ = false | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 113 | |
| 55454 | 114 | fun same_type_constrs (Type (r, _), Type (q, _)) = (r = q) | 
| 115 | | same_type_constrs _ = false | |
| 116 | ||
| 117 | fun Targs (Type (_, args)) = args | |
| 118 | | Targs _ = [] | |
| 119 | ||
| 120 | fun Tname (Type (name, _)) = name | |
| 121 | | Tname _ = "" | |
| 122 | ||
| 55945 | 123 | fun is_rel_fun (Const (@{const_name "rel_fun"}, _) $ _ $ _) = true
 | 
| 124 | | is_rel_fun _ = false | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 125 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 126 | fun relation_types typ = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 127 | case strip_type typ of | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 128 |     ([typ1, typ2], @{typ bool}) => (typ1, typ2)
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 129 | | _ => error "relation_types: not a relation" | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 130 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 131 | fun mk_HOL_eq r = r RS @{thm meta_eq_to_obj_eq}
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 132 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 133 | fun safe_HOL_meta_eq r = mk_HOL_eq r handle Thm.THM _ => r | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
51314diff
changeset | 134 | |
| 55731 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 135 | fun map_interrupt f l = | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 136 | let | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 137 | fun map_interrupt' _ [] l = SOME (rev l) | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 138 | | map_interrupt' f (x::xs) l = (case f x of | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 139 | NONE => NONE | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 140 | | SOME v => map_interrupt' f xs (v::l)) | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 141 | in | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 142 | map_interrupt' f l [] | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 143 | end | 
| 
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
 kuncar parents: 
55456diff
changeset | 144 | |
| 60223 | 145 | fun instT_thm ctxt env = | 
| 146 | let | |
| 147 | val cinst = env |> Vartab.dest | |
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60229diff
changeset | 148 | |> map (fn (x, (S, T)) => ((x, S), Thm.ctyp_of ctxt T)); | 
| 60223 | 149 | in | 
| 150 | Thm.instantiate (cinst, []) | |
| 151 | end; | |
| 152 | ||
| 153 | fun instT_morphism ctxt env = | |
| 154 | Morphism.morphism "Lifting_Util.instT" | |
| 155 |     {binding = [],
 | |
| 156 | typ = [Envir.subst_type env], | |
| 157 | term = [Envir.subst_term_types env], | |
| 158 | fact = [map (instT_thm ctxt env)]}; | |
| 159 | ||
| 60229 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 160 | fun conceal_naming_result f lthy = | 
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 161 | let val old_lthy = lthy | 
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 162 | in lthy |> Proof_Context.concealed |> f ||> Proof_Context.restore_naming old_lthy end; | 
| 
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
 kuncar parents: 
60223diff
changeset | 163 | |
| 53651 | 164 | end |