| author | Christian Urban <urbanc@in.tum.de> | 
| Sat, 21 Nov 2009 03:12:50 +0100 | |
| changeset 323 | 31509c8cf72e | 
| parent 322 | d741ccea80d3 | 
| child 324 | bdbb52979790 | 
| permissions | -rw-r--r-- | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | signature QUOTIENT_INFO = | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 2 | sig | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 |   type maps_info = {mapfun: string, relfun: string}
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | val maps_lookup: theory -> string -> maps_info option | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | val maps_update_thy: string -> maps_info -> theory -> theory | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | val maps_update: string -> maps_info -> Proof.context -> Proof.context | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 |   type quotient_info = {qtyp: typ, rtyp: typ, rel: term, equiv_thm: thm}
 | 
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 9 | val print_quotinfo: Proof.context -> unit | 
| 311 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 10 | val quotdata_lookup_thy: theory -> string -> quotient_info option | 
| 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 11 | val quotdata_lookup: Proof.context -> string -> quotient_info option | 
| 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 12 | val quotdata_update_thy: string -> (typ * typ * term * thm) -> theory -> theory | 
| 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 13 | val quotdata_update: string -> (typ * typ * term * thm) -> Proof.context -> Proof.context | 
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 14 | |
| 322 
d741ccea80d3
flagged qenv-stuff as obsolete
 Christian Urban <urbanc@in.tum.de> parents: 
321diff
changeset | 15 | (*FIXME: obsolete *) | 
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 16 | type qenv = (typ * typ) list | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 17 | val mk_qenv: Proof.context -> qenv | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 18 | val lookup_qenv: ((typ * typ) -> bool) -> qenv -> typ -> (typ * typ) option | 
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 19 | |
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 20 |   type qconsts_info = {qconst: term, rconst: term}
 | 
| 318 
746b17e1d6d8
fixed the storage of qconst definitions
 Christian Urban <urbanc@in.tum.de> parents: 
314diff
changeset | 21 | val qconsts_transfer: morphism -> qconsts_info -> qconsts_info | 
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 22 | val qconsts_lookup: theory -> string -> qconsts_info option | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 23 | val qconsts_update_thy: string -> qconsts_info -> theory -> theory | 
| 321 
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
 Christian Urban <urbanc@in.tum.de> parents: 
320diff
changeset | 24 | val qconsts_update_gen: string -> qconsts_info -> Context.generic -> Context.generic | 
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 25 | val print_qconstinfo: Proof.context -> unit | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 26 | end; | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 27 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 28 | structure Quotient_Info: QUOTIENT_INFO = | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 29 | struct | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 30 | |
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 31 | |
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 32 | (* data containers *) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 33 | (*******************) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 34 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 35 | (* info about map- and rel-functions *) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 36 | type maps_info = {mapfun: string, relfun: string}
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 37 | |
| 306 
e7279efbe3dd
updated to new Theory_Data and to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
268diff
changeset | 38 | structure MapsData = Theory_Data | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 39 | (type T = maps_info Symtab.table | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 40 | val empty = Symtab.empty | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 41 | val extend = I | 
| 306 
e7279efbe3dd
updated to new Theory_Data and to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
268diff
changeset | 42 | val merge = Symtab.merge (K true)) | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 43 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 44 | val maps_lookup = Symtab.lookup o MapsData.get | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 45 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 46 | fun maps_update_thy k minfo = MapsData.map (Symtab.update (k, minfo)) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 47 | fun maps_update k minfo = ProofContext.theory (maps_update_thy k minfo) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 48 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 49 | fun maps_attribute_aux s minfo = Thm.declaration_attribute | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 50 | (fn thm => Context.mapping (maps_update_thy s minfo) (maps_update s minfo)) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 51 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 52 | (* attribute to be used in declare statements *) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 53 | fun maps_attribute (ctxt, (tystr, (mapstr, relstr))) = | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 54 | let | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 55 | val thy = ProofContext.theory_of ctxt | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 56 | val tyname = Sign.intern_type thy tystr | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 57 | val mapname = Sign.intern_const thy mapstr | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 58 | val relname = Sign.intern_const thy relstr | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 59 | in | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 60 |   maps_attribute_aux tyname {mapfun = mapname, relfun = relname}
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 61 | end | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 62 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 63 | val maps_attr_parser = | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 | Args.context -- Scan.lift | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 65 | ((Args.name --| OuterParse.$$$ "=") -- | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 66 |          (OuterParse.$$$ "(" |-- Args.name --| OuterParse.$$$ "," -- 
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 67 | Args.name --| OuterParse.$$$ ")")) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 68 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 69 | val _ = Context.>> (Context.map_theory | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 70 |          (Attrib.setup @{binding "map"} (maps_attr_parser >> maps_attribute) 
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 71 | "declaration of map information")) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 72 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 | (* info about the quotient types *) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 75 | type quotient_info = {qtyp: typ, rtyp: typ, rel: term, equiv_thm: thm}
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 76 | |
| 306 
e7279efbe3dd
updated to new Theory_Data and to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
268diff
changeset | 77 | structure QuotData = Theory_Data | 
| 311 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 78 | (type T = quotient_info Symtab.table | 
| 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 79 | val empty = Symtab.empty | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 80 | val extend = I | 
| 311 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 81 | val merge = Symtab.merge (K true)) | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 82 | |
| 320 
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
 Christian Urban <urbanc@in.tum.de> parents: 
