src/HOL/Tools/Quotient/quotient_info.ML
author wenzelm
Thu, 27 Oct 2011 21:52:57 +0200
changeset 45281 29e88714ffe4
parent 45280 9fd6fce8a230
child 45294 3c5d3d286055
permissions -rw-r--r--
more standard attribute setup;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 36960
diff changeset
     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
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
     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
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
     9
  type quotmaps = {mapfun: string, relmap: string}
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
    10
  val lookup_quotmaps: Proof.context -> string -> quotmaps option
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    11
  val print_quotmaps: Proof.context -> unit
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    13
  type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    14
  val transform_quotients: morphism -> quotients -> quotients
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
    15
  val lookup_quotients: Proof.context -> string -> quotients option
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    16
  val update_quotients: string -> quotients -> Context.generic -> Context.generic
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    17
  val dest_quotients: Proof.context -> quotients list
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    18
  val print_quotients: Proof.context -> unit
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    20
  type quotconsts = {qconst: term, rconst: term, def: thm}
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    21
  val transform_quotconsts: morphism -> quotconsts -> quotconsts
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
    22
  val lookup_quotconsts: Proof.context -> term -> quotconsts option
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    23
  val update_quotconsts: string -> quotconsts -> Context.generic -> Context.generic
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    24
  val dest_quotconsts: Proof.context -> quotconsts list
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    25
  val print_quotconsts: Proof.context -> unit
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    27
  val equiv_rules: Proof.context -> thm list
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
  val equiv_rules_add: attribute
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    29
  val rsp_rules: Proof.context -> thm list
35314
cbdf785a1eb3 export prs_rules and rsp_rules attributes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35278
diff changeset
    30
  val rsp_rules_add: attribute
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    31
  val prs_rules: Proof.context -> thm list
35314
cbdf785a1eb3 export prs_rules and rsp_rules attributes
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 35278
diff changeset
    32
  val prs_rules_add: attribute
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    33
  val id_simps: Proof.context -> thm list
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    34
  val quotient_rules: Proof.context -> thm list
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
  val quotient_rules_add: attribute
41452
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
    36
  val setup: theory -> theory
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
end;
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
structure Quotient_Info: QUOTIENT_INFO =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
struct
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
(** data containers **)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
41472
f6ab14e61604 misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents: 41452
diff changeset
    44
(* 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
    45
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
(* info about map- and rel-functions for a type *)
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    47
type quotmaps = {mapfun: string, relmap: string}
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
    49
structure Quotmaps = Generic_Data
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    50
(
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    51
  type T = quotmaps Symtab.table
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    52
  val empty = Symtab.empty
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    53
  val extend = I
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    54
  fun merge data = Symtab.merge (K true) data
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    55
)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
    57
val lookup_quotmaps = Symtab.lookup o Quotmaps.get o Context.Proof
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
45281
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    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
45281
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    61
val quotmaps_attribute_setup =
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    62
  Attrib.setup @{binding map}
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    63
    ((Args.type_name false --| Scan.lift (Parse.$$$ "=")) --  (* FIXME Args.type_name true (requires "set" type) *)
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    64
      (Scan.lift (Parse.$$$ "(") |-- Args.const_proper true --| Scan.lift (Parse.$$$ ",") --
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    65
        Args.const_proper true --| Scan.lift (Parse.$$$ ")")) >>
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    66
      (fn (tyname, (mapname, relname)) =>
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    67
        let val minfo = {mapfun = mapname, relmap = relname}
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    68
        in Thm.declaration_attribute (fn _ => Quotmaps.map (Symtab.update (tyname, minfo))) end))
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
    69
    "declaration of map information"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    71
fun print_quotmaps ctxt =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    72
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    73
    fun prt_map (ty_name, {mapfun, relmap}) =
41445
wenzelm
parents: 41444
diff changeset
    74
      Pretty.block (separate (Pretty.brk 2)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    75
        (map Pretty.str
41445
wenzelm
parents: 41444
diff changeset
    76
         ["type:", ty_name,
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    77
          "map:", mapfun,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    78
          "relation map:", relmap]))
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    79
  in
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
    80
    map prt_map (Symtab.dest (Quotmaps.get (Context.Proof ctxt)))
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    81
    |> Pretty.big_list "maps for type constructors:"
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    82
    |> Pretty.writeln
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    83
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
(* info about quotient types *)
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    87
type quotients = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
    89
structure Quotients = Generic_Data
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    90
(
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    91
  type T = quotients Symtab.table
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    92
  val empty = Symtab.empty
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    93
  val extend = I
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    94
  fun merge data = Symtab.merge (K true) data
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
    95
)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
    97
fun transform_quotients phi {qtyp, rtyp, equiv_rel, equiv_thm} =
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
  {qtyp = Morphism.typ phi qtyp,
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
   rtyp = Morphism.typ phi rtyp,
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
   equiv_rel = Morphism.term phi equiv_rel,
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
   equiv_thm = Morphism.thm phi equiv_thm}
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   103
val lookup_quotients = Symtab.lookup o Quotients.get o Context.Proof
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   105
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
   106
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   107
fun dest_quotients ctxt =  (* FIXME slightly expensive way to retrieve data *)
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   108
  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
   109
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   110
fun print_quotients ctxt =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   111
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   112
    fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm} =
