author | wenzelm |
Sun, 07 Nov 2021 19:53:37 +0100 | |
changeset 74726 | 33ed2eb06d68 |
parent 74561 | 8e6c973003c8 |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOL/Tools/Quotient/quotient_info.ML |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
2 |
Author: Cezary Kaliszyk and Christian Urban |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
|
45279 | 4 |
Context data for the quotient package. |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
*) |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
signature QUOTIENT_INFO = |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
8 |
sig |
47094 | 9 |
type quotmaps = {relmap: string, quot_thm: thm} |
45280 | 10 |
val lookup_quotmaps: Proof.context -> string -> quotmaps option |
45340
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents:
45294
diff
changeset
|
11 |
val lookup_quotmaps_global: theory -> string -> quotmaps option |
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
12 |
val update_quotmaps: string * quotmaps -> Context.generic -> Context.generic |
45279 | 13 |
val print_quotmaps: Proof.context -> unit |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
14 |
|
45534
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
15 |
type abs_rep = {abs : term, rep : term} |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
16 |
val transform_abs_rep: morphism -> abs_rep -> abs_rep |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
17 |
val lookup_abs_rep: Proof.context -> string -> abs_rep option |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
18 |
val lookup_abs_rep_global: theory -> string -> abs_rep option |
59157 | 19 |
val update_abs_rep: string * abs_rep -> Context.generic -> Context.generic |
45534
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
20 |
val print_abs_rep: Proof.context -> unit |
59157 | 21 |
|
47093 | 22 |
type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm, quot_thm: thm} |
45279 | 23 |
val transform_quotients: morphism -> quotients -> quotients |
45280 | 24 |
val lookup_quotients: Proof.context -> string -> quotients option |
45340
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents:
45294
diff
changeset
|
25 |
val lookup_quotients_global: theory -> string -> quotients option |
59157 | 26 |
val update_quotients: string * quotients -> Context.generic -> Context.generic |
45279 | 27 |
val dest_quotients: Proof.context -> quotients list |
28 |
val print_quotients: Proof.context -> unit |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
29 |
|
45279 | 30 |
type quotconsts = {qconst: term, rconst: term, def: thm} |
31 |
val transform_quotconsts: morphism -> quotconsts -> quotconsts |
|
45340
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents:
45294
diff
changeset
|
32 |
val lookup_quotconsts_global: theory -> term -> quotconsts option |
59157 | 33 |
val update_quotconsts: string * quotconsts -> Context.generic -> Context.generic |
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
34 |
val dest_quotconsts: Proof.context -> quotconsts list |
45350
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents:
45340
diff
changeset
|
35 |
val dest_quotconsts_global: theory -> quotconsts list |
45279 | 36 |
val print_quotconsts: Proof.context -> unit |
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
37 |
end |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
38 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
structure Quotient_Info: QUOTIENT_INFO = |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
40 |
struct |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
41 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
42 |
(** data containers **) |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
43 |
|
59157 | 44 |
(*info about map- and rel-functions for a type*) |
47094 | 45 |
type quotmaps = {relmap: string, quot_thm: thm} |
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
46 |
fun transform_quotmaps phi : quotmaps -> quotmaps = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
47 |
fn {relmap, quot_thm} => {relmap = relmap, quot_thm = Morphism.thm phi quot_thm} |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
48 |
|
59157 | 49 |
(*info about abs/rep terms*) |
50 |
type abs_rep = {abs : term, rep : term} |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
51 |
fun transform_abs_rep phi : abs_rep -> abs_rep = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
52 |
fn {abs, rep} => {abs = Morphism.term phi abs, rep = Morphism.term phi rep} |
59157 | 53 |
|
54 |
(*info about quotient types*) |
|
55 |
type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm, quot_thm: thm} |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
56 |
fun transform_quotients phi : quotients -> quotients = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
57 |
fn {qtyp, rtyp, equiv_rel, equiv_thm, quot_thm} => |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
58 |
{qtyp = Morphism.typ phi qtyp, |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
59 |
rtyp = Morphism.typ phi rtyp, |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
60 |
equiv_rel = Morphism.term phi equiv_rel, |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
61 |
equiv_thm = Morphism.thm phi equiv_thm, |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
62 |
quot_thm = Morphism.thm phi quot_thm} |
59157 | 63 |
|
64 |
(*info about quotient constants*) |
|
65 |
(*We need to be able to lookup instances of lifted constants, |
|
66 |
for example given "nat fset" we need to find "'a fset"; |
|
67 |
but overloaded constants share the same name.*) |
|
68 |
type quotconsts = {qconst: term, rconst: term, def: thm} |
|
69 |
fun eq_quotconsts (x: quotconsts, y: quotconsts) = #qconst x = #qconst y |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
70 |
fun transform_quotconsts phi : quotconsts -> quotconsts = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
71 |
fn {qconst, rconst, def} => |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
72 |
{qconst = Morphism.term phi qconst, |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
73 |
rconst = Morphism.term phi rconst, |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
74 |
def = Morphism.thm phi def} |
59157 | 75 |
|
76 |
structure Data = Generic_Data |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
77 |
( |
59157 | 78 |
type T = |
79 |
quotmaps Symtab.table * |
|
80 |
abs_rep Symtab.table * |
|
81 |
quotients Symtab.table * |
|
82 |
quotconsts list Symtab.table |
|
83 |
val empty: T = (Symtab.empty, Symtab.empty, Symtab.empty, Symtab.empty) |
|
84 |
fun merge |
|
85 |
((quotmaps1, abs_rep1, quotients1, quotconsts1), |
|
86 |
(quotmaps2, abs_rep2, quotients2, quotconsts2)) : T = |
|
87 |
(Symtab.merge (K true) (quotmaps1, quotmaps2), |
|
88 |
Symtab.merge (K true) (abs_rep1, abs_rep2), |
|
89 |
Symtab.merge (K true) (quotients1, quotients2), |
|
90 |
Symtab.merge_list eq_quotconsts (quotconsts1, quotconsts2)) |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
91 |
) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
92 |
|
59157 | 93 |
val get_quotmaps = #1 o Data.get |
94 |
val get_abs_rep = #2 o Data.get |
|
95 |
val get_quotients = #3 o Data.get |
|
96 |
val get_quotconsts = #4 o Data.get |
|
97 |
||
98 |
val map_quotmaps = Data.map o @{apply 4(1)} |
|
99 |
val map_abs_rep = Data.map o @{apply 4(2)} |
|
100 |
val map_quotients = Data.map o @{apply 4(3)} |
|
101 |
val map_quotconsts = Data.map o @{apply 4(4)} |
|
102 |
||
103 |
||
104 |
(* quotmaps *) |
|
105 |
||
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
106 |
fun lookup_quotmaps_generic context name = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
107 |
Symtab.lookup (get_quotmaps context) name |
67664 | 108 |
|> Option.map (transform_quotmaps (Morphism.transfer_morphism'' context)) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
109 |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
110 |
val lookup_quotmaps = lookup_quotmaps_generic o Context.Proof |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
111 |
val lookup_quotmaps_global = lookup_quotmaps_generic o Context.Theory |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
112 |
|
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
113 |
val update_quotmaps = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
114 |
map_quotmaps o Symtab.update o apsnd (transform_quotmaps Morphism.trim_context_morphism) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
115 |
|
57960 | 116 |
val _ = |
117 |
Theory.setup |
|
69593 | 118 |
(Attrib.setup \<^binding>\<open>mapQ3\<close> |
119 |
((Args.type_name {proper = true, strict = true} --| Scan.lift \<^keyword>\<open>=\<close>) -- |
|
120 |
(Scan.lift \<^keyword>\<open>(\<close> |-- |
|
121 |
Args.const {proper = true, strict = true} --| Scan.lift \<^keyword>\<open>,\<close> -- |
|
122 |
Attrib.thm --| Scan.lift \<^keyword>\<open>)\<close>) >> |
|
67632 | 123 |
(fn (tyname, (relmap, quot_thm)) => |
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
124 |
Thm.declaration_attribute |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
125 |
(K (update_quotmaps (tyname, {relmap = relmap, quot_thm = quot_thm}))))) |
57960 | 126 |
"declaration of map information") |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
127 |
|
45279 | 128 |
fun print_quotmaps ctxt = |
41444 | 129 |
let |
47094 | 130 |
fun prt_map (ty_name, {relmap, quot_thm}) = |
41445 | 131 |
Pretty.block (separate (Pretty.brk 2) |
59157 | 132 |
[Pretty.str "type:", |
47094 | 133 |
Pretty.str ty_name, |
59157 | 134 |
Pretty.str "relation map:", |
47094 | 135 |
Pretty.str relmap, |
59157 | 136 |
Pretty.str "quot. theorem:", |
59582 | 137 |
Syntax.pretty_term ctxt (Thm.prop_of quot_thm)]) |
41444 | 138 |
in |
59157 | 139 |
map prt_map (Symtab.dest (get_quotmaps (Context.Proof ctxt))) |
41444 | 140 |
|> Pretty.big_list "maps for type constructors:" |
141 |
|> Pretty.writeln |
|
142 |
end |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
143 |
|
59157 | 144 |
|
145 |
(* abs_rep *) |
|
45534
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
146 |
|
59157 | 147 |
val lookup_abs_rep = Symtab.lookup o get_abs_rep o Context.Proof |
148 |
val lookup_abs_rep_global = Symtab.lookup o get_abs_rep o Context.Theory |
|
45534
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
149 |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
150 |
val update_abs_rep = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
151 |
map_abs_rep o Symtab.update o apsnd (transform_abs_rep Morphism.trim_context_morphism) |
45534
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
152 |
|
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
153 |
fun print_abs_rep ctxt = |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
154 |
let |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
155 |
fun prt_abs_rep (s, {abs, rep}) = |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
156 |
Pretty.block (separate (Pretty.brk 2) |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
157 |
[Pretty.str "type constructor:", |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
158 |
Pretty.str s, |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
159 |
Pretty.str "abs term:", |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
160 |
Syntax.pretty_term ctxt abs, |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
161 |
Pretty.str "rep term:", |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
162 |
Syntax.pretty_term ctxt rep]) |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
163 |
in |
59157 | 164 |
map prt_abs_rep (Symtab.dest (get_abs_rep (Context.Proof ctxt))) |
45534
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
165 |
|> Pretty.big_list "abs/rep terms:" |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
166 |
|> Pretty.writeln |
4ab21521b393
adding database of abs and rep terms to the quotient package; registering abs and rep terms in quotient_type and using them in quotient_definition
bulwahn
parents:
45350
diff
changeset
|
167 |
end |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
168 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
169 |
|
59157 | 170 |
(* quotients *) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
171 |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
172 |
fun lookup_quotients_generic context name = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
173 |
Symtab.lookup (get_quotients context) name |
67664 | 174 |
|> Option.map (transform_quotients (Morphism.transfer_morphism'' context)) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
175 |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
176 |
val lookup_quotients = lookup_quotients_generic o Context.Proof |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
177 |
val lookup_quotients_global = lookup_quotients_generic o Context.Theory |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
178 |
|
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
179 |
val update_quotients = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
180 |
map_quotients o Symtab.update o apsnd (transform_quotients Morphism.trim_context_morphism) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
181 |
|
59157 | 182 |
fun dest_quotients ctxt = |
183 |
map snd (Symtab.dest (get_quotients (Context.Proof ctxt))) |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
184 |
|
45279 | 185 |
fun print_quotients ctxt = |
41444 | 186 |
let |
47093 | 187 |
fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm, quot_thm} = |
41445 | 188 |
Pretty.block (separate (Pretty.brk 2) |
41444 | 189 |
[Pretty.str "quotient type:", |
190 |
Syntax.pretty_typ ctxt qtyp, |
|
191 |
Pretty.str "raw type:", |
|
192 |
Syntax.pretty_typ ctxt rtyp, |
|
193 |
Pretty.str "relation:", |
|
194 |
Syntax.pretty_term ctxt equiv_rel, |
|
195 |
Pretty.str "equiv. thm:", |
|
59582 | 196 |
Syntax.pretty_term ctxt (Thm.prop_of equiv_thm), |
47093 | 197 |
Pretty.str "quot. thm:", |
59582 | 198 |
Syntax.pretty_term ctxt (Thm.prop_of quot_thm)]) |
41444 | 199 |
in |
59157 | 200 |
map (prt_quot o snd) (Symtab.dest (get_quotients (Context.Proof ctxt))) |
41444 | 201 |
|> Pretty.big_list "quotients:" |
202 |
|> Pretty.writeln |
|
203 |
end |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
204 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
205 |
|
59157 | 206 |
(* quotconsts *) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
207 |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
208 |
val update_quotconsts = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
209 |
map_quotconsts o Symtab.cons_list o apsnd (transform_quotconsts Morphism.trim_context_morphism) |
45280 | 210 |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
211 |
fun dest_quotconsts_generic context = |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
212 |
maps #2 (Symtab.dest (get_quotconsts context)) |
67664 | 213 |
|> map (transform_quotconsts (Morphism.transfer_morphism'' context)) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
214 |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
215 |
val dest_quotconsts = dest_quotconsts_generic o Context.Proof |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
216 |
val dest_quotconsts_global = dest_quotconsts_generic o Context.Theory |
45350
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents:
45340
diff
changeset
|
217 |
|
45340
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents:
45294
diff
changeset
|
218 |
fun lookup_quotconsts_global thy t = |
45280 | 219 |
let |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
220 |
val (name, qty) = dest_Const t |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
221 |
in |
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
222 |
Symtab.lookup_list (get_quotconsts (Context.Theory thy)) name |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
223 |
|> find_first (fn {qconst, ...} => |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
224 |
let val (name', qty') = dest_Const qconst |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
225 |
in name = name' andalso Sign.typ_instance thy (qty, qty') end) |
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
226 |
|> Option.map (transform_quotconsts (Morphism.transfer_morphism thy)) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
227 |
end |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
228 |
|
45279 | 229 |
fun print_quotconsts ctxt = |
41444 | 230 |
let |
231 |
fun prt_qconst {qconst, rconst, def} = |
|
232 |
Pretty.block (separate (Pretty.brk 1) |
|
233 |
[Syntax.pretty_term ctxt qconst, |
|
234 |
Pretty.str ":=", |
|
235 |
Syntax.pretty_term ctxt rconst, |
|
236 |
Pretty.str "as", |
|
59582 | 237 |
Syntax.pretty_term ctxt (Thm.prop_of def)]) |
41444 | 238 |
in |
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
239 |
map prt_qconst (dest_quotconsts ctxt) |
41444 | 240 |
|> Pretty.big_list "quotient constants:" |
241 |
|> Pretty.writeln |
|
242 |
end |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
243 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
244 |
|
41452 | 245 |
(* outer syntax commands *) |
246 |
||
247 |
val _ = |
|
69593 | 248 |
Outer_Syntax.command \<^command_keyword>\<open>print_quotmapsQ3\<close> "print quotient map functions" |
45279 | 249 |
(Scan.succeed (Toplevel.keep (print_quotmaps o Toplevel.context_of))) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
250 |
|
41452 | 251 |
val _ = |
69593 | 252 |
Outer_Syntax.command \<^command_keyword>\<open>print_quotientsQ3\<close> "print quotients" |
45279 | 253 |
(Scan.succeed (Toplevel.keep (print_quotients o Toplevel.context_of))) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
254 |
|
41452 | 255 |
val _ = |
69593 | 256 |
Outer_Syntax.command \<^command_keyword>\<open>print_quotconsts\<close> "print quotient constants" |
45279 | 257 |
(Scan.succeed (Toplevel.keep (print_quotconsts o Toplevel.context_of))) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
258 |
|
67633
9a1212f4393e
clarified data operations, with trim_context and transfer;
wenzelm
parents:
67632
diff
changeset
|
259 |
end |