318diff
changeset | 83 | fun quotdata_lookup_thy thy str = | 
| 
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
 Christian Urban <urbanc@in.tum.de> parents: 
318diff
changeset | 84 | Symtab.lookup (QuotData.get thy) (Sign.intern_type thy str) | 
| 
7d3d86beacd6
started regularize of rtrm/qtrm version; looks quite promising
 Christian Urban <urbanc@in.tum.de> parents: 
318diff
changeset | 85 | |
| 311 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 86 | val quotdata_lookup = quotdata_lookup_thy o ProofContext.theory_of | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 87 | |
| 314 | 88 | fun quotdata_update_thy qty_name (qty, rty, rel, equiv_thm) = | 
| 311 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 89 |       QuotData.map (Symtab.update (qty_name, {qtyp = qty, rtyp = rty, rel = rel, equiv_thm = equiv_thm}))
 | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 90 | |
| 311 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 91 | fun quotdata_update qty_name (qty, rty, rel, equiv_thm) = | 
| 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 92 | ProofContext.theory (quotdata_update_thy qty_name (qty, rty, rel, equiv_thm)) | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 93 | |
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 94 | fun print_quotinfo ctxt = | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 95 | let | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 96 |   fun prt_quot {qtyp, rtyp, rel, equiv_thm} = 
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 97 | Pretty.block (Library.separate (Pretty.brk 2) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 98 |           [Pretty.str ("quotient type:"), 
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 99 | Syntax.pretty_typ ctxt qtyp, | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 100 |            Pretty.str ("raw type:"), 
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 101 | Syntax.pretty_typ ctxt rtyp, | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 102 |            Pretty.str ("relation:"), 
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 103 | Syntax.pretty_term ctxt rel, | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 104 |            Pretty.str ("equiv. thm:"), 
 | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 105 | Syntax.pretty_term ctxt (prop_of equiv_thm)]) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 106 | in | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 107 | QuotData.get (ProofContext.theory_of ctxt) | 
| 311 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 108 | |> Symtab.dest | 
| 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 109 | |> map (prt_quot o snd) | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 110 | |> Pretty.big_list "quotients:" | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 111 | |> Pretty.writeln | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 112 | end | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 113 | |
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 114 | val _ = | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 115 | OuterSyntax.improper_command "print_quotients" "print out all quotients" | 
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 116 | OuterKeyword.diag (Scan.succeed (Toplevel.keep (print_quotinfo o Toplevel.context_of))) | 
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 117 | |
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 118 | |
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 119 | (* environments of quotient and raw types *) | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 120 | type qenv = (typ * typ) list | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 121 | |
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 122 | fun mk_qenv ctxt = | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 123 | let | 
| 311 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 124 | val thy = ProofContext.theory_of ctxt | 
| 
77fc6f3c0343
changed the quotdata to be a symtab table (needs fixing)
 Christian Urban <urbanc@in.tum.de> parents: 
