src/Pure/consts.ML
author wenzelm
Fri, 17 Feb 2006 20:03:10 +0100
changeset 19098 fc736dbbe333
parent 19027 adf6fb0db28a
child 19364 38799cfa34e5
permissions -rw-r--r--
constrain: assert const declaration, optional type (i.e. may delete constraints);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/consts.ML
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
     4
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
     5
Polymorphic constants: declarations, abbreviations, additional type
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
     6
constraints.
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
     7
*)
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
     8
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
     9
signature CONSTS =
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    10
sig
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    11
  type T
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    12
  val eq_consts: T * T -> bool
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    13
  val dest: T ->
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    14
   {constants: (typ * term option) NameSpace.table,
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    15
    constraints: typ NameSpace.table}
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    16
  val space_of: T -> NameSpace.T
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    17
  val abbrevs_of: T -> (term * term) list
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    18
  val declaration: T -> string -> typ                           (*exception TYPE*)
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    19
  val monomorphic: T -> string -> bool                          (*exception TYPE*)
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    20
  val constraint: T -> string -> typ                            (*exception TYPE*)
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    21
  val certify: Pretty.pp -> Type.tsig -> T -> term -> term      (*exception TYPE*)
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    22
  val read_const: T -> string -> term
18146
47463b1825c6 uncurried Consts.typargs;
wenzelm
parents: 18117
diff changeset
    23
  val typargs: T -> string * typ -> typ list
18163
9b729737bf14 added instance;
wenzelm
parents: 18146
diff changeset
    24
  val instance: T -> string * typ list -> typ
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    25
  val declare: NameSpace.naming -> bstring * typ -> T -> T
18992
910fbe64033c abbrevs: store in reverted orientation;
wenzelm
parents: 18965
diff changeset
    26
  val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> bool -> bstring * term -> T -> T
