src/Pure/consts.ML
author wenzelm
Sun, 18 Mar 2012 13:04:22 +0100
changeset 47876 421760a1efe7
parent 46537 d83797ef0d2d
child 50007 0518bf89c777
permissions -rw-r--r--
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
more explicit Context.generic for Name_Space.declare/define and derivatives (NB: naming changed after Proof_Context.init_global);
prefer Context.pretty in low-level operations of structure Sorts/Type (defer full Syntax.init_pretty until error output);
simplified signatures;
     1 (*  Title:      Pure/consts.ML
     2     Author:     Makarius
     3 
     4 Polymorphic constants: declarations, abbreviations, additional type
     5 constraints.
     6 *)
     7 
     8 signature CONSTS =
     9 sig
    10   type T
    11   val eq_consts: T * T -> bool
    12   val retrieve_abbrevs: T -> string list -> term -> (term * term) list
    13   val dest: T ->
    14    {constants: (typ * term option) Name_Space.table,
    15     constraints: typ Name_Space.table}
    16   val the_const: T -> string -> string * typ                   (*exception TYPE*)
    17   val the_abbreviation: T -> string -> typ * term              (*exception TYPE*)
    18   val type_scheme: T -> string -> typ                          (*exception TYPE*)
    19   val is_monomorphic: T -> string -> bool                      (*exception TYPE*)
    20   val the_constraint: T -> string -> typ                       (*exception TYPE*)
    21   val space_of: T -> Name_Space.T
    22   val alias: Name_Space.naming -> binding -> string -> T -> T
    23   val is_concealed: T -> string -> bool
    24   val intern: T -> xstring -> string
    25   val extern: Proof.context -> T -> string -> xstring
    26   val intern_syntax: T -> xstring -> string
    27   val extern_syntax: Proof.context -> T -> string -> xstring
    28   val read_const: T -> string * Position.T -> term
    29   val certify: Context.pretty -> Type.tsig -> bool -> T -> term -> term  (*exception TYPE*)
    30   val typargs: T -> string * typ -> typ list
    31   val instance: T -> string * typ list -> typ
    32   val declare: Context.generic -> binding * typ -> T -> T
    33   val constrain: string * typ option -> T -> T
    34   val abbreviate: Context.generic -> Type.tsig -> string -> binding * term -> T -> (term * term) * T
    35   val revert_abbrev: string -> string -> T -> T
    36   val hide: bool -> string -> T -> T
    37   val empty: T
    38   val merge: T * T -> T
    39 end;
    40 
    41 structure Consts: CONSTS =
    42 struct
    43 
    44 (** consts type **)
    45 
    46 (* datatype T *)
    47 
    48 type decl = {T: typ, typargs: int list list};
    49 type abbrev = {rhs: term, normal_rhs: term, force_expand: bool};
    50 
    51 datatype T = Consts of
    52  {decls: (decl * abbrev option) Name_Space.table,
    53   constraints: typ Symtab.table,
    54   rev_abbrevs: (term * term) Item_Net.T Symtab.table};
    55 
    56 fun eq_consts
    57    (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1},
    58     Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) =
    59   pointer_eq (decls1, decls2) andalso
    60   pointer_eq (constraints1, constraints2) andalso
    61   pointer_eq (rev_abbrevs1, rev_abbrevs2);
    62 
    63 fun make_consts (decls, constraints, rev_abbrevs) =
    64   Consts {decls = decls, constraints = constraints, rev_abbrevs = rev_abbrevs};
    65 
    66 fun map_consts f (Consts {decls, constraints, rev_abbrevs}) =
    67   make_consts (f (decls, constraints, rev_abbrevs));
    68 
    69 
    70 (* reverted abbrevs *)
    71 
    72 val empty_abbrevs =
    73   Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') (single o #1);
    74 
    75 fun update_abbrevs mode abbrs =
    76   Symtab.map_default (mode, empty_abbrevs) (Item_Net.update abbrs);
    77 
    78 fun retrieve_abbrevs (Consts {rev_abbrevs, ...}) modes =
    79   let val nets = map_filter (Symtab.lookup rev_abbrevs) modes
    80   in fn t => maps (fn net => Item_Net.retrieve net t) nets end;
    81 
    82 
    83 (* dest consts *)
    84 
    85 fun dest (Consts {decls = (space, decls), constraints, ...}) =
    86  {constants = (space,
    87     Symtab.fold (fn (c, ({T, ...}, abbr)) =>
    88       Symtab.update (c, (T, Option.map #rhs abbr))) decls Symtab.empty),
    89   constraints = (space, constraints)};
    90 
    91 
    92 (* lookup consts *)
    93 
    94 fun the_entry (Consts {decls = (_, tab), ...}) c =
    95   (case Symtab.lookup_key tab c of
    96     SOME entry => entry
    97   | NONE => raise TYPE ("Unknown constant: " ^ quote c, [], []));
    98 
    99 fun the_const consts c =
   100   (case the_entry consts c of
   101     (c', ({T, ...}, NONE)) => (c', T)
   102   | _ => raise TYPE ("Not a logical constant: " ^ quote c, [], []));
   103 
   104 fun the_abbreviation consts c =
   105   (case the_entry consts c of
   106     (_, ({T, ...}, SOME {rhs, ...})) => (T, rhs)
   107   | _ => raise TYPE ("Not an abbreviated constant: " ^ quote c, [], []));
   108 
   109 fun the_decl consts = #1 o #2 o the_entry consts;
   110 val type_scheme = #T oo the_decl;
   111 val type_arguments = #typargs oo the_decl;
   112 
   113 val is_monomorphic = null oo type_arguments;
   114 
   115 fun the_constraint (consts as Consts {constraints, ...}) c =
   116   (case Symtab.lookup constraints c of
   117     SOME T => T
   118   | NONE => type_scheme consts c);
   119 
   120 
   121 (* name space and syntax *)
   122 
   123 fun space_of (Consts {decls = (space, _), ...}) = space;
   124 
   125 fun alias naming binding name = map_consts (fn ((space, decls), constraints, rev_abbrevs) =>
   126   ((Name_Space.alias naming binding name space, decls), constraints, rev_abbrevs));
   127 
   128 val is_concealed = Name_Space.is_concealed o space_of;
   129 
   130 val intern = Name_Space.intern o space_of;
   131 fun extern ctxt = Name_Space.extern ctxt o space_of;
   132 
   133 fun intern_syntax consts s =
   134   (case try Lexicon.unmark_const s of
   135     SOME c => c
   136   | NONE => intern consts s);
   137 
   138 fun extern_syntax ctxt consts s =
   139   (case try Lexicon.unmark_const s of
   140     SOME c => extern ctxt consts c
   141   | NONE => s);
   142 
   143 
   144 (* read_const *)
   145 
   146 fun read_const consts (raw_c, pos) =
   147   let
   148     val c = intern consts raw_c;
   149     val T = type_scheme consts c handle TYPE (msg, _, _) => error (msg ^ Position.str_of pos);
   150   in Const (c, T) end;
   151 
   152 
   153 (* certify *)
   154 
   155 fun certify pp tsig do_expand consts =
   156   let
   157     fun err msg (c, T) =
   158       raise TYPE (msg ^ " " ^ quote c ^ " :: " ^
   159         Syntax.string_of_typ (Syntax.init_pretty pp) T, [], []);
   160     val certT = Type.cert_typ tsig;
   161     fun cert tm =
   162       let
   163         val (head, args) = Term.strip_comb tm;
   164         val args' = map cert args;
   165         fun comb head' = Term.list_comb (head', args');
   166       in
   167         (case head of
   168           Abs (x, T, t) => comb (Abs (x, certT T, cert t))
   169         | Const (c, T) =>
   170             let
   171               val T' = certT T;
   172               val (_, ({T = U, ...}, abbr)) = the_entry consts c;
   173               fun expand u =
   174                 Term.betapplys (Envir.expand_atom T' (U, u) handle TYPE _ =>
   175                   err "Illegal type for abbreviation" (c, T), args');
   176             in
   177               if not (Type.raw_instance (T', U)) then
   178                 err "Illegal type for constant" (c, T)
   179               else
   180                 (case abbr of
   181                   SOME {rhs, normal_rhs, force_expand} =>
   182                     if do_expand then expand normal_rhs
   183                     else if force_expand then expand rhs
   184                     else comb head
   185                 | _ => comb head)
   186             end
   187         | _ => comb head)
   188       end;
   189   in cert end;
   190 
   191 
   192 (* typargs -- view actual const type as instance of declaration *)
   193 
   194 local
   195 
   196 fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos
   197   | args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos)
   198   | args_of (TFree _) _ = I
   199 and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is
   200   | args_of_list [] _ _ = I;
   201 
   202 fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is
   203   | subscript T [] = T
   204   | subscript _ _ = raise Subscript;
   205 
   206 in
   207 
   208 fun typargs_of T = map #2 (rev (args_of T [] []));
   209 
   210 fun typargs consts (c, T) = map (subscript T) (type_arguments consts c);
   211 
   212 end;
   213 
   214 fun instance consts (c, Ts) =
   215   let
   216     val declT = type_scheme consts c;
   217     val vars = map Term.dest_TVar (typargs consts (c, declT));
   218     val inst = vars ~~ Ts handle ListPair.UnequalLengths =>
   219       raise TYPE ("Consts.instance", Ts, [Const (c, dummyT)]);
   220   in declT |> Term_Subst.instantiateT inst end;
   221 
   222 
   223 
   224 (** build consts **)
   225 
   226 (* name space *)
   227 
   228 fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs) =>
   229   (apfst (Name_Space.hide fully c) decls, constraints, rev_abbrevs));
   230 
   231 
   232 (* declarations *)
   233 
   234 fun declare context (b, declT) =
   235   map_consts (fn (decls, constraints, rev_abbrevs) =>
   236     let
   237       val decl = {T = declT, typargs = typargs_of declT};
   238       val _ = Binding.check b;
   239       val (_, decls') = decls |> Name_Space.define context true (b, (decl, NONE));
   240     in (decls', constraints, rev_abbrevs) end);
   241 
   242 
   243 (* constraints *)
   244 
   245 fun constrain (c, C) consts =
   246   consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>
   247     (#2 (the_entry consts c) handle TYPE (msg, _, _) => error msg;
   248       (decls,
   249         constraints |> (case C of SOME T => Symtab.update (c, T) | NONE => Symtab.delete_safe c),
   250         rev_abbrevs)));
   251 
   252 
   253 (* abbreviations *)
   254 
   255 local
   256 
   257 fun strip_abss (t as Abs (x, T, b)) =
   258       if Term.is_dependent b then strip_abss b |>> cons (x, T)  (* FIXME decr!? *)
   259       else ([], t)
   260   | strip_abss t = ([], t);
   261 
   262 fun rev_abbrev lhs rhs =
   263   let
   264     val (xs, body) = strip_abss (Envir.beta_eta_contract rhs);
   265     val vars = fold (fn (x, T) => cons (Var ((x, 0), T))) (Term.rename_wrt_term body xs) [];
   266   in (Term.subst_bounds (rev vars, body), Term.list_comb (lhs, vars)) end;
   267 
   268 in
   269 
   270 fun abbreviate context tsig mode (b, raw_rhs) consts =
   271   let
   272     val pp = Context.pretty_generic context;
   273     val cert_term = certify pp tsig false consts;
   274     val expand_term = certify pp tsig true consts;
   275     val force_expand = mode = Print_Mode.internal;
   276 
   277     val _ = Term.exists_subterm Term.is_Var raw_rhs andalso
   278       error ("Illegal schematic variables on rhs of abbreviation " ^ Binding.print b);
   279 
   280     val rhs = raw_rhs
   281       |> Term.map_types (Type.cert_typ tsig)
   282       |> cert_term
   283       |> Term.close_schematic_term;
   284     val normal_rhs = expand_term rhs;
   285     val T = Term.fastype_of rhs;
   286     val lhs = Const (Name_Space.full_name (Name_Space.naming_of context) b, T);
   287   in
   288     consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>
   289       let
   290         val decl = {T = T, typargs = typargs_of T};
   291         val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand};
   292         val _ = Binding.check b;
   293         val (_, decls') = decls
   294           |> Name_Space.define context true (b, (decl, SOME abbr));
   295         val rev_abbrevs' = rev_abbrevs
   296           |> update_abbrevs mode (rev_abbrev lhs rhs);
   297       in (decls', constraints, rev_abbrevs') end)
   298     |> pair (lhs, rhs)
   299   end;
   300 
   301 fun revert_abbrev mode c consts = consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>
   302   let
   303     val (T, rhs) = the_abbreviation consts c;
   304     val rev_abbrevs' = rev_abbrevs
   305       |> update_abbrevs mode (rev_abbrev (Const (c, T)) rhs);
   306   in (decls, constraints, rev_abbrevs') end);
   307 
   308 end;
   309 
   310 
   311 (* empty and merge *)
   312 
   313 val empty =
   314   make_consts (Name_Space.empty_table Isabelle_Markup.constantN, Symtab.empty, Symtab.empty);
   315 
   316 fun merge
   317    (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1},
   318     Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) =
   319   let
   320     val decls' = Name_Space.merge_tables (decls1, decls2);
   321     val constraints' = Symtab.join (K fst) (constraints1, constraints2);
   322     val rev_abbrevs' = Symtab.join (K Item_Net.merge) (rev_abbrevs1, rev_abbrevs2);
   323   in make_consts (decls', constraints', rev_abbrevs') end;
   324 
   325 end;