310diff
changeset | 125 | val qinfo = (QuotData.get thy) |> Symtab.dest |> map snd | 
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 126 | in | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 127 |   map (fn {qtyp, rtyp, ...} => (qtyp, rtyp)) qinfo
 | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 128 | end | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 129 | |
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 130 | fun lookup_qenv _ [] _ = NONE | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 131 | | lookup_qenv eq ((qty, rty)::xs) qty' = | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 132 | if eq (qty', qty) then SOME (qty, rty) | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 133 | else lookup_qenv eq xs qty' | 
| 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 134 | |
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 135 | (* information about quotient constants *) | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 136 | type qconsts_info = {qconst: term, rconst: term}
 | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 137 | |
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 138 | structure QConstsData = Theory_Data | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 139 | (type T = qconsts_info Symtab.table | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 140 | val empty = Symtab.empty | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 141 | val extend = I | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 142 | val merge = Symtab.merge (K true)) | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 143 | |
| 318 
746b17e1d6d8
fixed the storage of qconst definitions
 Christian Urban <urbanc@in.tum.de> parents: 
314diff
changeset | 144 | fun qconsts_transfer phi {qconst, rconst} =
 | 
| 
746b17e1d6d8
fixed the storage of qconst definitions
 Christian Urban <urbanc@in.tum.de> parents: 
314diff
changeset | 145 |     {qconst = Morphism.term phi qconst,
 | 
| 
746b17e1d6d8
fixed the storage of qconst definitions
 Christian Urban <urbanc@in.tum.de> parents: 
314diff
changeset | 146 | rconst = Morphism.term phi rconst} | 
| 
746b17e1d6d8
fixed the storage of qconst definitions
 Christian Urban <urbanc@in.tum.de> parents: 
314diff
changeset | 147 | |
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 148 | val qconsts_lookup = Symtab.lookup o QConstsData.get | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 149 | |
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 150 | fun qconsts_update_thy k qcinfo = QConstsData.map (Symtab.update (k, qcinfo)) | 
| 321 
f46dc0ca08c3
simplified get_fun so that it uses directly rty and qty, instead of qenv
 Christian Urban <urbanc@in.tum.de> parents: 
320diff
changeset | 151 | fun qconsts_update_gen k qcinfo = | 
| 318 
746b17e1d6d8
fixed the storage of qconst definitions
 Christian Urban <urbanc@in.tum.de> parents: 
314diff
changeset | 152 | Context.mapping (qconsts_update_thy k qcinfo) I | 
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 153 | |
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 154 | fun print_qconstinfo ctxt = | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 155 | let | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 156 |   fun prt_qconst {qconst, rconst} = 
 | 
| 318 
746b17e1d6d8
fixed the storage of qconst definitions
 Christian Urban <urbanc@in.tum.de> parents: 
314diff
changeset | 157 | Pretty.block (separate (Pretty.brk 1) | 
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 158 | [Syntax.pretty_term ctxt qconst, | 
| 318 
746b17e1d6d8
fixed the storage of qconst definitions
 Christian Urban <urbanc@in.tum.de> parents: 
314diff
changeset | 159 | Pretty.str ":=", | 
| 310 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 160 | Syntax.pretty_term ctxt rconst]) | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 161 | in | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 162 | QConstsData.get (ProofContext.theory_of ctxt) | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 163 | |> Symtab.dest | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 164 | |> map (prt_qconst o snd) | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 165 | |> Pretty.big_list "quotient constants:" | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 166 | |> Pretty.writeln | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 167 | end | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 168 | |
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 169 | val _ = | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 170 | OuterSyntax.improper_command "print_quotconsts" "print out all quotient constants" | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 171 | OuterKeyword.diag (Scan.succeed (Toplevel.keep (print_qconstinfo o Toplevel.context_of))) | 
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 172 | |
| 
fec6301a1989
added a container for quotient constants (does not work yet though)
 Christian Urban <urbanc@in.tum.de> parents: 
306diff
changeset | 173 | |
| 264 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 174 | end; (* structure *) | 
| 
d0581fbc096c
split quotient.ML into two files
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 175 | |
| 314 | 176 | open Quotient_Info |