19098
fc736dbbe333 constrain: assert const declaration, optional type (i.e. may delete constraints);
wenzelm
parents: 19027
diff changeset
    27
  val constrain: string * typ option -> T -> T
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    28
  val hide: bool -> string -> T -> T
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    29
  val empty: T
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    30
  val merge: T * T -> T
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    31
end
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    32
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    33
structure Consts: CONSTS =
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    34
struct
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    35
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    36
(** datatype T **)
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    37
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    38
datatype decl =
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    39
  LogicalConst of typ * int list list |
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    40
  Abbreviation of typ * term;
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    41
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    42
datatype T = Consts of
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    43
 {decls: (decl * serial) NameSpace.table,
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    44
  abbrevs: (term * term) list,
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    45
  constraints: typ Symtab.table} * stamp;
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    46
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    47
fun eq_consts (Consts (_, s1), Consts (_, s2)) = s1 = s2;
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    48
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    49
fun make_consts (decls, abbrevs, constraints) =
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    50
  Consts ({decls = decls, abbrevs = abbrevs, constraints = constraints}, stamp ());
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    51
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    52
fun map_consts f (Consts ({decls, abbrevs, constraints}, _)) =
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    53
  make_consts (f (decls, abbrevs, constraints));
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    54
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    55
fun dest (Consts ({decls = (space, decls), abbrevs = _, constraints}, _)) =
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    56
 {constants = (space, Symtab.fold
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    57
    (fn (c, (LogicalConst (T, _), _)) => Symtab.update (c, (T, NONE))
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    58
      | (c, (Abbreviation (T, t), _)) => Symtab.update (c, (T, SOME t))) decls Symtab.empty),
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    59
  constraints = (space, constraints)};
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    60
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    61
fun space_of (Consts ({decls = (space, _), ...}, _)) = space;
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    62
fun abbrevs_of (Consts ({abbrevs, ...}, _)) = abbrevs;
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    63
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    64
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    65
(* lookup consts *)
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    66
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    67
fun err_undeclared c = raise TYPE ("Undeclared constant: " ^ quote c, [], []);
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    68
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    69
fun the_const (Consts ({decls = (_, tab), ...}, _)) c =
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    70
  (case Symtab.lookup tab c of SOME (decl, _) => decl | NONE => err_undeclared c);
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    71
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    72
fun logical_const consts c =
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    73
  (case the_const consts c of LogicalConst d => d | _ => err_undeclared c);
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    74
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    75
fun declaration consts c = #1 (logical_const consts c);
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    76
fun monomorphic consts c = null (#2 (logical_const consts c));
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    77
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
    78
fun constraint (consts as Consts ({constraints, ...}, _)) c =
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    79
  (case Symtab.lookup constraints c of
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    80
    SOME T => T
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    81
  | NONE => (case the_const consts c of LogicalConst (T, _) => T | Abbreviation (T, _) => T));
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    82
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
    83
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    84
(* certify: check/expand consts *)
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
    85
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    86
fun certify pp tsig consts =
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    87
  let
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    88
    fun err msg (c, T) =
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    89
      raise TYPE (msg ^ " " ^ quote c ^ " :: " ^ Pretty.string_of_typ pp T, [], []);
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    90
    fun cert tm =
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    91
      let
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    92
        val (head, args) = Term.strip_comb tm;
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    93
        val args' = map cert args;
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    94
      in
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    95
        (case head of
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    96
          Const (c, T) =>
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    97
            (case the_const consts c of
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    98
              LogicalConst (U, _) =>
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
    99
                if Type.typ_instance tsig (T, U) then Term.list_comb (head, args')
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   100
                else err "Illegal type for constant" (c, T)
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   101
            | Abbreviation (U, u) =>
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   102
                Term.betapplys (Envir.expand_atom tsig T (U, u) handle TYPE _ =>
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   103
                  err "Illegal type for abbreviation" (c, T), args'))
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   104
        | Abs (x, T, t) => Term.list_comb (Abs (x, T, cert t), args')
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   105
        | _ => Term.list_comb (head, args'))
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   106
      end;
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   107
  in cert end;
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   108
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   109
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   110
(* read_const *)
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   111
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   112
fun read_const consts raw_c =
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   113
  let
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   114
    val c = NameSpace.intern (space_of consts) raw_c;
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   115
    val _ = the_const consts c handle TYPE (msg, _, _) => error msg;
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   116
  in Const (c, dummyT) end;
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   117
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   118
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   119
(* typargs -- view actual const type as instance of declaration *)
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   120
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   121
fun sub (Type (_, Ts)) (i :: is) = sub (nth Ts i) is
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   122
  | sub T [] = T
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   123
  | sub T _ = raise Subscript;
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   124
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   125
fun typargs consts (c, T) = map (sub T) (#2 (logical_const consts c));
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   126
18163
9b729737bf14 added instance;
wenzelm
parents: 18146
diff changeset
   127
fun instance consts (c, Ts) =
9b729737bf14 added instance;
wenzelm
parents: 18146
diff changeset
   128
  let
9b729737bf14 added instance;
wenzelm
parents: 18146
diff changeset
   129
    val declT = declaration consts c;
9b729737bf14 added instance;
wenzelm
parents: 18146
diff changeset
   130
    val vars = map Term.dest_TVar (typargs consts (c, declT));
9b729737bf14 added instance;
wenzelm
parents: 18146
diff changeset
   131
  in declT |> Term.instantiateT (vars ~~ Ts) end;
9b729737bf14 added instance;
wenzelm
parents: 18146
diff changeset
   132
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   133
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   134
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   135
(** declare consts **)
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   136
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   137
fun err_dup_consts cs =
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   138
  error ("Duplicate declaration of constant(s) " ^ commas_quote cs);
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   139
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   140
fun err_inconsistent_constraints cs =
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   141
  error ("Inconsistent type constraints for constant(s) " ^ commas_quote cs);
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   142
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   143
fun extend_decls naming decl tab = NameSpace.extend_table naming (tab, [decl])
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   144
  handle Symtab.DUPS cs => err_dup_consts cs;
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   145
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   146
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   147
(* name space *)
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   148
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   149
fun hide fully c = map_consts (fn (decls, abbrevs, constraints) =>
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   150
  (apfst (NameSpace.hide fully c) decls, abbrevs, constraints));
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   151
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   152
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   153
(* declarations *)
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   154
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   155
fun declare naming (c, declT) = map_consts (fn (decls, abbrevs, constraints) =>
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   156
  let
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   157
    fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   158
      | args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos)
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   159
      | args_of (TFree _) _ = I
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   160
    and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   161
      | args_of_list [] _ _ = I;
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   162
    val decl = (c, (LogicalConst (declT, map #2 (rev (args_of declT [] []))), serial ()));
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   163
  in (extend_decls naming decl decls, abbrevs, constraints) end);
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   164
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   165
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   166
(* abbreviations *)
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   167
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   168
local
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   169
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   170
fun strip_abss (Abs (a, T, t)) =
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   171
      ([], t) :: map (fn (xs, b) => ((a, T) :: xs, b)) (strip_abss t)
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   172
  | strip_abss t = [([], t)];
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   173
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   174
fun revert_abbrev const rhs =
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   175
  let
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   176
    fun abbrev (xs, body) =
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   177
      let val vars = fold (fn (x, T) => cons (Var ((x, 0), T))) (Term.rename_wrt_term body xs) []
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   178
      in (Term.subst_bounds (rev vars, body), Term.list_comb (Const const, vars)) end;
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   179
  in map abbrev (rev (strip_abss (Envir.beta_eta_contract rhs))) end;
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   180
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   181
in
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   182
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   183
fun abbreviate pp tsig naming revert (c, raw_rhs) consts =
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   184
    consts |> map_consts (fn (decls, abbrevs, constraints) =>
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   185
  let
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   186
    val rhs = certify pp tsig consts raw_rhs;
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   187
    val T = Term.fastype_of rhs;
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   188
    val decls' = decls |> extend_decls naming (c, (Abbreviation (T, rhs), serial ()));
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   189
    val abbrevs' =
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   190
      if revert then revert_abbrev (NameSpace.full naming c, T) rhs @ abbrevs else abbrevs;
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   191
  in (decls', abbrevs', constraints) end);
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   192
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   193
end;
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   194
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   195
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   196
(* constraints *)
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   197
19098
fc736dbbe333 constrain: assert const declaration, optional type (i.e. may delete constraints);
wenzelm
parents: 19027
diff changeset
   198
fun constrain (c, opt_T) consts = consts |> map_consts (fn (decls, abbrevs, constraints) =>
fc736dbbe333 constrain: assert const declaration, optional type (i.e. may delete constraints);
wenzelm
parents: 19027
diff changeset
   199
  (the_const consts c handle TYPE (msg, _, _) => error msg;
fc736dbbe333 constrain: assert const declaration, optional type (i.e. may delete constraints);
wenzelm
parents: 19027
diff changeset
   200
    (decls, abbrevs, constraints |>
fc736dbbe333 constrain: assert const declaration, optional type (i.e. may delete constraints);
wenzelm
parents: 19027
diff changeset
   201
      (case opt_T of SOME T => Symtab.update (c, T) | NONE => Symtab.delete_safe c))));
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   202
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   203
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   204
(* empty and merge *)
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   205
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   206
val empty = make_consts (NameSpace.empty_table, [], Symtab.empty);
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   207
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   208
fun merge
19027
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   209
   (Consts ({decls = decls1, abbrevs = abbrevs1, constraints = constraints1}, _),
adf6fb0db28a added eq_consts;
wenzelm
parents: 18992
diff changeset
   210
    Consts ({decls = decls2, abbrevs = abbrevs2, constraints = constraints2}, _)) =
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   211
  let
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   212
    val decls' = NameSpace.merge_tables (eq_snd (op =)) (decls1, decls2)
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   213
      handle Symtab.DUPS cs => err_dup_consts cs;
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   214
    val abbrevs' =
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   215
      Library.merge (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') (abbrevs1, abbrevs2);
18935
f22be3d61ed5 added abbreviations;
wenzelm
parents: 18163
diff changeset
   216
    val constraints' = Symtab.merge (op =) (constraints1, constraints2)
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   217
      handle Symtab.DUPS cs => err_inconsistent_constraints cs;
18965
3b76383e3ab3 renamed space to space_of;
wenzelm
parents: 18935
diff changeset
   218
  in make_consts (decls', abbrevs', constraints') end;
18060
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   219
afcc28d16629 Polymorphic constants.
wenzelm
parents:
diff changeset
   220
end;