41445
wenzelm
parents: 41444
diff changeset
   113
      Pretty.block (separate (Pretty.brk 2)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   114
       [Pretty.str "quotient type:",
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   115
        Syntax.pretty_typ ctxt qtyp,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   116
        Pretty.str "raw type:",
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   117
        Syntax.pretty_typ ctxt rtyp,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   118
        Pretty.str "relation:",
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   119
        Syntax.pretty_term ctxt equiv_rel,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   120
        Pretty.str "equiv. thm:",
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   121
        Syntax.pretty_term ctxt (prop_of equiv_thm)])
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   122
  in
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   123
    map (prt_quot o snd) (Symtab.dest (Quotients.get (Context.Proof ctxt)))
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   124
    |> Pretty.big_list "quotients:"
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   125
    |> Pretty.writeln
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   126
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   127
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
(* info about quotient constants *)
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   130
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
   131
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   132
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
   133
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
(* 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
   135
   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
   136
   but overloaded constants share the same name *)
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   137
structure Quotconsts = Generic_Data
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
   138
(
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   139
  type T = quotconsts list Symtab.table
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
   140
  val empty = Symtab.empty
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
   141
  val extend = I
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   142
  val merge = Symtab.merge_list eq_quotconsts
38759
37a9092de102 simplification/standardization of some theory data;
wenzelm
parents: 38756
diff changeset
   143
)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   145
fun transform_quotconsts phi {qconst, rconst, def} =
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
  {qconst = Morphism.term phi qconst,
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
   rconst = Morphism.term phi rconst,
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
   def = Morphism.thm phi def}
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   150
fun update_quotconsts name qcinfo = Quotconsts.map (Symtab.cons_list (name, qcinfo))
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   151
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   152
fun dest_quotconsts ctxt =
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   153
  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
   154
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   155
fun lookup_quotconsts ctxt t =
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   156
  let
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   157
    val thy = Proof_Context.theory_of ctxt
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
    val (name, qty) = dest_Const t
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   160
    fun matches (x: quotconsts) =
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   161
      let val (name', qty') = dest_Const (#qconst x);
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   162
      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
   163
  in
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   164
    (case Symtab.lookup (Quotconsts.get (Context.Proof ctxt)) 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
   165
      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
   166
    | SOME l => find_first matches l)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   167
  end
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   169
fun print_quotconsts ctxt =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   170
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   171
    fun prt_qconst {qconst, rconst, def} =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   172
      Pretty.block (separate (Pretty.brk 1)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   173
       [Syntax.pretty_term ctxt qconst,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   174
        Pretty.str ":=",
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   175
        Syntax.pretty_term ctxt rconst,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   176
        Pretty.str "as",
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   177
        Syntax.pretty_term ctxt (prop_of def)])
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   178
  in
45280
9fd6fce8a230 localized quotient data;
wenzelm
parents: 45279
diff changeset
   179
    map prt_qconst (maps snd (Symtab.dest (Quotconsts.get (Context.Proof ctxt))))
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   180
    |> Pretty.big_list "quotient constants:"
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   181
    |> Pretty.writeln
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   182
  end
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
(* equivalence relation theorems *)
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   185
structure Equiv_Rules = Named_Thms
38764
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   186
(
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   187
  val name = "quot_equiv"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   188
  val description = "equivalence relation theorems"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   189
)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   190
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   191
val equiv_rules = Equiv_Rules.get
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   192
val equiv_rules_add = Equiv_Rules.add
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   193
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   194
(* respectfulness theorems *)
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   195
structure Rsp_Rules = Named_Thms
38764
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   196
(
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   197
  val name = "quot_respect"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   198
  val description = "respectfulness theorems"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   199
)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   200
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   201
val rsp_rules = Rsp_Rules.get
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   202
val rsp_rules_add = Rsp_Rules.add
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   203
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   204
(* preservation theorems *)
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   205
structure Prs_Rules = Named_Thms
38764
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   206
(
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   207
  val name = "quot_preserve"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   208
  val description = "preservation theorems"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   209
)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   210
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   211
val prs_rules = Prs_Rules.get
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   212
val prs_rules_add = Prs_Rules.add
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   213
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   214
(* id simplification theorems *)
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   215
structure Id_Simps = Named_Thms
38764
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   216
(
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   217
  val name = "id_simps"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   218
  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: 38759
diff changeset
   219
)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   220
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   221
val id_simps = Id_Simps.get
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   223
(* quotient theorems *)
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   224
structure Quotient_Rules = Named_Thms
38764
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   225
(
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   226
  val name = "quot_thm"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   227
  val description = "quotient theorems"
e6a18808873c more uniform descriptions, which end up in the collective output of 'print_attributes' for example;
wenzelm
parents: 38759
diff changeset
   228
)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   229
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   230
val quotient_rules = Quotient_Rules.get
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   231
val quotient_rules_add = Quotient_Rules.add
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   232
41452
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   233
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   234
(* theory setup *)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   235
41452
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   236
val setup =
45281
29e88714ffe4 more standard attribute setup;
wenzelm
parents: 45280
diff changeset
   237
  quotmaps_attribute_setup #>
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   238
  Equiv_Rules.setup #>
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   239
  Rsp_Rules.setup #>
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   240
  Prs_Rules.setup #>
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   241
  Id_Simps.setup #>
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   242
  Quotient_Rules.setup
41452
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   243
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   244
41452
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   245
(* outer syntax commands *)
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   246
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   247
val _ =
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   248
  Outer_Syntax.improper_command "print_quotmaps" "print quotient map functions" Keyword.diag
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   249
    (Scan.succeed (Toplevel.keep (print_quotmaps o Toplevel.context_of)))
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   250
41452
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   251
val _ =
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   252
  Outer_Syntax.improper_command "print_quotients" "print quotients" Keyword.diag
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   253
    (Scan.succeed (Toplevel.keep (print_quotients o Toplevel.context_of)))
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   254
41452
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   255
val _ =
c291e0826902 more standard package setup;
wenzelm
parents: 41451
diff changeset
   256
  Outer_Syntax.improper_command "print_quotconsts" "print quotient constants" Keyword.diag
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   257
    (Scan.succeed (Toplevel.keep (print_quotconsts o Toplevel.context_of)))
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   258
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 45278
diff changeset
   259
end;