# HG changeset patch # User wenzelm # Date 1139338607 -3600 # Node ID 3b76383e3ab3aaee8b146524cb95a5ab0bc94e53 # Parent 67f572e0323695af6c50a67355b310bf7978db33 renamed space to space_of; removed expansion; added abbrevs_of; added read_const; certify: substitute arguments into expanded const; tuned; diff -r 67f572e03236 -r 3b76383e3ab3 src/Pure/consts.ML --- a/src/Pure/consts.ML Tue Feb 07 19:56:45 2006 +0100 +++ b/src/Pure/consts.ML Tue Feb 07 19:56:47 2006 +0100 @@ -12,16 +12,17 @@ val dest: T -> {constants: (typ * term option) NameSpace.table, constraints: typ NameSpace.table} - val space: T -> NameSpace.T + val space_of: T -> NameSpace.T + val abbrevs_of: T -> (term * term) list val declaration: T -> string -> typ (*exception TYPE*) val monomorphic: T -> string -> bool (*exception TYPE*) val constraint: T -> string -> typ (*exception TYPE*) - val expansion: Type.tsig -> T -> string * typ -> term option (*exception TYPE*) val certify: Pretty.pp -> Type.tsig -> T -> term -> term (*exception TYPE*) + val read_const: T -> string -> term val typargs: T -> string * typ -> typ list val instance: T -> string * typ list -> typ val declare: NameSpace.naming -> bstring * typ -> T -> T - val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> string * term -> T -> T + val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> bool -> string * term -> T -> T val constrain: string * typ -> T -> T val hide: bool -> string -> T -> T val empty: T @@ -39,18 +40,23 @@ datatype T = Consts of {decls: (decl * serial) NameSpace.table, + abbrevs: (term * term) list, constraints: typ Symtab.table}; -fun make_consts (decls, constraints) = Consts {decls = decls, constraints = constraints}; -fun map_consts f (Consts {decls, constraints}) = make_consts (f (decls, constraints)); +fun make_consts (decls, abbrevs, constraints) = + Consts {decls = decls, abbrevs = abbrevs, constraints = constraints}; -fun dest (Consts {decls = (space, decls), constraints}) = +fun map_consts f (Consts {decls, abbrevs, constraints}) = + make_consts (f (decls, abbrevs, constraints)); + +fun dest (Consts {decls = (space, decls), abbrevs = _, constraints}) = {constants = (space, Symtab.fold (fn (c, (LogicalConst (T, _), _)) => Symtab.update (c, (T, NONE)) | (c, (Abbreviation (T, t), _)) => Symtab.update (c, (T, SOME t))) decls Symtab.empty), constraints = (space, constraints)}; -fun space (Consts {decls, ...}) = #1 decls; +fun space_of (Consts {decls = (space, _), ...}) = space; +fun abbrevs_of (Consts {abbrevs, ...}) = abbrevs; (* lookup consts *) @@ -66,36 +72,45 @@ fun declaration consts c = #1 (logical_const consts c); fun monomorphic consts c = null (#2 (logical_const consts c)); -fun expansion tsig consts (c, T) = - (case the_const consts c of - Abbreviation a => SOME (Envir.expand_atom tsig T a) - | _ => NONE); - fun constraint (consts as Consts {constraints, ...}) c = (case Symtab.lookup constraints c of SOME T => T | NONE => (case the_const consts c of LogicalConst (T, _) => T | Abbreviation (T, _) => T)); -(* certify: check/expand consts -- without beta normalization *) - -fun certify pp tsig consts tm = - let - fun err msg = raise TYPE (msg, [], [tm]); - fun show_const (c, T) = quote c ^ " :: " ^ Pretty.string_of_typ pp T; +(* certify: check/expand consts *) - fun cert (t as Const (c, T)) = - (case the_const consts c of - LogicalConst (U, _) => - if Type.typ_instance tsig (T, U) then t - else err ("Illegal type for constant " ^ show_const (c, T)) - | Abbreviation (U, u) => - Envir.expand_atom tsig T (U, u) handle TYPE _ => - err ("Illegal type for abbreviation " ^ show_const (c, T))) - | cert (t $ u) = cert t $ cert u - | cert (Abs (x, T, t)) = Abs (x, T, cert t) - | cert a = a; - in cert tm end; +fun certify pp tsig consts = + let + fun err msg (c, T) = + raise TYPE (msg ^ " " ^ quote c ^ " :: " ^ Pretty.string_of_typ pp T, [], []); + fun cert tm = + let + val (head, args) = Term.strip_comb tm; + val args' = map cert args; + in + (case head of + Const (c, T) => + (case the_const consts c of + LogicalConst (U, _) => + if Type.typ_instance tsig (T, U) then Term.list_comb (head, args') + else err "Illegal type for constant" (c, T) + | Abbreviation (U, u) => + Term.betapplys (Envir.expand_atom tsig T (U, u) handle TYPE _ => + err "Illegal type for abbreviation" (c, T), args')) + | Abs (x, T, t) => Term.list_comb (Abs (x, T, cert t), args') + | _ => Term.list_comb (head, args')) + end; + in cert end; + + +(* read_const *) + +fun read_const consts raw_c = + let + val c = NameSpace.intern (space_of consts) raw_c; + val _ = the_const consts c handle TYPE (msg, _, _) => error msg; + in Const (c, dummyT) end; (* typargs -- view actual const type as instance of declaration *) @@ -128,13 +143,13 @@ (* name space *) -fun hide fully c = map_consts (fn (decls, constraints) => - (apfst (NameSpace.hide fully c) decls, constraints)); +fun hide fully c = map_consts (fn (decls, abbrevs, constraints) => + (apfst (NameSpace.hide fully c) decls, abbrevs, constraints)); (* declarations *) -fun declare naming (c, declT) = map_consts (fn (decls, constraints) => +fun declare naming (c, declT) = map_consts (fn (decls, abbrevs, constraints) => let fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos | args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos) @@ -142,37 +157,48 @@ and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is | args_of_list [] _ _ = I; val decl = (c, (LogicalConst (declT, map #2 (rev (args_of declT [] []))), serial ())); - in (extend_decls naming decl decls, constraints) end); + in (extend_decls naming decl decls, abbrevs, constraints) end); (* abbreviations *) -fun abbreviate pp tsig naming (c, rhs) consts = - consts |> map_consts (fn (decls, constraints) => +fun revert_abbrev const rhs = let - val rhs' = Envir.beta_norm (certify pp tsig consts rhs); - val decl = (c, (Abbreviation (Term.fastype_of rhs', rhs'), serial ())); - in (extend_decls naming decl decls, constraints) end); + val (xs, body) = Term.strip_abs (Envir.beta_eta_contract rhs); + val vars = fold (fn (x, T) => cons (Var ((x, 0), T))) (Term.rename_wrt_term body xs) []; + in (Term.list_comb (Const const, vars), Term.subst_bounds (rev vars, body)) end; + +fun abbreviate pp tsig naming revert (c, raw_rhs) consts = + consts |> map_consts (fn (decls, abbrevs, constraints) => + let + val rhs = certify pp tsig consts raw_rhs; + val T = Term.fastype_of rhs; + val decls' = decls |> extend_decls naming (c, (Abbreviation (T, rhs), serial ())); + val abbrevs' = + if revert then revert_abbrev (NameSpace.full naming c, T) rhs :: abbrevs else abbrevs; + in (decls', abbrevs', constraints) end); (* constraints *) -fun constrain (c, T) = - map_consts (fn (decls, constraints) => (decls, Symtab.update (c, T) constraints)); +fun constrain (c, T) = map_consts (fn (decls, abbrevs, constraints) => + (decls, abbrevs, Symtab.update (c, T) constraints)); (* empty and merge *) -val empty = make_consts (NameSpace.empty_table, Symtab.empty); +val empty = make_consts (NameSpace.empty_table, [], Symtab.empty); fun merge - (Consts {decls = decls1, constraints = constraints1}, - Consts {decls = decls2, constraints = constraints2}) = + (Consts {decls = decls1, abbrevs = abbrevs1, constraints = constraints1}, + Consts {decls = decls2, abbrevs = abbrevs2, constraints = constraints2}) = let val decls' = NameSpace.merge_tables (eq_snd (op =)) (decls1, decls2) handle Symtab.DUPS cs => err_dup_consts cs; + val abbrevs' = + Library.merge (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') (abbrevs1, abbrevs2); val constraints' = Symtab.merge (op =) (constraints1, constraints2) handle Symtab.DUPS cs => err_inconsistent_constraints cs; - in make_consts (decls', constraints') end; + in make_consts (decls', abbrevs', constraints') end; end;