# HG changeset patch # User wenzelm # Date 1139255942 -3600 # Node ID f22be3d61ed53416199bd5ed3b74d8da4cb37c9d # Parent 0342b7c2138887f78009f816d8ea69f7db1c3b06 added abbreviations; added certify (mostly from sign.ML); diff -r 0342b7c21388 -r f22be3d61ed5 src/Pure/consts.ML --- a/src/Pure/consts.ML Mon Feb 06 20:59:01 2006 +0100 +++ b/src/Pure/consts.ML Mon Feb 06 20:59:02 2006 +0100 @@ -2,20 +2,26 @@ ID: $Id$ Author: Makarius -Polymorphic constants. +Polymorphic constants: declarations, abbreviations, additional type +constraints. *) signature CONSTS = sig type T - val dest: T -> {declarations: typ NameSpace.table, constraints: typ NameSpace.table} + val dest: T -> + {constants: (typ * term option) NameSpace.table, + constraints: typ NameSpace.table} val space: T -> NameSpace.T - val declaration: T -> string -> typ (*exception TYPE*) - val constraint: T -> string -> typ (*exception TYPE*) - val monomorphic: T -> string -> bool + 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 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 constrain: string * typ -> T -> T val hide: bool -> string -> T -> T val empty: T @@ -25,40 +31,71 @@ structure Consts: CONSTS = struct - (** datatype T **) +datatype decl = + LogicalConst of typ * int list list | + Abbreviation of typ * term; + datatype T = Consts of - {declarations: ((typ * int list list) * serial) NameSpace.table, + {decls: (decl * serial) NameSpace.table, constraints: typ Symtab.table}; -fun make_consts (declarations, constraints) = - Consts {declarations = declarations, constraints = constraints}; +fun make_consts (decls, constraints) = Consts {decls = decls, constraints = constraints}; +fun map_consts f (Consts {decls, constraints}) = make_consts (f (decls, constraints)); -fun map_consts f (Consts {declarations, constraints}) = - make_consts (f (declarations, constraints)); - -fun dest (Consts {declarations = (space, decls), constraints}) = - {declarations = (space, Symtab.map (#1 o #1) decls), +fun dest (Consts {decls = (space, decls), 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 {declarations = (space, _), ...}) = space; +fun space (Consts {decls, ...}) = #1 decls; (* lookup consts *) -fun the_const (Consts {declarations = (_, decls), ...}) c = - (case Symtab.lookup decls c of SOME (decl, _) => decl - | NONE => raise TYPE ("Undeclared constant " ^ quote c, [], [])); +fun err_undeclared c = raise TYPE ("Undeclared constant: " ^ quote c, [], []); + +fun the_const (Consts {decls = (_, tab), ...}) c = + (case Symtab.lookup tab c of SOME (decl, _) => decl | NONE => err_undeclared c); + +fun logical_const consts c = + (case the_const consts c of LogicalConst d => d | _ => err_undeclared c); -fun declaration consts c = #1 (the_const consts c); +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 => declaration consts c); + | 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; -fun monomorphic consts c = null (#2 (the_const consts c)); + 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; (* typargs -- view actual const type as instance of declaration *) @@ -67,7 +104,7 @@ | sub T [] = T | sub T _ = raise Subscript; -fun typargs consts (c, T) = map (sub T) (#2 (the_const consts c)); +fun typargs consts (c, T) = map (sub T) (#2 (logical_const consts c)); fun instance consts (c, Ts) = let @@ -77,7 +114,7 @@ -(** build consts **) +(** declare consts **) fun err_dup_consts cs = error ("Duplicate declaration of constant(s) " ^ commas_quote cs); @@ -85,28 +122,43 @@ fun err_inconsistent_constraints cs = error ("Inconsistent type constraints for constant(s) " ^ commas_quote cs); +fun extend_decls naming decl tab = NameSpace.extend_table naming (tab, [decl]) + handle Symtab.DUPS cs => err_dup_consts cs; -(* declarations etc. *) + +(* name space *) -fun args_of declT = +fun hide fully c = map_consts (fn (decls, constraints) => + (apfst (NameSpace.hide fully c) decls, constraints)); + + +(* declarations *) + +fun declare naming (c, declT) = map_consts (fn (decls, constraints) => let - fun args (Type (_, Ts)) pos = args_list Ts 0 pos - | args (TVar v) pos = insert (eq_fst op =) (v, rev pos) - | args (TFree _) _ = I - and args_list (T :: Ts) i is = args T (i :: is) #> args_list Ts (i + 1) is - | args_list [] _ _ = I; - in map #2 (rev (args declT [] [])) end; + fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos + | args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos) + | args_of (TFree _) _ = I + 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); -fun declare naming (c, T) = map_consts (apfst (fn declarations => + +(* abbreviations *) + +fun abbreviate pp tsig naming (c, rhs) consts = + consts |> map_consts (fn (decls, constraints) => let - val decl = (c, ((T, args_of T), serial ())); - val declarations' = NameSpace.extend_table naming (declarations, [decl]) - handle Symtab.DUPS cs => err_dup_consts cs; - in declarations' end)); + 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 constrain = map_consts o apsnd o Symtab.update; +(* constraints *) -fun hide fully c = map_consts (apfst (apfst (NameSpace.hide fully c))); +fun constrain (c, T) = + map_consts (fn (decls, constraints) => (decls, Symtab.update (c, T) constraints)); (* empty and merge *) @@ -114,13 +166,13 @@ val empty = make_consts (NameSpace.empty_table, Symtab.empty); fun merge - (Consts {declarations = declarations1, constraints = constraints1}, - Consts {declarations = declarations2, constraints = constraints2}) = + (Consts {decls = decls1, constraints = constraints1}, + Consts {decls = decls2, constraints = constraints2}) = let - val declarations = NameSpace.merge_tables (eq_snd (op =)) (declarations1, declarations2) + val decls' = NameSpace.merge_tables (eq_snd (op =)) (decls1, decls2) handle Symtab.DUPS cs => err_dup_consts cs; - val constraints = Symtab.merge (op =) (constraints1, constraints2) + val constraints' = Symtab.merge (op =) (constraints1, constraints2) handle Symtab.DUPS cs => err_inconsistent_constraints cs; - in make_consts (declarations, constraints) end; + in make_consts (decls', constraints') end; end;