| author | wenzelm | 
| Sat, 31 Aug 2013 22:32:43 +0200 | |
| changeset 53349 | ae8c9380bbc4 | 
| parent 47308 | 9caab698dbe4 | 
| child 55951 | c07d184aebe9 | 
| 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: 
45294diff
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: 
45350diff
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: 
45350diff
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: 
45350diff
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: 
45350diff
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: 
45350diff
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: 
45350diff
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: 
45350diff
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: 
45294diff
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: 
45294diff
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: 
45340diff
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 | |
| 45279 | 37 | val equiv_rules: Proof.context -> thm list | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 38 | val equiv_rules_add: attribute | 
| 45279 | 39 | val rsp_rules: Proof.context -> thm list | 
| 35314 
cbdf785a1eb3
export prs_rules and rsp_rules attributes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
35278diff
changeset | 40 | val rsp_rules_add: attribute | 
| 45279 | 41 | val prs_rules: Proof.context -> thm list | 
| 35314 
cbdf785a1eb3
export prs_rules and rsp_rules attributes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
35278diff
changeset | 42 | val prs_rules_add: attribute | 
| 45279 | 43 | val id_simps: Proof.context -> thm list | 
| 44 | val quotient_rules: Proof.context -> thm list | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 45 | val quotient_rules_add: attribute | 
| 41452 | 46 | val setup: theory -> theory | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 47 | end; | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 48 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 49 | structure Quotient_Info: QUOTIENT_INFO = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 50 | struct | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 51 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 52 | (** data containers **) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 53 | |
| 41472 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 wenzelm parents: 
41452diff
changeset | 54 | (* 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: 
41452diff
changeset | 55 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 56 | (* info about map- and rel-functions for a type *) | 
| 47094 | 57 | type quotmaps = {relmap: string, quot_thm: thm}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 58 | |
| 45280 | 59 | structure Quotmaps = Generic_Data | 
| 38759 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 60 | ( | 
| 45279 | 61 | type T = quotmaps Symtab.table | 
| 38759 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 62 | val empty = Symtab.empty | 
| 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 63 | val extend = I | 
| 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 64 | fun merge data = Symtab.merge (K true) data | 
| 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 65 | ) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 66 | |
| 45280 | 67 | 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: 
45294diff
changeset | 68 | 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 | 69 | |
| 45281 | 70 | (* FIXME export proper internal update operation!? *) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 71 | |
| 45281 | 72 | val quotmaps_attribute_setup = | 
| 47308 | 73 |   Attrib.setup @{binding mapQ3}
 | 
| 47110 | 74 |     ((Args.type_name true --| Scan.lift @{keyword "="}) --
 | 
| 75 |       (Scan.lift @{keyword "("} |-- Args.const_proper true --| Scan.lift @{keyword ","} --
 | |
| 76 |         Attrib.thm --| Scan.lift @{keyword ")"}) >>
 | |
| 47094 | 77 | (fn (tyname, (relname, qthm)) => | 
| 78 |         let val minfo = {relmap = relname, quot_thm = qthm}
 | |
| 45281 | 79 | in Thm.declaration_attribute (fn _ => Quotmaps.map (Symtab.update (tyname, minfo))) end)) | 
| 80 | "declaration of map information" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 81 | |
| 45279 | 82 | fun print_quotmaps ctxt = | 
| 41444 | 83 | let | 
| 47094 | 84 |     fun prt_map (ty_name, {relmap, quot_thm}) =
 | 
| 41445 | 85 | Pretty.block (separate (Pretty.brk 2) | 
| 47094 | 86 | [Pretty.str "type:", | 
| 87 | Pretty.str ty_name, | |
| 88 | Pretty.str "relation map:", | |
| 89 | Pretty.str relmap, | |
| 90 | Pretty.str "quot. theorem:", | |
| 91 | Syntax.pretty_term ctxt (prop_of quot_thm)]) | |
| 41444 | 92 | in | 
| 45280 | 93 | map prt_map (Symtab.dest (Quotmaps.get (Context.Proof ctxt))) | 
| 41444 | 94 | |> Pretty.big_list "maps for type constructors:" | 
| 95 | |> Pretty.writeln | |
| 96 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 97 | |
| 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: 
45350diff
changeset | 98 | (* 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: 
45350diff
changeset | 99 | 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: 
45350diff
changeset | 100 | |
| 
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: 
45350diff
changeset | 101 | 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: 
45350diff
changeset | 102 | ( | 
| 
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: 
45350diff
changeset | 103 | 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: 
45350diff
changeset | 104 | 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: 
45350diff
changeset | 105 | 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: 
45350diff
changeset | 106 | 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: 
45350diff
changeset | 107 | ) | 
| 
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: 
45350diff
changeset | 108 | |
| 
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: 
45350diff
changeset | 109 | 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: 
45350diff
changeset | 110 | |
| 
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: 
45350diff
changeset | 111 | 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: 
45350diff
changeset | 112 | 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: 
45350diff
changeset | 113 | |
| 
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: 
45350diff
changeset | 114 | 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: 
45350diff
changeset | 115 | |
| 
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: 
45350diff
changeset | 116 | 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: 
45350diff
changeset | 117 | 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: 
45350diff
changeset | 118 |     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: 
45350diff
changeset | 119 | 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: 
45350diff
changeset | 120 | [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: 
45350diff
changeset | 121 | 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: 
45350diff
changeset | 122 | 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: 
45350diff
changeset | 123 | 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: 
45350diff
changeset | 124 | 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: 
45350diff
changeset | 125 | 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: 
45350diff
changeset | 126 | 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: 
45350diff
changeset | 127 | 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: 
45350diff
changeset | 128 | |> 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: 
45350diff
changeset | 129 | |> 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: 
45350diff
changeset | 130 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 131 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 132 | (* info about quotient types *) | 
| 47093 | 133 | 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 | 134 | |
| 45280 | 135 | structure Quotients = Generic_Data | 
| 38759 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 136 | ( | 
| 45279 | 137 | type T = quotients Symtab.table | 
| 38759 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 138 | val empty = Symtab.empty | 
| 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 139 | val extend = I | 
| 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 140 | fun merge data = Symtab.merge (K true) data | 
| 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 141 | ) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 142 | |
| 47093 | 143 | 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 | 144 |   {qtyp = Morphism.typ phi qtyp,
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 145 | rtyp = Morphism.typ phi rtyp, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 146 | equiv_rel = Morphism.term phi equiv_rel, | 
| 47093 | 147 | equiv_thm = Morphism.thm phi equiv_thm, | 
| 148 | quot_thm = Morphism.thm phi quot_thm} | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 149 | |
| 45280 | 150 | 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: 
45294diff
changeset | 151 | 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 | 152 | |
| 45280 | 153 | 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 | 154 | |
| 45280 | 155 | fun dest_quotients ctxt = (* FIXME slightly expensive way to retrieve data *) | 
| 156 | 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 | 157 | |
| 45279 | 158 | fun print_quotients ctxt = | 
| 41444 | 159 | let | 
| 47093 | 160 |     fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm, quot_thm} =
 | 
| 41445 | 161 | Pretty.block (separate (Pretty.brk 2) | 
| 41444 | 162 | [Pretty.str "quotient type:", | 
| 163 | Syntax.pretty_typ ctxt qtyp, | |
| 164 | Pretty.str "raw type:", | |
| 165 | Syntax.pretty_typ ctxt rtyp, | |
| 166 | Pretty.str "relation:", | |
| 167 | Syntax.pretty_term ctxt equiv_rel, | |
| 168 | Pretty.str "equiv. thm:", | |
| 47093 | 169 | Syntax.pretty_term ctxt (prop_of equiv_thm), | 
| 170 | Pretty.str "quot. thm:", | |
| 171 | Syntax.pretty_term ctxt (prop_of quot_thm)]) | |
| 41444 | 172 | in | 
| 45280 | 173 | map (prt_quot o snd) (Symtab.dest (Quotients.get (Context.Proof ctxt))) | 
| 41444 | 174 | |> Pretty.big_list "quotients:" | 
| 175 | |> Pretty.writeln | |
| 176 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 177 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 178 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 179 | (* info about quotient constants *) | 
| 45279 | 180 | 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 | 181 | |
| 45279 | 182 | 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 | 183 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 184 | (* 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 | 185 | 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 | 186 | but overloaded constants share the same name *) | 
| 45280 | 187 | structure Quotconsts = Generic_Data | 
| 38759 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 188 | ( | 
| 45279 | 189 | type T = quotconsts list Symtab.table | 
| 38759 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 190 | val empty = Symtab.empty | 
| 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 191 | val extend = I | 
| 45279 | 192 | val merge = Symtab.merge_list eq_quotconsts | 
| 38759 
37a9092de102
simplification/standardization of some theory data;
 wenzelm parents: 
38756diff
changeset | 193 | ) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 194 | |
| 45279 | 195 | fun transform_quotconsts phi {qconst, rconst, def} =
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 196 |   {qconst = Morphism.term phi qconst,
 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 197 | rconst = Morphism.term phi rconst, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 198 | def = Morphism.thm phi def} | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 199 | |
| 45280 | 200 | fun update_quotconsts name qcinfo = Quotconsts.map (Symtab.cons_list (name, qcinfo)) | 
| 201 | ||
| 202 | fun dest_quotconsts ctxt = | |
| 203 | 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 | 204 | |
| 45350 
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
 Christian Urban <urbanc@in.tum.de> parents: 
45340diff
changeset | 205 | fun dest_quotconsts_global thy = | 
| 
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
 Christian Urban <urbanc@in.tum.de> parents: 
45340diff
changeset | 206 | 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: 
45340diff
changeset | 207 | |
| 
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
 Christian Urban <urbanc@in.tum.de> parents: 
45340diff
changeset | 208 | |
| 
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
 Christian Urban <urbanc@in.tum.de> parents: 
45340diff
changeset | 209 | |
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45294diff
changeset | 210 | fun lookup_quotconsts_global thy t = | 
| 45280 | 211 | let | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 212 | val (name, qty) = dest_Const t | 
| 45279 | 213 | fun matches (x: quotconsts) = | 
| 45280 | 214 | let val (name', qty') = dest_Const (#qconst x); | 
| 215 | 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 | 216 | in | 
| 45340 
98ec8b51af9c
prefer global Quotient_Info lookup to accomodate Quotient_Term, which is not quite localized yet (cf. 9fd6fce8a230);
 wenzelm parents: 
45294diff
changeset | 217 | (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: 
45273diff
changeset | 218 | 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: 
45273diff
changeset | 219 | | SOME l => find_first matches l) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 220 | end | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 221 | |
| 45279 | 222 | fun print_quotconsts ctxt = | 
| 41444 | 223 | let | 
| 224 |     fun prt_qconst {qconst, rconst, def} =
 | |
| 225 | Pretty.block (separate (Pretty.brk 1) | |
| 226 | [Syntax.pretty_term ctxt qconst, | |
| 227 | Pretty.str ":=", | |
| 228 | Syntax.pretty_term ctxt rconst, | |
| 229 | Pretty.str "as", | |
| 230 | Syntax.pretty_term ctxt (prop_of def)]) | |
| 231 | in | |
| 45280 | 232 | map prt_qconst (maps snd (Symtab.dest (Quotconsts.get (Context.Proof ctxt)))) | 
| 41444 | 233 | |> Pretty.big_list "quotient constants:" | 
| 234 | |> Pretty.writeln | |
| 235 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 236 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 237 | (* equivalence relation theorems *) | 
| 45279 | 238 | structure Equiv_Rules = Named_Thms | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 239 | ( | 
| 45294 | 240 |   val name = @{binding quot_equiv}
 | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 241 | val description = "equivalence relation theorems" | 
| 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 242 | ) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 243 | |
| 45279 | 244 | val equiv_rules = Equiv_Rules.get | 
| 245 | val equiv_rules_add = Equiv_Rules.add | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 246 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 247 | (* respectfulness theorems *) | 
| 45279 | 248 | structure Rsp_Rules = Named_Thms | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 249 | ( | 
| 45294 | 250 |   val name = @{binding quot_respect}
 | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 251 | val description = "respectfulness theorems" | 
| 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 252 | ) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 253 | |
| 45279 | 254 | val rsp_rules = Rsp_Rules.get | 
| 255 | val rsp_rules_add = Rsp_Rules.add | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 256 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 257 | (* preservation theorems *) | 
| 45279 | 258 | structure Prs_Rules = Named_Thms | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 259 | ( | 
| 45294 | 260 |   val name = @{binding quot_preserve}
 | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 261 | val description = "preservation theorems" | 
| 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 262 | ) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 263 | |
| 45279 | 264 | val prs_rules = Prs_Rules.get | 
| 265 | val prs_rules_add = Prs_Rules.add | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 266 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 267 | (* id simplification theorems *) | 
| 45279 | 268 | structure Id_Simps = Named_Thms | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 269 | ( | 
| 45294 | 270 |   val name = @{binding id_simps}
 | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 271 | val description = "identity simp rules for maps" | 
| 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 272 | ) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 273 | |
| 45279 | 274 | val id_simps = Id_Simps.get | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 275 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 276 | (* quotient theorems *) | 
| 45279 | 277 | structure Quotient_Rules = Named_Thms | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 278 | ( | 
| 45294 | 279 |   val name = @{binding quot_thm}
 | 
| 38764 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 280 | val description = "quotient theorems" | 
| 
e6a18808873c
more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
 wenzelm parents: 
38759diff
changeset | 281 | ) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 282 | |
| 45279 | 283 | val quotient_rules = Quotient_Rules.get | 
| 284 | val quotient_rules_add = Quotient_Rules.add | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 285 | |
| 41452 | 286 | |
| 287 | (* theory setup *) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 288 | |
| 41452 | 289 | val setup = | 
| 45281 | 290 | quotmaps_attribute_setup #> | 
| 45279 | 291 | Equiv_Rules.setup #> | 
| 292 | Rsp_Rules.setup #> | |
| 293 | Prs_Rules.setup #> | |
| 294 | Id_Simps.setup #> | |
| 295 | Quotient_Rules.setup | |
| 41452 | 296 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 297 | |
| 41452 | 298 | (* outer syntax commands *) | 
| 299 | ||
| 300 | val _ = | |
| 47308 | 301 |   Outer_Syntax.improper_command @{command_spec "print_quotmapsQ3"} "print quotient map functions"
 | 
| 45279 | 302 | (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 | 303 | |
| 41452 | 304 | val _ = | 
| 47308 | 305 |   Outer_Syntax.improper_command @{command_spec "print_quotientsQ3"} "print quotients"
 | 
| 45279 | 306 | (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 | 307 | |
| 41452 | 308 | val _ = | 
| 46961 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 wenzelm parents: 
46949diff
changeset | 309 |   Outer_Syntax.improper_command @{command_spec "print_quotconsts"} "print quotient constants"
 | 
| 45279 | 310 | (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 | 311 | |
| 45279 | 312 | end; |