author | kuncar |
Thu, 29 Aug 2013 20:15:13 +0200 | |
changeset 53284 | d0153a0a9b2b |
parent 53219 | ca237b9e4542 |
child 53650 | 71a0a8687d6c |
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:
51994
diff
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:
51994
diff
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:
51994
diff
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:
51994
diff
changeset
|
12 |
|
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
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:
51994
diff
changeset
|
14 |
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:
51994
diff
changeset
|
15 |
val transform_quotient: morphism -> quotient -> quotient |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
16 |
val lookup_quotients: Proof.context -> string -> quotient option |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
17 |
val update_quotients: string -> quotient -> Context.generic -> Context.generic |
47308 | 18 |
val print_quotients: Proof.context -> unit |
19 |
||
47634
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
20 |
val get_invariant_commute_rules: Proof.context -> thm list |
47936
756f30eac792
infrastructure that makes possible to prove that a relation is reflexive
kuncar
parents:
47784
diff
changeset
|
21 |
|
47982
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47951
diff
changeset
|
22 |
val get_reflexivity_rules: Proof.context -> thm list |
51994 | 23 |
val add_reflexivity_rule_raw_attribute: attribute |
47982
7aa35601ff65
prove reflexivity also for the quotient composition relation; reflp_preserve renamed to reflexivity_rule
kuncar
parents:
47951
diff
changeset
|
24 |
val add_reflexivity_rule_attribute: attribute |
47634
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
25 |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
26 |
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:
50227
diff
changeset
|
27 |
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:
50227
diff
changeset
|
28 |
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:
51994
diff
changeset
|
29 |
|
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
30 |
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:
51994
diff
changeset
|
31 |
val get_quotients : Proof.context -> quotient Symtab.table |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
32 |
val get_relator_distr_data : Proof.context -> relator_distr_data Symtab.table |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
33 |
|
47308 | 34 |
val setup: theory -> theory |
35 |
end; |
|
36 |
||
37 |
structure Lifting_Info: LIFTING_INFO = |
|
38 |
struct |
|
39 |
||
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
40 |
open Lifting_Util |
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
41 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
42 |
(** data container **) |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
43 |
|
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
44 |
type quot_map = {rel_quot_thm: thm} |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
45 |
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:
51994
diff
changeset
|
46 |
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:
51994
diff
changeset
|
47 |
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:
51994
diff
changeset
|
48 |
pos_distr_rules: thm list, neg_distr_rules: thm list} |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
49 |
|
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
50 |
structure Data = Generic_Data |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
51 |
( |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
52 |
type T = |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
53 |
{ quot_maps : quot_map Symtab.table, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
54 |
quotients : quotient Symtab.table, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
55 |
reflexivity_rules : thm Item_Net.T, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
56 |
relator_distr_data : relator_distr_data Symtab.table |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
57 |
} |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
58 |
val empty = |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
59 |
{ quot_maps = Symtab.empty, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
60 |
quotients = Symtab.empty, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
61 |
reflexivity_rules = Thm.full_rules, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
62 |
relator_distr_data = Symtab.empty |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
63 |
} |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
64 |
val extend = I |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
65 |
fun merge |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
66 |
( { quot_maps = qm1, quotients = q1, reflexivity_rules = rr1, relator_distr_data = rdd1 }, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
67 |
{ quot_maps = qm2, quotients = q2, reflexivity_rules = rr2, relator_distr_data = rdd2 } ) = |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
68 |
{ quot_maps = Symtab.merge (K true) (qm1, qm2), |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
69 |
quotients = Symtab.merge (K true) (q1, q2), |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
70 |
reflexivity_rules = Item_Net.merge (rr1, rr2), |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
71 |
relator_distr_data = Symtab.merge (K true) (rdd1, rdd2) } |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
72 |
) |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
73 |
|
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
74 |
fun map_data f1 f2 f3 f4 |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
75 |
{ quot_maps, quotients, reflexivity_rules, relator_distr_data} = |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
76 |
{ quot_maps = f1 quot_maps, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
77 |
quotients = f2 quotients, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
78 |
reflexivity_rules = f3 reflexivity_rules, |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
79 |
relator_distr_data = f4 relator_distr_data } |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
80 |
|
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
81 |
fun map_quot_maps f = map_data f I I I |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
82 |
fun map_quotients f = map_data I f I I |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
83 |
fun map_reflexivity_rules f = map_data I I f I |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
84 |
fun map_relator_distr_data f = map_data I I I f |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
85 |
|
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
86 |
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:
51994
diff
changeset
|
87 |
val get_quotients' = #quotients o Data.get |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
88 |
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:
51994
diff
changeset
|
89 |
val get_relator_distr_data' = #relator_distr_data o Data.get |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
90 |
|
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
91 |
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:
51994
diff
changeset
|
92 |
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:
51994
diff
changeset
|
93 |
fun get_relator_distr_data ctxt = get_relator_distr_data' (Context.Proof ctxt) |
47308 | 94 |
|
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
95 |
(* info about Quotient map theorems *) |
47308 | 96 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
97 |
val lookup_quot_maps = Symtab.lookup o get_quot_maps |
47308 | 98 |
|
47784 | 99 |
fun quot_map_thm_sanity_check rel_quot_thm ctxt = |
100 |
let |
|
101 |
fun quot_term_absT ctxt quot_term = |
|
102 |
let |
|
103 |
val (_, abs, _, _) = (dest_Quotient o HOLogic.dest_Trueprop) quot_term |
|
104 |
handle TERM (_, [t]) => error (Pretty.string_of (Pretty.block |
|
105 |
[Pretty.str "The Quotient map theorem is not in the right form.", |
|
106 |
Pretty.brk 1, |
|
107 |
Pretty.str "The following term is not the Quotient predicate:", |
|
108 |
Pretty.brk 1, |
|
109 |
Syntax.pretty_term ctxt t])) |
|
110 |
in |
|
111 |
fastype_of abs |
|
112 |
end |
|
113 |
||
114 |
val ((_, [rel_quot_thm_fixed]), ctxt') = Variable.importT [rel_quot_thm] ctxt |
|
115 |
val rel_quot_thm_prop = prop_of rel_quot_thm_fixed |
|
116 |
val rel_quot_thm_concl = Logic.strip_imp_concl rel_quot_thm_prop |
|
117 |
val rel_quot_thm_prems = Logic.strip_imp_prems rel_quot_thm_prop; |
|
118 |
val concl_absT = quot_term_absT ctxt' rel_quot_thm_concl |
|
119 |
val concl_tfrees = Term.add_tfree_namesT (concl_absT) [] |
|
120 |
val prems_tfrees = fold (fn typ => fn list => Term.add_tfree_namesT (quot_term_absT ctxt' typ) list) |
|
121 |
rel_quot_thm_prems [] |
|
122 |
val extra_prem_tfrees = |
|
123 |
case subtract (op =) concl_tfrees prems_tfrees of |
|
124 |
[] => [] |
|
125 |
| extras => [Pretty.block ([Pretty.str "Extra type variables in the premises:", |
|
126 |
Pretty.brk 1] @ |
|
127 |
((Pretty.commas o map (Pretty.str o quote)) extras) @ |
|
128 |
[Pretty.str "."])] |
|
129 |
val errs = extra_prem_tfrees |
|
130 |
in |
|
131 |
if null errs then () else error (cat_lines (["Sanity check of the quotient map theorem failed:",""] |
|
132 |
@ (map Pretty.string_of errs))) |
|
133 |
end |
|
134 |
||
135 |
||
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
136 |
fun add_quot_map rel_quot_thm ctxt = |
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
137 |
let |
47784 | 138 |
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:
47634
diff
changeset
|
139 |
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:
47634
diff
changeset
|
140 |
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:
47634
diff
changeset
|
141 |
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:
47634
diff
changeset
|
142 |
val minfo = {rel_quot_thm = rel_quot_thm} |
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
143 |
in |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
144 |
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:
47634
diff
changeset
|
145 |
end |
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
146 |
|
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
147 |
val quot_map_attribute_setup = |
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
148 |
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:
47634
diff
changeset
|
149 |
"declaration of the Quotient map theorem" |
47308 | 150 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
151 |
fun print_quot_maps ctxt = |
47308 | 152 |
let |
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
153 |
fun prt_map (ty_name, {rel_quot_thm}) = |
47308 | 154 |
Pretty.block (separate (Pretty.brk 2) |
155 |
[Pretty.str "type:", |
|
156 |
Pretty.str ty_name, |
|
157 |
Pretty.str "quot. theorem:", |
|
47777
f29e7dcd7c40
use a quot_map theorem attribute instead of the complicated map attribute
kuncar
parents:
47634
diff
changeset
|
158 |
Syntax.pretty_term ctxt (prop_of rel_quot_thm)]) |
47308 | 159 |
in |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
160 |
map prt_map (Symtab.dest (get_quot_maps ctxt)) |
47308 | 161 |
|> Pretty.big_list "maps for type constructors:" |
162 |
|> Pretty.writeln |
|
163 |
end |
|
164 |
||
165 |
(* info about quotient types *) |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
166 |
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:
50227
diff
changeset
|
167 |
{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:
50227
diff
changeset
|
168 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
169 |
fun transform_quotient phi {quot_thm, pcr_info} = |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
170 |
{quot_thm = Morphism.thm phi quot_thm, pcr_info = Option.map (transform_pcr_info phi) pcr_info} |
47308 | 171 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
172 |
fun lookup_quotients ctxt type_name = Symtab.lookup (get_quotients ctxt) type_name |
47308 | 173 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
174 |
fun update_quotients type_name qinfo ctxt = |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
175 |
Data.map (map_quotients (Symtab.update (type_name, qinfo))) ctxt |
47308 | 176 |
|
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
177 |
fun delete_quotients quot_thm ctxt = |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
178 |
let |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
179 |
val (_, qtyp) = quot_thm_rty_qty quot_thm |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
180 |
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:
51994
diff
changeset
|
181 |
val symtab = get_quotients' ctxt |
53284 | 182 |
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:
47936
diff
changeset
|
183 |
in |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
184 |
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:
51994
diff
changeset
|
185 |
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:
51994
diff
changeset
|
186 |
else ctxt |
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
187 |
end |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
188 |
|
47308 | 189 |
fun print_quotients ctxt = |
190 |
let |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
191 |
fun prt_quot (qty_name, {quot_thm, pcr_info}: quotient) = |
47308 | 192 |
Pretty.block (separate (Pretty.brk 2) |
193 |
[Pretty.str "type:", |
|
194 |
Pretty.str qty_name, |
|
195 |
Pretty.str "quot. thm:", |
|
50227 | 196 |
Syntax.pretty_term ctxt (prop_of quot_thm), |
197 |
Pretty.str "pcrel_def thm:", |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
198 |
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:
50227
diff
changeset
|
199 |
Pretty.str "pcr_cr_eq thm:", |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
200 |
option_fold (Pretty.str "-") ((Syntax.pretty_term ctxt) o prop_of o #pcr_cr_eq) pcr_info]) |
47308 | 201 |
in |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
202 |
map prt_quot (Symtab.dest (get_quotients ctxt)) |
47308 | 203 |
|> Pretty.big_list "quotients:" |
204 |
|> Pretty.writeln |
|
205 |
end |
|
206 |
||
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
207 |
val quot_del_attribute_setup = |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
208 |
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:
47936
diff
changeset
|
209 |
"deletes the Quotient theorem" |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
210 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
211 |
(* theorems that a relator of an invariant is an invariant of the corresponding predicate *) |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
212 |
|
47634
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
213 |
structure Invariant_Commute = Named_Thms |
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
214 |
( |
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
215 |
val name = @{binding invariant_commute} |
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
216 |
val description = "theorems that a relator of an invariant is an invariant of the corresponding predicate" |
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
217 |
) |
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
218 |
|
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
219 |
fun get_invariant_commute_rules ctxt = map safe_mk_meta_eq (Invariant_Commute.get ctxt) |
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
220 |
|
51994 | 221 |
(* info about reflexivity rules *) |
222 |
||
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
223 |
fun get_reflexivity_rules ctxt = Item_Net.content (get_reflexivity_rules' (Context.Proof ctxt)) |
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
224 |
|
51994 | 225 |
|
226 |
(* Conversion to create a reflp' variant of a reflexivity rule *) |
|
227 |
fun safe_reflp_conv ct = |
|
228 |
Conv.try_conv (HOLogic.Trueprop_conv (Conv.rewr_conv (Thm.symmetric @{thm reflp'_def}))) ct |
|
229 |
||
230 |
fun prep_reflp_conv ct = ( |
|
231 |
Conv.implies_conv safe_reflp_conv prep_reflp_conv |
|
232 |
else_conv |
|
233 |
safe_reflp_conv |
|
234 |
else_conv |
|
235 |
Conv.all_conv) ct |
|
236 |
||
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
237 |
fun add_reflexivity_rule_raw thm = Data.map (map_reflexivity_rules (Item_Net.update thm)) |
51994 | 238 |
val add_reflexivity_rule_raw_attribute = Thm.declaration_attribute add_reflexivity_rule_raw |
239 |
||
240 |
fun add_reflexivity_rule thm = add_reflexivity_rule_raw thm #> |
|
241 |
add_reflexivity_rule_raw (Conv.fconv_rule prep_reflp_conv thm) |
|
242 |
val add_reflexivity_rule_attribute = Thm.declaration_attribute add_reflexivity_rule |
|
243 |
||
244 |
||
245 |
val relfexivity_rule_setup = |
|
246 |
let |
|
247 |
val name = @{binding reflexivity_rule} |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
248 |
fun del_thm_raw thm = Data.map (map_reflexivity_rules (Item_Net.remove thm)) |
51994 | 249 |
fun del_thm thm = del_thm_raw thm #> |
250 |
del_thm_raw (Conv.fconv_rule prep_reflp_conv thm) |
|
251 |
val del = Thm.declaration_attribute del_thm |
|
252 |
val text = "rules that are used to prove that a relation is reflexive" |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
253 |
val content = Item_Net.content o get_reflexivity_rules' |
51994 | 254 |
in |
255 |
Attrib.setup name (Attrib.add_del add_reflexivity_rule_attribute del) text |
|
256 |
#> Global_Theory.add_thms_dynamic (name, content) |
|
257 |
end |
|
47936
756f30eac792
infrastructure that makes possible to prove that a relation is reflexive
kuncar
parents:
47784
diff
changeset
|
258 |
|
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
259 |
(* info about relator distributivity theorems *) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
260 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
261 |
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:
50227
diff
changeset
|
262 |
{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:
50227
diff
changeset
|
263 |
{pos_mono_rule = f1 pos_mono_rule, |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
264 |
neg_mono_rule = f2 neg_mono_rule, |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
265 |
pos_distr_rules = f3 pos_distr_rules, |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
266 |
neg_distr_rules = f4 neg_distr_rules} |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
267 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
268 |
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:
51994
diff
changeset
|
269 |
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:
51994
diff
changeset
|
270 |
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:
51994
diff
changeset
|
271 |
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:
50227
diff
changeset
|
272 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
273 |
fun introduce_polarities rule = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
274 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
275 |
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:
50227
diff
changeset
|
276 |
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:
50227
diff
changeset
|
277 |
val equal_prems = filter op= prems_pairs |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
278 |
val _ = if null equal_prems then () |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
279 |
else error "The rule contains reflexive assumptions." |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
280 |
val concl_pairs = rule |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
281 |
|> concl_of |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
282 |
|> HOLogic.dest_Trueprop |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
283 |
|> dest_less_eq |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
284 |
|> pairself (snd o strip_comb) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
285 |
|> op~~ |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
286 |
|> filter_out op= |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
287 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
288 |
val _ = if has_duplicates op= concl_pairs |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
289 |
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:
50227
diff
changeset
|
290 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
291 |
fun rewrite_prem prem_pair = |
51427
08bb00239652
proper use of "member", without embarking on delicate questions about SML equality types;
wenzelm
parents:
51426
diff
changeset
|
292 |
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:
50227
diff
changeset
|
293 |
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:
51426
diff
changeset
|
294 |
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:
50227
diff
changeset
|
295 |
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:
50227
diff
changeset
|
296 |
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:
50227
diff
changeset
|
297 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
298 |
fun rewrite_prems [] = Conv.all_conv |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
299 |
| 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:
50227
diff
changeset
|
300 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
301 |
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:
50227
diff
changeset
|
302 |
val rewrite_concl_conv = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
303 |
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:
50227
diff
changeset
|
304 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
305 |
(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:
50227
diff
changeset
|
306 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
307 |
handle |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
308 |
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:
50227
diff
changeset
|
309 |
| 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:
50227
diff
changeset
|
310 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
311 |
fun negate_mono_rule mono_rule = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
312 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
313 |
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:
50227
diff
changeset
|
314 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
315 |
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:
50227
diff
changeset
|
316 |
end; |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
317 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
318 |
fun add_mono_rule mono_rule ctxt = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
319 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
320 |
val mono_rule = introduce_polarities mono_rule |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
321 |
val mono_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:
50227
diff
changeset
|
322 |
dest_Const o head_of o HOLogic.dest_Trueprop o concl_of) mono_rule |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
323 |
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:
50227
diff
changeset
|
324 |
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:
50227
diff
changeset
|
325 |
else () |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
326 |
val neg_mono_rule = negate_mono_rule mono_rule |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
327 |
val relator_distr_data = {pos_mono_rule = mono_rule, neg_mono_rule = neg_mono_rule, |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
328 |
pos_distr_rules = [], neg_distr_rules = []} |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
329 |
in |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
330 |
Data.map (map_relator_distr_data (Symtab.update (mono_ruleT_name, relator_distr_data))) ctxt |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
331 |
end; |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
332 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
333 |
local |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
334 |
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:
50227
diff
changeset
|
335 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
336 |
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:
50227
diff
changeset
|
337 |
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:
50227
diff
changeset
|
338 |
in |
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
339 |
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:
51994
diff
changeset
|
340 |
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:
51994
diff
changeset
|
341 |
ctxt |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
342 |
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:
50227
diff
changeset
|
343 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
344 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
345 |
fun rewrite_concl_conv thm ctm = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
346 |
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:
50227
diff
changeset
|
347 |
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:
50227
diff
changeset
|
348 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
349 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
350 |
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:
50227
diff
changeset
|
351 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
352 |
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:
50227
diff
changeset
|
353 |
fun update_entry distr_rule data = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
354 |
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:
50227
diff
changeset
|
355 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
356 |
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:
50227
diff
changeset
|
357 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
358 |
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:
50227
diff
changeset
|
359 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
360 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
361 |
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:
50227
diff
changeset
|
362 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
363 |
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:
50227
diff
changeset
|
364 |
fun update_entry distr_rule data = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
365 |
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:
50227
diff
changeset
|
366 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
367 |
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:
50227
diff
changeset
|
368 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
369 |
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:
50227
diff
changeset
|
370 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
371 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
372 |
local |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
373 |
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:
50227
diff
changeset
|
374 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
375 |
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:
50227
diff
changeset
|
376 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
377 |
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:
50227
diff
changeset
|
378 |
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:
50227
diff
changeset
|
379 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
380 |
ctxt |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
381 |
|> add_pos_distr_rule pos_distr_rule |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
382 |
|> add_neg_distr_rule neg_distr_rule |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
383 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
384 |
end; |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
385 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
386 |
local |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
387 |
fun sanity_check rule = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
388 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
389 |
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:
50227
diff
changeset
|
390 |
val concl = (perhaps (try HOLogic.dest_Trueprop)) (concl_of rule); |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
391 |
val (lhs, rhs) = case concl of |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
392 |
Const ("Orderings.ord_class.less_eq", _) $ (lhs as Const ("Relation.relcompp",_) $ _ $ _) $ rhs => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
393 |
(lhs, rhs) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
394 |
| Const ("Orderings.ord_class.less_eq", _) $ rhs $ (lhs as Const ("Relation.relcompp",_) $ _ $ _) => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
395 |
(lhs, rhs) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
396 |
| Const ("HOL.eq", _) $ (lhs as Const ("Relation.relcompp",_) $ _ $ _) $ rhs => (lhs, rhs) |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
397 |
| _ => error "The rule has a wrong format." |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
398 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
399 |
val lhs_vars = Term.add_vars lhs [] |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
400 |
val rhs_vars = Term.add_vars rhs [] |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
401 |
val assms_vars = fold Term.add_vars assms []; |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
402 |
val _ = if has_duplicates op= lhs_vars then error "Left-hand side has variable duplicates" else () |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
403 |
val _ = if subset op= (rhs_vars, lhs_vars) then () |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
404 |
else error "Extra variables in the right-hand side of the rule" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
405 |
val _ = if subset op= (assms_vars, lhs_vars) then () |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
406 |
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:
50227
diff
changeset
|
407 |
val rhs_args = (snd o strip_comb) rhs; |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
408 |
fun check_comp t = case t of |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
409 |
Const ("Relation.relcompp", _) $ Var (_, _) $ Var (_,_) => () |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
410 |
| _ => error "There is an argument on the rhs that is not a composition." |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
411 |
val _ = map check_comp rhs_args |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
412 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
413 |
() |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
414 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
415 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
416 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
417 |
fun add_distr_rule distr_rule ctxt = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
418 |
let |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
419 |
val _ = sanity_check distr_rule |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
420 |
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:
50227
diff
changeset
|
421 |
in |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
422 |
case concl of |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
423 |
Const ("Orderings.ord_class.less_eq", _) $ (Const ("Relation.relcompp",_) $ _ $ _) $ _ => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
424 |
add_pos_distr_rule distr_rule ctxt |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
425 |
| Const ("Orderings.ord_class.less_eq", _) $ _ $ (Const ("Relation.relcompp",_) $ _ $ _) => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
426 |
add_neg_distr_rule distr_rule ctxt |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
427 |
| Const ("HOL.eq", _) $ (Const ("Relation.relcompp",_) $ _ $ _) $ _ => |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
428 |
add_eq_distr_rule distr_rule ctxt |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
429 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
430 |
end |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
431 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
432 |
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:
50227
diff
changeset
|
433 |
(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:
51994
diff
changeset
|
434 |
(get_relator_distr_data' ctxt) [] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
435 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
436 |
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:
50227
diff
changeset
|
437 |
(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:
51994
diff
changeset
|
438 |
(get_relator_distr_data' ctxt) [] |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
439 |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
440 |
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:
50227
diff
changeset
|
441 |
|
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
442 |
val relator_distr_attribute_setup = |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
443 |
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:
50227
diff
changeset
|
444 |
"declaration of relator's monoticity" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
445 |
#> 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:
50227
diff
changeset
|
446 |
"declaration of relator's distributivity over OO" |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
447 |
#> Global_Theory.add_thms_dynamic |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
448 |
(@{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:
50227
diff
changeset
|
449 |
#> Global_Theory.add_thms_dynamic |
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
450 |
(@{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:
50227
diff
changeset
|
451 |
|
47308 | 452 |
(* theory setup *) |
453 |
||
454 |
val setup = |
|
47951
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
455 |
quot_map_attribute_setup |
8c8a03765de7
quot_del attribute, it allows us to deregister quotient types
kuncar
parents:
47936
diff
changeset
|
456 |
#> quot_del_attribute_setup |
47634
091bcd569441
hide the invariant constant for relators: invariant_commute infrastracture
kuncar
parents:
47308
diff
changeset
|
457 |
#> Invariant_Commute.setup |
51994 | 458 |
#> relfexivity_rule_setup |
51374
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
kuncar
parents:
50227
diff
changeset
|
459 |
#> relator_distr_attribute_setup |
47308 | 460 |
|
461 |
(* outer syntax commands *) |
|
462 |
||
463 |
val _ = |
|
53219
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
kuncar
parents:
51994
diff
changeset
|
464 |
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:
51994
diff
changeset
|
465 |
(Scan.succeed (Toplevel.keep (print_quot_maps o Toplevel.context_of))) |
47308 | 466 |
|
467 |
val _ = |
|
468 |
Outer_Syntax.improper_command @{command_spec "print_quotients"} "print quotients" |
|
469 |
(Scan.succeed (Toplevel.keep (print_quotients o Toplevel.context_of))) |
|
470 |
||
51426 | 471 |
end; |