| author | wenzelm |
| Sat, 04 Oct 2014 18:16:24 +0200 | |
| changeset 58536 | 402a8e8107a7 |
| parent 57960 | ee1ba4848896 |
| child 58893 | 9e0ecb66d6a7 |
| 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 |
| 45279 | 12 |
val print_quotmaps: Proof.context -> unit |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
13 |
|
|
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
|
14 |
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
|
15 |
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
|
16 |
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
|
17 |
val lookup_abs_rep_global: theory -> 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 update_abs_rep: string -> abs_rep -> Context.generic -> Context.generic |
|
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
|
19 |
val print_abs_rep: Proof.context -> unit |
|
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 |
|
| 47093 | 21 |
type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm, quot_thm: thm}
|
| 45279 | 22 |
val transform_quotients: morphism -> quotients -> quotients |
| 45280 | 23 |
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
|
24 |
val lookup_quotients_global: theory -> string -> quotients option |
| 45279 | 25 |
val update_quotients: string -> quotients -> Context.generic -> Context.generic |
26 |
val dest_quotients: Proof.context -> quotients list |
|
27 |
val print_quotients: Proof.context -> unit |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
28 |
|
| 45279 | 29 |
type quotconsts = {qconst: term, rconst: term, def: thm}
|
30 |
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
|
31 |
val lookup_quotconsts_global: theory -> term -> quotconsts option |
| 45279 | 32 |
val update_quotconsts: string -> quotconsts -> Context.generic -> Context.generic |
|
45350
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents:
45340
diff
changeset
|
33 |
val dest_quotconsts_global: theory -> quotconsts list |
| 45279 | 34 |
val dest_quotconsts: Proof.context -> quotconsts list |
35 |
val print_quotconsts: Proof.context -> unit |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
36 |
end; |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
37 |
|
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
38 |
structure Quotient_Info: QUOTIENT_INFO = |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
struct |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
40 |
|
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
41 |
(** data containers **) |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
42 |
|
|
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41452
diff
changeset
|
43 |
(* FIXME just one data slot (record) per program unit *) |
|
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
41452
diff
changeset
|
44 |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
45 |
(* info about map- and rel-functions for a type *) |
| 47094 | 46 |
type quotmaps = {relmap: string, quot_thm: thm}
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
47 |
|
| 45280 | 48 |
structure Quotmaps = Generic_Data |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
49 |
( |
| 45279 | 50 |
type T = quotmaps Symtab.table |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
51 |
val empty = Symtab.empty |
|
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
52 |
val extend = I |
|
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
53 |
fun merge data = Symtab.merge (K true) data |
|
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
54 |
) |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
55 |
|
| 45280 | 56 |
val lookup_quotmaps = Symtab.lookup o Quotmaps.get o Context.Proof |
|
45340
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents:
45294
diff
changeset
|
57 |
val lookup_quotmaps_global = Symtab.lookup o Quotmaps.get o Context.Theory |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
58 |
|
| 45281 | 59 |
(* FIXME export proper internal update operation!? *) |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
60 |
|
| 57960 | 61 |
val _ = |
62 |
Theory.setup |
|
63 |
(Attrib.setup @{binding mapQ3}
|
|
64 |
((Args.type_name {proper = true, strict = true} --| Scan.lift @{keyword "="}) --
|
|
65 |
(Scan.lift @{keyword "("} |--
|
|
66 |
Args.const {proper = true, strict = true} --| Scan.lift @{keyword ","} --
|
|
67 |
Attrib.thm --| Scan.lift @{keyword ")"}) >>
|
|
68 |
(fn (tyname, (relname, qthm)) => |
|
69 |
let val minfo = {relmap = relname, quot_thm = qthm}
|
|
70 |
in Thm.declaration_attribute (fn _ => Quotmaps.map (Symtab.update (tyname, minfo))) end)) |
|
71 |
"declaration of map information") |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
72 |
|
| 45279 | 73 |
fun print_quotmaps ctxt = |
| 41444 | 74 |
let |
| 47094 | 75 |
fun prt_map (ty_name, {relmap, quot_thm}) =
|
| 41445 | 76 |
Pretty.block (separate (Pretty.brk 2) |
| 47094 | 77 |
[Pretty.str "type:", |
78 |
Pretty.str ty_name, |
|
79 |
Pretty.str "relation map:", |
|
80 |
Pretty.str relmap, |
|
81 |
Pretty.str "quot. theorem:", |
|
82 |
Syntax.pretty_term ctxt (prop_of quot_thm)]) |
|
| 41444 | 83 |
in |
| 45280 | 84 |
map prt_map (Symtab.dest (Quotmaps.get (Context.Proof ctxt))) |
| 41444 | 85 |
|> Pretty.big_list "maps for type constructors:" |
86 |
|> Pretty.writeln |
|
87 |
end |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
88 |
|
|
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
|
89 |
(* info about 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
|
90 |
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
|
91 |
|
|
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
|
92 |
structure Abs_Rep = Generic_Data |
|
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
|
93 |
( |
|
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
|
94 |
type T = abs_rep Symtab.table |
|
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
|
95 |
val empty = Symtab.empty |
|
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
|
96 |
val extend = I |
|
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
|
97 |
fun merge data = Symtab.merge (K true) data |
|
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
|
98 |
) |
|
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
|
99 |
|
|
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
|
100 |
fun transform_abs_rep phi {abs, rep} = {abs = Morphism.term phi abs, rep = Morphism.term phi 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
|
101 |
|
|
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
|
102 |
val lookup_abs_rep = Symtab.lookup o Abs_Rep.get o Context.Proof |
|
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
|
103 |
val lookup_abs_rep_global = Symtab.lookup o Abs_Rep.get o Context.Theory |
|
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
|
104 |
|
|
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
|
105 |
fun update_abs_rep str data = Abs_Rep.map (Symtab.update (str, data)) |
|
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
|
106 |
|
|
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
|
107 |
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
|
108 |
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
|
109 |
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
|
110 |
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
|
111 |
[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
|
112 |
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
|
113 |
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
|
114 |
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
|
115 |
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
|
116 |
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
|
117 |
in |
|
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
|
118 |
map prt_abs_rep (Symtab.dest (Abs_Rep.get (Context.Proof 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
|
119 |
|> 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
|
120 |
|> 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
|
121 |
end |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
122 |
|
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
123 |
(* info about quotient types *) |
| 47093 | 124 |
type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm, quot_thm: thm}
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
125 |
|
| 45280 | 126 |
structure Quotients = Generic_Data |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
127 |
( |
| 45279 | 128 |
type T = quotients Symtab.table |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
129 |
val empty = Symtab.empty |
|
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
130 |
val extend = I |
|
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
131 |
fun merge data = Symtab.merge (K true) data |
|
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
132 |
) |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
133 |
|
| 47093 | 134 |
fun transform_quotients phi {qtyp, rtyp, equiv_rel, equiv_thm, quot_thm} =
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
135 |
{qtyp = Morphism.typ phi qtyp,
|
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
136 |
rtyp = Morphism.typ phi rtyp, |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
137 |
equiv_rel = Morphism.term phi equiv_rel, |
| 47093 | 138 |
equiv_thm = Morphism.thm phi equiv_thm, |
139 |
quot_thm = Morphism.thm phi quot_thm} |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
140 |
|
| 45280 | 141 |
val lookup_quotients = Symtab.lookup o Quotients.get o Context.Proof |
|
45340
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents:
45294
diff
changeset
|
142 |
val lookup_quotients_global = Symtab.lookup o Quotients.get o Context.Theory |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
143 |
|
| 45280 | 144 |
fun update_quotients str qinfo = Quotients.map (Symtab.update (str, qinfo)) |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
145 |
|
| 45280 | 146 |
fun dest_quotients ctxt = (* FIXME slightly expensive way to retrieve data *) |
147 |
map snd (Symtab.dest (Quotients.get (Context.Proof ctxt))) |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
148 |
|
| 45279 | 149 |
fun print_quotients ctxt = |
| 41444 | 150 |
let |
| 47093 | 151 |
fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm, quot_thm} =
|
| 41445 | 152 |
Pretty.block (separate (Pretty.brk 2) |
| 41444 | 153 |
[Pretty.str "quotient type:", |
154 |
Syntax.pretty_typ ctxt qtyp, |
|
155 |
Pretty.str "raw type:", |
|
156 |
Syntax.pretty_typ ctxt rtyp, |
|
157 |
Pretty.str "relation:", |
|
158 |
Syntax.pretty_term ctxt equiv_rel, |
|
159 |
Pretty.str "equiv. thm:", |
|
| 47093 | 160 |
Syntax.pretty_term ctxt (prop_of equiv_thm), |
161 |
Pretty.str "quot. thm:", |
|
162 |
Syntax.pretty_term ctxt (prop_of quot_thm)]) |
|
| 41444 | 163 |
in |
| 45280 | 164 |
map (prt_quot o snd) (Symtab.dest (Quotients.get (Context.Proof ctxt))) |
| 41444 | 165 |
|> Pretty.big_list "quotients:" |
166 |
|> Pretty.writeln |
|
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 |
|
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
170 |
(* info about quotient constants *) |
| 45279 | 171 |
type quotconsts = {qconst: term, rconst: term, def: thm}
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
172 |
|
| 45279 | 173 |
fun eq_quotconsts (x : quotconsts, y : quotconsts) = #qconst x = #qconst y |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
174 |
|
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
175 |
(* We need to be able to lookup instances of lifted constants, |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
176 |
for example given "nat fset" we need to find "'a fset"; |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
177 |
but overloaded constants share the same name *) |
| 45280 | 178 |
structure Quotconsts = Generic_Data |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
179 |
( |
| 45279 | 180 |
type T = quotconsts list Symtab.table |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
181 |
val empty = Symtab.empty |
|
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
182 |
val extend = I |
| 45279 | 183 |
val merge = Symtab.merge_list eq_quotconsts |
|
38759
37a9092de102
simplification/standardization of some theory data;
wenzelm
parents:
38756
diff
changeset
|
184 |
) |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
185 |
|
| 45279 | 186 |
fun transform_quotconsts phi {qconst, rconst, def} =
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
187 |
{qconst = Morphism.term phi qconst,
|
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
188 |
rconst = Morphism.term phi rconst, |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
189 |
def = Morphism.thm phi def} |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
190 |
|
| 45280 | 191 |
fun update_quotconsts name qcinfo = Quotconsts.map (Symtab.cons_list (name, qcinfo)) |
192 |
||
193 |
fun dest_quotconsts ctxt = |
|
194 |
flat (map snd (Symtab.dest (Quotconsts.get (Context.Proof ctxt)))) |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
195 |
|
|
45350
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents:
45340
diff
changeset
|
196 |
fun dest_quotconsts_global thy = |
|
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents:
45340
diff
changeset
|
197 |
flat (map snd (Symtab.dest (Quotconsts.get (Context.Theory thy)))) |
|
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents:
45340
diff
changeset
|
198 |
|
|
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents:
45340
diff
changeset
|
199 |
|
|
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents:
45340
diff
changeset
|
200 |
|
|
45340
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents:
45294
diff
changeset
|
201 |
fun lookup_quotconsts_global thy t = |
| 45280 | 202 |
let |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
203 |
val (name, qty) = dest_Const t |
| 45279 | 204 |
fun matches (x: quotconsts) = |
| 45280 | 205 |
let val (name', qty') = dest_Const (#qconst x); |
206 |
in name = name' andalso Sign.typ_instance thy (qty, qty') end |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
207 |
in |
|
45340
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
wenzelm
parents:
45294
diff
changeset
|
208 |
(case Symtab.lookup (Quotconsts.get (Context.Theory thy)) name of |
|
45274
252cd58847e0
respecting isabelle's programming style in the quotient package by simplifying qconsts_lookup function for data access; removing odd NotFound exception
bulwahn
parents:
45273
diff
changeset
|
209 |
NONE => NONE |
|
252cd58847e0
respecting isabelle's programming style in the quotient package by simplifying qconsts_lookup function for data access; removing odd NotFound exception
bulwahn
parents:
45273
diff
changeset
|
210 |
| SOME l => find_first matches l) |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
211 |
end |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
212 |
|
| 45279 | 213 |
fun print_quotconsts ctxt = |
| 41444 | 214 |
let |
215 |
fun prt_qconst {qconst, rconst, def} =
|
|
216 |
Pretty.block (separate (Pretty.brk 1) |
|
217 |
[Syntax.pretty_term ctxt qconst, |
|
218 |
Pretty.str ":=", |
|
219 |
Syntax.pretty_term ctxt rconst, |
|
220 |
Pretty.str "as", |
|
221 |
Syntax.pretty_term ctxt (prop_of def)]) |
|
222 |
in |
|
| 45280 | 223 |
map prt_qconst (maps snd (Symtab.dest (Quotconsts.get (Context.Proof ctxt)))) |
| 41444 | 224 |
|> Pretty.big_list "quotient constants:" |
225 |
|> Pretty.writeln |
|
226 |
end |
|
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
227 |
|
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
228 |
|
| 41452 | 229 |
(* outer syntax commands *) |
230 |
||
231 |
val _ = |
|
| 47308 | 232 |
Outer_Syntax.improper_command @{command_spec "print_quotmapsQ3"} "print quotient map functions"
|
| 45279 | 233 |
(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
|
234 |
|
| 41452 | 235 |
val _ = |
| 47308 | 236 |
Outer_Syntax.improper_command @{command_spec "print_quotientsQ3"} "print quotients"
|
| 45279 | 237 |
(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
|
238 |
|
| 41452 | 239 |
val _ = |
|
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
240 |
Outer_Syntax.improper_command @{command_spec "print_quotconsts"} "print quotient constants"
|
| 45279 | 241 |
(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
|
242 |
|
| 45279 | 243 |
end; |