author | wenzelm |
Tue, 04 Jun 2019 15:14:56 +0200 | |
changeset 70320 | 59258a3192bf |
parent 69593 | 3dda49e08b9d |
child 70323 | 2943934a169d |
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 |
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47699
diff
changeset
|
11 |
val dest_Quotient: term -> term * term * term * term |
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
12 |
|
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
13 |
val quot_thm_rel: thm -> term |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
14 |
val quot_thm_abs: thm -> term |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
15 |
val quot_thm_rep: thm -> term |
50226 | 16 |
val quot_thm_crel: thm -> term |
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
17 |
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:
60223
diff
changeset
|
18 |
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:
60223
diff
changeset
|
19 |
val Quotient_R_conv: conv -> conv |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
20 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
21 |
val undisch: thm -> thm |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
22 |
val undisch_all: thm -> thm |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
23 |
val is_fun_type: typ -> bool |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
24 |
val get_args: int -> term -> term list |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
25 |
val strip_args: int -> term -> term |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
26 |
val all_args_conv: conv -> conv |
55454 | 27 |
val same_type_constrs: typ * typ -> bool |
28 |
val Targs: typ -> typ list |
|
29 |
val Tname: typ -> string |
|
55945 | 30 |
val is_rel_fun: term -> bool |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
31 |
val relation_types: typ -> typ * typ |
55731
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents:
55456
diff
changeset
|
32 |
val map_interrupt: ('a -> 'b option) -> 'a list -> 'b list option |
60229
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
60223
diff
changeset
|
33 |
val conceal_naming_result: (local_theory -> 'a * local_theory) -> local_theory -> 'a * local_theory |
53651 | 34 |
end |
47698 | 35 |
|
36 |
||
37 |
structure Lifting_Util: LIFTING_UTIL = |
|
38 |
struct |
|
39 |
||
40 |
infix 0 MRSL |
|
41 |
||
42 |
fun ants MRSL thm = fold (fn rl => fn thm => rl RS thm) ants thm |
|
43 |
||
50225 | 44 |
fun option_fold a _ NONE = a |
45 |
| option_fold _ f (SOME x) = f x |
|
46 |
||
69593 | 47 |
fun dest_Quotient (Const (\<^const_name>\<open>Quotient\<close>, _) $ rel $ abs $ rep $ cr) |
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47699
diff
changeset
|
48 |
= (rel, abs, rep, cr) |
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47699
diff
changeset
|
49 |
| dest_Quotient t = raise TERM ("dest_Quotient", [t]) |
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47699
diff
changeset
|
50 |
|
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
51 |
(* |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
52 |
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:
47777
diff
changeset
|
53 |
for destructing quotient theorems (Quotient R Abs Rep T). |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
54 |
*) |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
55 |
|
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
56 |
fun quot_thm_rel quot_thm = |
59582 | 57 |
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:
47777
diff
changeset
|
58 |
(rel, _, _, _) => rel |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
59 |
|
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
60 |
fun quot_thm_abs quot_thm = |
59582 | 61 |
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:
47777
diff
changeset
|
62 |
(_, abs, _, _) => abs |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
63 |
|
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
64 |
fun quot_thm_rep 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:
47777
diff
changeset
|
66 |
(_, _, rep, _) => rep |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
67 |
|
50226 | 68 |
fun quot_thm_crel quot_thm = |
59582 | 69 |
case (dest_Quotient o HOLogic.dest_Trueprop o Thm.prop_of) quot_thm of |
50226 | 70 |
(_, _, _, crel) => crel |
71 |
||
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
72 |
fun quot_thm_rty_qty quot_thm = |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
73 |
let |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
74 |
val abs = quot_thm_abs quot_thm |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
75 |
val abs_type = fastype_of abs |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
76 |
in |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
77 |
(domain_type abs_type, range_type abs_type) |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
78 |
end |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47777
diff
changeset
|
79 |
|
60229
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
60223
diff
changeset
|
80 |
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:
60223
diff
changeset
|
81 |
(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:
60223
diff
changeset
|
82 |
|
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
60223
diff
changeset
|
83 |
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:
60223
diff
changeset
|
84 |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
85 |
fun undisch thm = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
86 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
87 |
val assm = Thm.cprem_of thm 1 |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
88 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
89 |
Thm.implies_elim thm (Thm.assume assm) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
90 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
91 |
|
59582 | 92 |
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:
51314
diff
changeset
|
93 |
|
69593 | 94 |
fun is_fun_type (Type (\<^type_name>\<open>fun\<close>, _)) = true |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
95 |
| is_fun_type _ = false |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
96 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
97 |
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:
51314
diff
changeset
|
98 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
99 |
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:
51314
diff
changeset
|
100 |
|
55456 | 101 |
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:
51314
diff
changeset
|
102 |
|
55454 | 103 |
fun same_type_constrs (Type (r, _), Type (q, _)) = (r = q) |
104 |
| same_type_constrs _ = false |
|
105 |
||
106 |
fun Targs (Type (_, args)) = args |
|
107 |
| Targs _ = [] |
|
108 |
||
109 |
fun Tname (Type (name, _)) = name |
|
110 |
| Tname _ = "" |
|
111 |
||
69593 | 112 |
fun is_rel_fun (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) = true |
55945 | 113 |
| is_rel_fun _ = false |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
114 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
115 |
fun relation_types typ = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
116 |
case strip_type typ of |
69593 | 117 |
([typ1, typ2], \<^typ>\<open>bool\<close>) => (typ1, typ2) |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
118 |
| _ => error "relation_types: not a relation" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
51314
diff
changeset
|
119 |
|
55731
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents:
55456
diff
changeset
|
120 |
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:
55456
diff
changeset
|
121 |
let |
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents:
55456
diff
changeset
|
122 |
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:
55456
diff
changeset
|
123 |
| 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:
55456
diff
changeset
|
124 |
NONE => NONE |
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents:
55456
diff
changeset
|
125 |
| 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:
55456
diff
changeset
|
126 |
in |
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents:
55456
diff
changeset
|
127 |
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:
55456
diff
changeset
|
128 |
end |
66df76dd2640
rewrite composition of quotients to a more readable form in a respectfulness goal that is presented to a user
kuncar
parents:
55456
diff
changeset
|
129 |
|
60229
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
60223
diff
changeset
|
130 |
fun conceal_naming_result f lthy = |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
69593
diff
changeset
|
131 |
lthy |> Proof_Context.concealed |> f ||> Proof_Context.restore_naming lthy; |
60229
4cd6462c1fda
Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype
kuncar
parents:
60223
diff
changeset
|
132 |
|
53651 | 133 |
end |