# HG changeset patch # User wenzelm # Date 1150565872 -7200 # Node ID 300bc6ce970d53095819cdddd0ae388575868551 # Parent 2b4a222fef3c758a706e5462a7f41dd0d5d8dffe major reworking of export functionality -- based on Term/Thm.generalize; tuned interfaces; diff -r 2b4a222fef3c -r 300bc6ce970d src/Pure/variable.ML --- a/src/Pure/variable.ML Sat Jun 17 19:37:51 2006 +0200 +++ b/src/Pure/variable.ML Sat Jun 17 19:37:52 2006 +0200 @@ -14,7 +14,7 @@ val fixed_names_of: Context.proof -> string list val binds_of: Context.proof -> (typ * term) Vartab.table val defaults_of: Context.proof -> - typ Vartab.table * sort Vartab.table * string list * term list Symtab.table + typ Vartab.table * sort Vartab.table * string list * string list Symtab.table val used_types: Context.proof -> string list val is_declared: Context.proof -> string -> bool val is_fixed: Context.proof -> string -> bool @@ -24,23 +24,29 @@ val declare_type: typ -> Context.proof -> Context.proof val declare_syntax: term -> Context.proof -> Context.proof val declare_term: term -> Context.proof -> Context.proof + val rename_wrt: Context.proof -> term list -> (string * 'a) list -> (string * 'a) list + val add_fixes: string list -> Context.proof -> string list * Context.proof + val invent_fixes: string list -> Context.proof -> string list * Context.proof val invent_types: sort list -> Context.proof -> (string * sort) list * Context.proof - val rename_wrt: Context.proof -> term list -> (string * 'a) list -> (string * 'a) list + val export_inst: Context.proof -> Context.proof -> string list * string list + val exportT_inst: Context.proof -> Context.proof -> string list + val export_terms: Context.proof -> Context.proof -> term list -> term list + val exportT_terms: Context.proof -> Context.proof -> term list -> term list + val exportT: Context.proof -> Context.proof -> thm list -> thm list + val export: Context.proof -> Context.proof -> thm list -> thm list + val importT_inst: term list -> Context.proof -> ((indexname * sort) * typ) list * Context.proof + val import_inst: bool -> term list -> Context.proof -> + (((indexname * sort) * typ) list * ((indexname * typ) * term) list) * Context.proof + val importT_terms: term list -> Context.proof -> term list * Context.proof + val import_terms: bool -> term list -> Context.proof -> term list * Context.proof + val importT: thm list -> Context.proof -> thm list * Context.proof + val import: bool -> thm list -> Context.proof -> thm list * Context.proof val warn_extra_tfrees: Context.proof -> Context.proof -> unit - val generalize_tfrees: Context.proof -> Context.proof -> string list -> string list - val generalize: Context.proof -> Context.proof -> term list -> term list + val monomorphic: Context.proof -> term list -> term list val polymorphic: Context.proof -> term list -> term list val hidden_polymorphism: term -> typ -> (indexname * sort) list - val monomorphic_inst: term list -> Context.proof -> - ((indexname * sort) * typ) list * Context.proof - val monomorphic: Context.proof -> term list -> term list val add_binds: (indexname * term option) list -> Context.proof -> Context.proof val expand_binds: Context.proof -> term -> term - val add_fixes: string list -> Context.proof -> string list * Context.proof - val invent_fixes: string list -> Context.proof -> string list * Context.proof - val import_types: bool -> typ list -> Context.proof -> typ list * Context.proof - val import_terms: bool -> term list -> Context.proof -> term list * Context.proof - val import: bool -> thm list -> Context.proof -> thm list * Context.proof end; structure Variable: VARIABLE = @@ -49,14 +55,14 @@ (** local context data **) datatype data = Data of - {is_body: bool, (*internal body mode*) - fixes: (string * string) list, (*term fixes*) + {is_body: bool, (*inner body mode*) + fixes: (string * string) list, (*term fixes -- extern/intern*) binds: (typ * term) Vartab.table, (*term bindings*) defaults: typ Vartab.table * (*type constraints*) sort Vartab.table * (*default sorts*) string list * (*used type variables*) - term list Symtab.table}; (*type variable occurrences*) + string list Symtab.table}; (*occurrences of type variables in term variables*) fun make_data (is_body, fixes, binds, defaults) = Data {is_body = is_body, fixes = fixes, binds = binds, defaults = defaults}; @@ -141,7 +147,8 @@ (fn TFree (x, _) => insert (op =) x | _ => I); val ins_occs = fold_term_types (fn t => - fold_atyps (fn TFree (x, _) => Symtab.update_list (x, t) | _ => I)); + let val x = case t of Free (x, _) => x | _ => "" + in fold_atyps (fn TFree (a, _) => Symtab.insert_list (op =) (a, x) | _ => I) end); fun ins_skolem def_ty = fold_rev (fn (x, x') => (case def_ty x' of @@ -162,6 +169,9 @@ fold_types ins_used t used, occ)); +fun declare_occs t = map_defaults (fn (types, sorts, used, occ) => + (types, sorts, used, ins_occs t occ)); + fun declare_term t ctxt = ctxt |> declare_syntax t @@ -174,15 +184,6 @@ end; -(* invent types *) - -fun invent_types Ss ctxt = - let - val tfrees = Term.invent_names (used_types ctxt) "'a" (length Ss) ~~ Ss; - val ctxt' = fold (declare_type o TFree) tfrees ctxt; - in (tfrees, ctxt') end; - - (* renaming term/type frees *) fun rename_wrt ctxt ts frees = @@ -198,24 +199,142 @@ -(** Hindley-Milner polymorphism **) +(** fixes **) + +fun no_dups [] = () + | no_dups dups = error ("Duplicate fixed variable(s): " ^ commas_quote dups); + +fun add_fixes xs ctxt = + let + val (ys, zs) = split_list (fixes_of ctxt); + val _ = no_dups (duplicates (op =) xs); + val _ = + (case filter (can Term.dest_skolem) xs of [] => () + | bads => error ("Illegal internal Skolem constant(s): " ^ commas_quote bads)); + val xs' = + if is_body ctxt then Term.variantlist (map Term.skolem xs, zs) + else (no_dups (xs inter_string ys); no_dups (xs inter_string zs); xs); + in + ctxt + |> map_fixes (fn fixes => rev (xs ~~ xs') @ fixes) + |> fold (declare_syntax o Syntax.free) xs' + |> pair xs' + end; + +fun invent_fixes xs ctxt = + ctxt + |> set_body true + |> add_fixes (Term.variantlist (xs, [])) + ||> restore_body ctxt; + +fun invent_types Ss ctxt = + let + val tfrees = Term.invent_names (used_types ctxt) "'a" (length Ss) ~~ Ss; + val ctxt' = fold (declare_type o TFree) tfrees ctxt; + in (tfrees, ctxt') end; + + + +(** export -- generalize type/term variables **) + +fun export_inst inner outer = + let + val types_outer = used_types outer; + val fixes_inner = fixes_of inner; + val fixes_outer = fixes_of outer; + + val gen_fixes = map #2 (Library.take (length fixes_inner - length fixes_outer, fixes_inner)); + val still_fixed = not o member (op =) ("" :: gen_fixes); + val gen_fixesT = + Symtab.fold (fn (a, xs) => + if member (op =) types_outer a orelse exists still_fixed xs + then I else cons a) (type_occs_of inner) []; + in (gen_fixesT, gen_fixes) end; + +fun exportT_inst inner outer = #1 (export_inst inner outer); + +fun exportT_terms inner outer ts = + map (Term.generalize (exportT_inst (fold declare_occs ts inner) outer, []) + (fold (Term.fold_types Term.maxidx_typ) ts ~1 + 1)) ts; + +fun export_terms inner outer ts = + map (Term.generalize (export_inst (fold declare_occs ts inner) outer) + (fold Term.maxidx_term ts ~1 + 1)) ts; + +fun gen_export inst inner outer ths = + let + val maxidx = fold Thm.maxidx_thm ths ~1; + val inner' = fold (declare_occs o Thm.full_prop_of) ths inner; + in map (Thm.generalize (inst inner' outer) (maxidx + 1)) ths end; + +val exportT = gen_export (rpair [] oo exportT_inst); +val export = gen_export export_inst; + + + +(** import -- fix schematic type/term variables **) + +fun importT_inst ts ctxt = + let + val tvars = rev (fold Term.add_tvars ts []); + val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; + in (tvars ~~ map TFree tfrees, ctxt') end; + +fun import_inst is_open ts ctxt = + let + val (instT, ctxt') = importT_inst ts ctxt; + val vars = map (apsnd (Term.instantiateT instT)) (rev (fold Term.add_vars ts [])); + val ren = if is_open then I else Term.internal; + val (xs, ctxt'') = invent_fixes (map (ren o #1 o #1) vars) ctxt'; + val inst = vars ~~ map Free (xs ~~ map #2 vars); + in ((instT, inst), ctxt'') end; + +fun importT_terms ts ctxt = + let val (instT, ctxt') = importT_inst ts ctxt + in (map (Term.instantiate (instT, [])) ts, ctxt') end; + +fun import_terms is_open ts ctxt = + let val (inst, ctxt') = import_inst is_open ts ctxt + in (map (Term.instantiate inst) ts, ctxt') end; + +fun importT ths ctxt = + let + val thy = Context.theory_of_proof ctxt; + val certT = Thm.ctyp_of thy; + val (instT, ctxt') = importT_inst (map Thm.full_prop_of ths) ctxt; + val instT' = map (fn (v, T) => (certT (TVar v), certT T)) instT; + val ths' = map (Thm.instantiate (instT', [])) ths; + in (ths', ctxt') end; + +fun import is_open ths ctxt = + let + val thy = Context.theory_of_proof ctxt; + val cert = Thm.cterm_of thy; + val certT = Thm.ctyp_of thy; + val ((instT, inst), ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; + val instT' = map (fn (v, T) => (certT (TVar v), certT T)) instT; + val inst' = map (fn (v, t) => (cert (Var v), cert t)) inst; + val ths' = map (Thm.instantiate (instT', inst')) ths; + in (ths', ctxt') end; + + + +(** implicit polymorphism **) (* warn_extra_tfrees *) fun warn_extra_tfrees ctxt1 ctxt2 = let - fun occs_typ a (Type (_, Ts)) = exists (occs_typ a) Ts - | occs_typ a (TFree (b, _)) = a = b - | occs_typ _ (TVar _) = false; - fun occs_free a (Free (x, _)) = + fun occs_typ a = Term.exists_subtype (fn TFree (b, _) => a = b | _ => false); + fun occs_free _ "" = I + | occs_free a x = (case def_type ctxt1 false (x, ~1) of SOME T => if occs_typ a T then I else cons (a, x) - | NONE => cons (a, x)) - | occs_free _ _ = I; + | NONE => cons (a, x)); val occs1 = type_occs_of ctxt1 and occs2 = type_occs_of ctxt2; - val extras = Symtab.fold (fn (a, ts) => - if Symtab.defined occs1 a then I else fold (occs_free a) ts) occs2 []; + val extras = Symtab.fold (fn (a, xs) => + if Symtab.defined occs1 a then I else fold (occs_free a) xs) occs2 []; val tfrees = map #1 extras |> sort_distinct string_ord; val frees = map #2 extras |> sort_distinct string_ord; in @@ -225,29 +344,16 @@ end; -(* generalize type variables *) +(* monomorphic vs. polymorphic terms *) -fun generalize_tfrees inner outer = - let - val extra_fixes = subtract (op =) (fixed_names_of outer) (fixed_names_of inner); - fun still_fixed (Free (x, _)) = not (member (op =) extra_fixes x) - | still_fixed _ = false; - val occs_inner = type_occs_of inner; - val occs_outer = type_occs_of outer; - fun add a gen = - if Symtab.defined occs_outer a orelse - exists still_fixed (Symtab.lookup_list occs_inner a) - then gen else a :: gen; - in fn tfrees => fold add tfrees [] end; - -fun generalize inner outer ts = - let - val tfrees = generalize_tfrees inner outer (map #1 (fold Term.add_tfrees ts [])); - fun gen (x, S) = if member (op =) tfrees x then TVar ((x, 0), S) else TFree (x, S); - in map (Term.map_term_types (Term.map_type_tfree gen)) ts end; +fun monomorphic ctxt ts = + #1 (importT_terms ts (fold declare_term ts ctxt)); fun polymorphic ctxt ts = - generalize (fold declare_term ts ctxt) ctxt ts; + exportT_terms (fold declare_term ts ctxt) ctxt ts; + + +(** term bindings **) fun hidden_polymorphism t T = let @@ -256,22 +362,6 @@ (fn TVar v => if member (op =) tvarsT v then I else insert (op =) v | _ => I)) t []; in extra_tvars end; - -(* monomorphic -- fixes type variables *) - -fun monomorphic_inst ts ctxt = - let - val tvars = rev (fold Term.add_tvars ts []); - val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; - in (tvars ~~ map TFree tfrees, ctxt') end; - -fun monomorphic ctxt ts = - map (Term.instantiate (#1 (monomorphic_inst ts (fold declare_term ts ctxt)), [])) ts; - - - -(** term abbreviations **) - fun add_bind (xi, NONE) = map_binds (Vartab.delete_safe xi) | add_bind ((x, i), SOME t) = let @@ -293,65 +383,4 @@ | expand t = t; in Envir.beta_norm o Term.map_aterms expand end; - - -(** fixes **) - -fun no_dups [] = () - | no_dups dups = error ("Duplicate fixed variable(s): " ^ commas_quote dups); - -fun add_fixes xs ctxt = - let - val (ys, zs) = split_list (fixes_of ctxt); - val _ = no_dups (duplicates (op =) xs); - val _ = - (case filter (can Syntax.dest_skolem) xs of [] => () - | bads => error ("Illegal internal Skolem constant(s): " ^ commas_quote bads)); - val xs' = - if is_body ctxt then Term.variantlist (map Syntax.skolem xs, zs) - else (no_dups (xs inter_string ys); no_dups (xs inter_string zs); xs); - in - ctxt - |> map_fixes (fn fixes => rev (xs ~~ xs') @ fixes) - |> fold (declare_syntax o Syntax.free) xs' - |> pair xs' - end; - -fun invent_fixes xs ctxt = - ctxt - |> set_body true - |> add_fixes (Term.variantlist (xs, [])) - ||> restore_body ctxt; - - -(* import -- fixes schematic variables *) - -fun import_inst is_open ts ctxt = - let - val (instT, ctxt') = monomorphic_inst ts ctxt; - val vars = map (apsnd (Term.instantiateT instT)) (rev (fold Term.add_vars ts [])); - val ren = if is_open then I else Syntax.internal; - val (xs, ctxt'') = invent_fixes (map (ren o #1 o #1) vars) ctxt'; - val inst = vars ~~ map Free (xs ~~ map #2 vars); - in ((instT, inst), ctxt'') end; - -fun import_terms is_open ts ctxt = - let val (inst, ctxt') = import_inst is_open ts ctxt - in (map (Term.instantiate inst) ts, ctxt') end; - -fun import_types is_open Ts ctxt = - import_terms is_open (map Logic.mk_type Ts) ctxt - |>> map Logic.dest_type; - -fun import is_open ths ctxt = - let - val thy = Context.theory_of_proof ctxt; - val cert = Thm.cterm_of thy; - val certT = Thm.ctyp_of thy; - val ((instT, inst), ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; - val instT' = map (fn (v, T) => (certT (TVar v), certT T)) instT; - val inst' = map (fn (v, t) => (cert (Var v), cert t)) inst; - val ths' = map (Thm.instantiate (instT', inst')) ths; - in (ths', ctxt') end; - end;