diff --git a/src/Pure/Isar/generic_target.ML b/src/Pure/Isar/generic_target.ML --- a/src/Pure/Isar/generic_target.ML +++ b/src/Pure/Isar/generic_target.ML @@ -1,475 +1,473 @@ (* Title: Pure/Isar/generic_target.ML Author: Makarius Author: Florian Haftmann, TU Muenchen Common target infrastructure. *) signature GENERIC_TARGET = sig (*auxiliary*) val export_abbrev: Proof.context -> (term -> term) -> term -> term * ((string * sort) list * (term list * term list)) val check_mixfix: Proof.context -> binding * (string * sort) list -> mixfix -> mixfix val check_mixfix_global: binding * bool -> mixfix -> mixfix (*background primitives*) val background_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val background_declaration: declaration -> local_theory -> local_theory val background_abbrev: binding * term -> term list -> local_theory -> (term * term) * local_theory val add_foundation_interpretation: (binding * (term * term list) -> Context.generic -> Context.generic) -> theory -> theory (*nested local theories primitives*) val standard_facts: local_theory -> Proof.context -> Attrib.fact list -> Attrib.fact list val standard_notes: (int * int -> bool) -> string -> Attrib.fact list -> local_theory -> local_theory val standard_declaration: (int * int -> bool) -> (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory val standard_const: (int * int -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val local_interpretation: Locale.registration -> local_theory -> local_theory (*lifting target primitives to local theory operations*) val define: (((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory) -> bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val notes: (string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory) -> string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: (Syntax.mode -> binding * mixfix -> term -> term list * term list -> local_theory -> local_theory) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory (*theory target primitives*) val theory_target_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val theory_target_notes: string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val theory_target_abbrev: Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory (*theory target operations*) val theory_abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val theory_declaration: declaration -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory (*locale target primitives*) val locale_target_notes: string -> string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val locale_target_abbrev: string -> Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory val locale_target_declaration: string -> bool -> declaration -> local_theory -> local_theory val locale_target_const: string -> (morphism -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory (*locale operations*) val locale_abbrev: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val locale_declaration: string -> {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory val locale_const: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val locale_dependency: string -> Locale.registration -> local_theory -> local_theory end structure Generic_Target: GENERIC_TARGET = struct (** consts **) fun export_abbrev lthy preprocess rhs = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); val rhs' = rhs |> Assumption.export_term lthy (Local_Theory.target_of lthy) |> preprocess; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) (Variable.add_fixed lthy rhs' [])); val u = fold_rev lambda term_params rhs'; val global_rhs = singleton (Variable.polymorphic thy_ctxt) u; val type_tfrees = Term_Subst.TFrees.build (Term_Subst.add_tfreesT (Term.fastype_of u)); val extra_tfrees = - (u, []) |-> (Term.fold_types o Term.fold_atyps) - (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I) - |> rev; + build_rev (u |> (Term.fold_types o Term.fold_atyps) + (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I)); val type_params = map (Logic.mk_type o TFree) extra_tfrees; in (global_rhs, (extra_tfrees, (type_params, term_params))) end; fun check_mixfix ctxt (b, extra_tfrees) mx = if null extra_tfrees then mx else (if Context_Position.is_visible ctxt then warning ("Additional type variable(s) in specification of " ^ Binding.print b ^ ": " ^ commas (map (Syntax.string_of_typ ctxt o TFree) (sort_by #1 extra_tfrees)) ^ (if Mixfix.is_empty mx then "" else "\nDropping mixfix syntax " ^ Pretty.string_of (Mixfix.pretty_mixfix mx))) else (); NoSyn); fun check_mixfix_global (b, no_params) mx = if no_params orelse Mixfix.is_empty mx then mx else (warning ("Dropping global mixfix syntax: " ^ Binding.print b ^ " " ^ Pretty.string_of (Mixfix.pretty_mixfix mx)); NoSyn); fun same_const (Const (c, _), Const (c', _)) = c = c' | same_const (t $ _, t' $ _) = same_const (t, t') | same_const (_, _) = false; fun const_decl phi_pred prmode ((b, mx), rhs) phi context = if phi_pred phi then let val b' = Morphism.binding phi b; val rhs' = Morphism.term phi rhs; val same_shape = Term.aconv_untyped (rhs, rhs'); val same_stem = same_shape orelse same_const (rhs, rhs'); val const_alias = if same_shape then (case rhs' of Const (c, T) => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; in (case Type_Infer_Context.const_type ctxt c of SOME T' => if Sign.typ_equiv thy (T, T') then SOME c else NONE | NONE => NONE) end | _ => NONE) else NONE; in (case const_alias of SOME c => context |> Context.mapping (Sign.const_alias b' c) (Proof_Context.const_alias b' c) |> Morphism.form (Proof_Context.generic_notation true prmode [(rhs', mx)]) | NONE => context |> Proof_Context.generic_add_abbrev Print_Mode.internal (b', Term.close_schematic_term rhs') |-> (fn (const as Const (c, _), _) => same_stem ? (Proof_Context.generic_revert_abbrev (#1 prmode) c #> same_shape ? Morphism.form (Proof_Context.generic_notation true prmode [(const, mx)])))) end else context; (** background primitives **) structure Foundation_Interpretations = Theory_Data ( type T = (binding * (term * term list) -> Context.generic -> Context.generic) Inttab.table; val empty = Inttab.empty; val extend = I; val merge = Inttab.merge (K true); ); fun add_foundation_interpretation f = Foundation_Interpretations.map (Inttab.update_new (serial (), f)); fun foundation_interpretation binding_const_params lthy = let val interps = Foundation_Interpretations.get (Proof_Context.theory_of lthy); val interp = Inttab.fold (fn (_, f) => f binding_const_params) interps; in lthy |> Local_Theory.background_theory (Context.theory_map interp) |> Local_Theory.map_contexts (K (Context.proof_map interp)) end; fun background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy = let val params = type_params @ term_params; val target_params = type_params @ take_prefix is_Free (Variable.export_terms lthy (Local_Theory.target_of lthy) term_params); val mx' = check_mixfix_global (b, null params) mx; val (const, lthy2) = lthy |> Local_Theory.background_theory_result (Sign.declare_const lthy ((b, U), mx')); val lhs = Term.list_comb (const, params); val ((_, def), lthy3) = lthy2 |> Local_Theory.background_theory_result (Thm.add_def (Proof_Context.defs_context lthy2) false false (Thm.def_binding_optional b b_def, Logic.mk_equals (lhs, rhs))) ||> foundation_interpretation (b, (const, target_params)); in ((lhs, def), lthy3) end; fun background_declaration decl lthy = let fun theory_decl context = Local_Theory.standard_form lthy (Proof_Context.init_global (Context.theory_of context)) decl context; in Local_Theory.background_theory (Context.theory_map theory_decl) lthy end; fun background_abbrev (b, global_rhs) params = Local_Theory.background_theory_result (Sign.add_abbrev Print_Mode.internal (b, global_rhs)) #>> apply2 (fn t => Term.list_comb (Logic.unvarify_global t, params)) (** nested local theories primitives **) fun standard_facts lthy ctxt = Attrib.transform_facts (Local_Theory.standard_morphism lthy ctxt); fun standard_notes pred kind facts lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Attrib.local_notes kind (standard_facts lthy ctxt facts) ctxt |> snd else ctxt) lthy; fun standard_declaration pred decl lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Context.proof_map (Local_Theory.standard_form lthy ctxt decl) ctxt else ctxt) lthy; fun standard_const pred prmode ((b, mx), rhs) = standard_declaration pred (const_decl (K true) prmode ((b, mx), rhs)); fun standard_registration pred registration lthy = Local_Theory.map_contexts (fn level => if pred (Local_Theory.level lthy, level) then Context.proof_map (Locale.add_registration registration) else I) lthy; val local_interpretation = standard_registration (fn (n, level) => level = n - 1); (** lifting target primitives to local theory operations **) (* define *) fun define foundation internal ((b, mx), ((b_def, atts), rhs)) lthy = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); (*term and type parameters*) val ((defs, _), rhs') = Thm.cterm_of lthy rhs |> Local_Defs.export_cterm lthy thy_ctxt ||> Thm.term_of; val xs = Variable.add_fixed lthy rhs' []; val T = Term.fastype_of rhs; val type_tfrees = Term_Subst.TFrees.build (Term_Subst.add_tfreesT T #> fold (Term_Subst.add_tfreesT o #2) xs); val extra_tfrees = - (rhs, []) |-> (Term.fold_types o Term.fold_atyps) - (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I) - |> rev; + build_rev (rhs |> (Term.fold_types o Term.fold_atyps) + (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I)); val mx' = check_mixfix lthy (b, extra_tfrees) mx; val type_params = map (Logic.mk_type o TFree) extra_tfrees; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) xs); val params = type_params @ term_params; val U = map Term.fastype_of params ---> T; (*foundation*) val ((lhs', global_def), lthy2) = lthy |> foundation (((b, U), mx'), (b_def, rhs')) (type_params, term_params); (*local definition*) val ([(lhs, (_, local_def))], lthy3) = lthy2 |> Context_Position.set_visible false |> Local_Defs.define [((b, NoSyn), (Binding.empty_atts, lhs'))] ||> Context_Position.restore_visible lthy2; (*result*) val def = Thm.transitive local_def global_def |> Local_Defs.contract lthy3 defs (Thm.cterm_of lthy3 (Logic.mk_equals (lhs, rhs))); val ([(res_name, [res])], lthy4) = lthy3 |> Local_Theory.notes [((if internal then Binding.empty else b_def, atts), [([def], [])])]; in ((lhs, (res_name, res)), lthy4) end; (* notes *) local fun import_export_proof ctxt (name, raw_th) = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of ctxt); (*export assumes/defines*) val th = Goal.norm_result ctxt raw_th; val ((defs, asms), th') = Local_Defs.export ctxt thy_ctxt th; val asms' = map (rewrite_rule ctxt (Drule.norm_hhf_eqs @ defs)) asms; (*export fixes*) val tfrees = map TFree (Thm.fold_terms Term.add_tfrees th' []); val frees = map Free (Thm.fold_terms Term.add_frees th' []); val (th'' :: vs) = (th' :: map (Drule.mk_term o Thm.cterm_of ctxt) (map Logic.mk_type tfrees @ frees)) |> Variable.export ctxt thy_ctxt |> Drule.zero_var_indexes_list; (*thm definition*) val result = Global_Theory.name_thm Global_Theory.official1 name th''; (*import fixes*) val (tvars, vars) = chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs) |>> map Logic.dest_type; val instT = Term_Subst.TVars.build (fold2 (fn a => fn b => (case a of TVar v => Term_Subst.TVars.add (v, b) | _ => I)) tvars tfrees); val cinstT = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt T)) instT []; val cinst = map_filter (fn (Var (xi, T), t) => SOME ((xi, Term_Subst.instantiateT instT T), Thm.cterm_of ctxt (Term.map_types (Term_Subst.instantiateT instT) t)) | _ => NONE) (vars ~~ frees); val result' = Thm.instantiate (cinstT, cinst) result; (*import assumes/defines*) val result'' = (fold (curry op COMP) asms' result' handle THM _ => raise THM ("Failed to re-import result", 0, result' :: asms')) |> Local_Defs.contract ctxt defs (Thm.cprop_of th) |> Goal.norm_result ctxt |> Global_Theory.name_thm Global_Theory.unofficial2 name; in (result'', result) end; fun bind_name lthy b = (Local_Theory.full_name lthy b, Binding.default_pos_of b); fun map_facts f = map (apsnd (map (apfst (map f)))); in fun notes target_notes kind facts lthy = let val facts' = facts |> map (fn (a, bs) => (a, Global_Theory.burrow_fact (Global_Theory.name_multi (bind_name lthy (fst a))) bs)) |> map_facts (import_export_proof lthy); val local_facts = map_facts #1 facts'; val global_facts = map_facts #2 facts'; in lthy |> target_notes kind global_facts (Attrib.partial_evaluation lthy local_facts) |> Attrib.local_notes kind local_facts end; end; (* abbrev *) fun abbrev target_abbrev prmode ((b, mx), rhs) lthy = let val (global_rhs, (extra_tfrees, (type_params, term_params))) = export_abbrev lthy I rhs; val mx' = check_mixfix lthy (b, extra_tfrees) mx; in lthy |> target_abbrev prmode (b, mx') global_rhs (type_params, term_params) |> Context_Position.set_visible false |> Proof_Context.add_abbrev Print_Mode.internal (b, rhs) |> snd |> Local_Defs.fixed_abbrev ((b, NoSyn), rhs) ||> Context_Position.restore_visible lthy end; (** theory target primitives **) fun theory_target_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) = background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) #-> (fn (lhs, def) => standard_const (op <>) Syntax.mode_default ((b, mx), lhs) #> pair (lhs, def)); fun theory_target_notes kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind global_facts #> snd) #> standard_notes (op <>) kind local_facts; fun theory_target_abbrev prmode (b, mx) global_rhs params = Local_Theory.background_theory_result (Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> (fn (lhs, _) => (* FIXME type_params!? *) Sign.notation true prmode [(lhs, check_mixfix_global (b, null (snd params)) mx)] #> pair lhs)) #-> (fn lhs => standard_const (op <>) prmode ((b, if null (snd params) then NoSyn else mx), Term.list_comb (Logic.unvarify_global lhs, snd params))); (** theory operations **) val theory_abbrev = abbrev theory_target_abbrev; fun theory_declaration decl = background_declaration decl #> standard_declaration (K true) decl; fun target_registration lthy {inst, mixin, export} = {inst = inst, mixin = mixin, export = export $> Proof_Context.export_morphism lthy (Local_Theory.target_of lthy)}; fun theory_registration registration lthy = lthy |> (Local_Theory.raw_theory o Context.theory_map) (Locale.add_registration (target_registration lthy registration)) |> standard_registration (K true) registration; (** locale target primitives **) fun locale_target_notes locale kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind (Attrib.map_facts (K []) global_facts) #> snd) #> (fn lthy => lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_facts locale kind (standard_facts lthy ctxt local_facts))) #> standard_notes (fn (this, other) => other <> 0 andalso this <> other) kind local_facts; fun locale_target_declaration locale syntax decl lthy = lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_declaration locale syntax (Morphism.transform (Local_Theory.standard_morphism lthy ctxt) decl)); fun locale_target_const locale phi_pred prmode ((b, mx), rhs) = locale_target_declaration locale true (const_decl phi_pred prmode ((b, mx), rhs)) (** locale operations **) fun locale_declaration locale {syntax, pervasive} decl = pervasive ? background_declaration decl #> locale_target_declaration locale syntax decl #> standard_declaration (fn (_, other) => other <> 0) decl; fun locale_const locale prmode ((b, mx), rhs) = locale_target_const locale (K true) prmode ((b, mx), rhs) #> standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), rhs); fun locale_dependency loc registration lthy = lthy |> Local_Theory.raw_theory (Locale.add_dependency loc registration) |> standard_registration (K true) registration; (** locale abbreviations **) fun locale_target_abbrev locale prmode (b, mx) global_rhs params = background_abbrev (b, global_rhs) (snd params) #-> (fn (lhs, _) => locale_const locale prmode ((b, mx), lhs)); fun locale_abbrev locale = abbrev (locale_target_abbrev locale); end; diff --git a/src/Pure/System/options.ML b/src/Pure/System/options.ML --- a/src/Pure/System/options.ML +++ b/src/Pure/System/options.ML @@ -1,223 +1,223 @@ (* Title: Pure/System/options.ML Author: Makarius System options with external string representation. *) signature OPTIONS = sig val boolT: string val intT: string val realT: string val stringT: string val unknownT: string type T val empty: T val dest: T -> (string * Position.T) list val typ: T -> string -> string val bool: T -> string -> bool val int: T -> string -> int val real: T -> string -> real val seconds: T -> string -> Time.time val string: T -> string -> string val put_bool: string -> bool -> T -> T val put_int: string -> int -> T -> T val put_real: string -> real -> T -> T val put_string: string -> string -> T -> T val declare: {pos: Position.T, name: string, typ: string, value: string} -> T -> T val update: string -> string -> T -> T val encode: T XML.Encode.T val decode: T XML.Decode.T val default: unit -> T val default_typ: string -> string val default_bool: string -> bool val default_int: string -> int val default_real: string -> real val default_seconds: string -> Time.time val default_string: string -> string val default_put_bool: string -> bool -> unit val default_put_int: string -> int -> unit val default_put_real: string -> real -> unit val default_put_string: string -> string -> unit val get_default: string -> string val put_default: string -> string -> unit val set_default: T -> unit val reset_default: unit -> unit val load_default: unit -> unit end; structure Options: OPTIONS = struct (* representation *) val boolT = "bool"; val intT = "int"; val realT = "real"; val stringT = "string"; val unknownT = "unknown"; datatype T = Options of {pos: Position.T, typ: string, value: string} Symtab.table; val empty = Options Symtab.empty; fun dest (Options tab) = Symtab.fold (fn (name, {pos, ...}) => cons (name, pos)) tab [] |> sort_by #1; (* check *) fun check_name (Options tab) name = let val opt = Symtab.lookup tab name in if is_some opt andalso #typ (the opt) <> unknownT then the opt else error ("Unknown system option " ^ quote name) end; fun check_type options name typ = let val opt = check_name options name in if #typ opt = typ then opt else error ("Ill-typed system option " ^ quote name ^ " : " ^ #typ opt ^ " vs. " ^ typ) end; (* typ *) fun typ options name = #typ (check_name options name); (* basic operations *) fun put T print name x (options as Options tab) = let val opt = check_type options name T in Options (Symtab.update (name, {pos = #pos opt, typ = #typ opt, value = print x}) tab) end; fun get T parse options name = let val opt = check_type options name T in (case parse (#value opt) of SOME x => x | NONE => error ("Malformed value for system option " ^ quote name ^ " : " ^ T ^ " =\n" ^ quote (#value opt))) end; (* internal lookup and update *) val bool = get boolT (try Value.parse_bool); val int = get intT (try Value.parse_int); val real = get realT (try Value.parse_real); val seconds = Time.fromReal oo real; val string = get stringT SOME; val put_bool = put boolT Value.print_bool; val put_int = put intT Value.print_int; val put_real = put realT Value.print_real; val put_string = put stringT I; (* external updates *) fun check_value options name = let val opt = check_name options name in if #typ opt = boolT then ignore (bool options name) else if #typ opt = intT then ignore (int options name) else if #typ opt = realT then ignore (real options name) else if #typ opt = stringT then ignore (string options name) else () end; fun declare {pos, name, typ, value} (Options tab) = let val options' = (case Symtab.lookup tab name of SOME other => error ("Duplicate declaration of system option " ^ quote name ^ Position.here pos ^ Position.here (#pos other)) | NONE => Options (Symtab.update (name, {pos = pos, typ = typ, value = value}) tab)); val _ = typ = boolT orelse typ = intT orelse typ = realT orelse typ = stringT orelse error ("Unknown type for system option " ^ quote name ^ " : " ^ quote typ ^ Position.here pos); val _ = check_value options' name; in options' end; fun update name value (options as Options tab) = let val opt = check_name options name; val options' = Options (Symtab.update (name, {pos = #pos opt, typ = #typ opt, value = value}) tab); val _ = check_value options' name; in options' end; (* XML data *) fun encode (Options tab) = let val opts = - (tab, []) |-> Symtab.fold (fn (name, {pos, typ, value}) => - cons (Position.properties_of pos, (name, (typ, value)))); + build (tab |> Symtab.fold (fn (name, {pos, typ, value}) => + cons (Position.properties_of pos, (name, (typ, value))))); open XML.Encode; in list (pair properties (pair string (pair string string))) opts end; fun decode body = let open XML.Decode; val decode_options = list (pair properties (pair string (pair string string))) #> map (fn (props, (name, (typ, value))) => {pos = Position.of_properties props, name = name, typ = typ, value = value}); in fold declare (decode_options body) empty end; (** global default **) val global_default = Synchronized.var "Options.default" (NONE: T option); fun err_no_default () = error "Missing default for system options within Isabelle process"; fun change_default f x y = Synchronized.change global_default (fn SOME options => SOME (f x y options) | NONE => err_no_default ()); fun default () = (case Synchronized.value global_default of SOME options => options | NONE => err_no_default ()); fun default_typ name = typ (default ()) name; fun default_bool name = bool (default ()) name; fun default_int name = int (default ()) name; fun default_real name = real (default ()) name; fun default_seconds name = seconds (default ()) name; fun default_string name = string (default ()) name; val default_put_bool = change_default put_bool; val default_put_int = change_default put_int; val default_put_real = change_default put_real; val default_put_string = change_default put_string; fun get_default name = let val options = default () in get (typ options name) SOME options name end; val put_default = change_default update; fun set_default options = Synchronized.change global_default (K (SOME options)); fun reset_default () = Synchronized.change global_default (K NONE); fun load_default () = (case getenv "ISABELLE_PROCESS_OPTIONS" of "" => () | name => let val path = Path.explode name in (case try File.read path of SOME s => set_default (decode (YXML.parse_body s)) | NONE => ()) end); val _ = load_default (); val _ = ML_Print_Depth.set_print_depth (default_int "ML_print_depth"); end; diff --git a/src/Pure/Thy/export_theory.ML b/src/Pure/Thy/export_theory.ML --- a/src/Pure/Thy/export_theory.ML +++ b/src/Pure/Thy/export_theory.ML @@ -1,438 +1,438 @@ (* Title: Pure/Thy/export_theory.ML Author: Makarius Export foundational theory content and locale/class structure. *) signature EXPORT_THEORY = sig val export_enabled: Thy_Info.presentation_context -> bool val export_body: theory -> string -> XML.body -> unit end; structure Export_Theory: EXPORT_THEORY = struct (* approximative syntax *) val get_syntax = Syntax.get_approx o Proof_Context.syn_of; fun get_syntax_type ctxt = get_syntax ctxt o Lexicon.mark_type; fun get_syntax_const ctxt = get_syntax ctxt o Lexicon.mark_const; fun get_syntax_fixed ctxt = get_syntax ctxt o Lexicon.mark_fixed; fun get_syntax_param ctxt loc x = let val thy = Proof_Context.theory_of ctxt in if Class.is_class thy loc then (case AList.lookup (op =) (Class.these_params thy [loc]) x of NONE => NONE | SOME (_, (c, _)) => get_syntax_const ctxt c) else get_syntax_fixed ctxt x end; val encode_syntax = XML.Encode.variant [fn NONE => ([], []), fn SOME (Syntax.Prefix delim) => ([delim], []), fn SOME (Syntax.Infix {assoc, delim, pri}) => let val ass = (case assoc of Printer.No_Assoc => 0 | Printer.Left_Assoc => 1 | Printer.Right_Assoc => 2); open XML.Encode Term_XML.Encode; in ([], triple int string int (ass, delim, pri)) end]; (* free variables: not declared in the context *) val is_free = not oo Name.is_declared; fun add_frees used = fold_aterms (fn Free (x, T) => is_free used x ? insert (op =) (x, T) | _ => I); fun add_tfrees used = (fold_types o fold_atyps) (fn TFree (a, S) => is_free used a ? insert (op =) (a, S) | _ => I); (* locales *) fun locale_content thy loc = let val ctxt = Locale.init loc thy; val args = Locale.params_of thy loc |> map (fn ((x, T), _) => ((x, T), get_syntax_param ctxt loc x)); val axioms = let val (asm, defs) = Locale.specification_of thy loc; val cprops = map (Thm.cterm_of ctxt) (the_list asm @ defs); val (intro1, intro2) = Locale.intros_of thy loc; val intros_tac = Method.try_intros_tac ctxt (the_list intro1 @ the_list intro2) []; val res = Goal.init (Conjunction.mk_conjunction_balanced cprops) |> (ALLGOALS Goal.conjunction_tac THEN intros_tac) |> try Seq.hd; in (case res of SOME goal => Thm.prems_of goal | NONE => raise Fail ("Cannot unfold locale " ^ quote loc)) end; val typargs = rev (fold Term.add_tfrees (map (Free o #1) args @ axioms) []); in {typargs = typargs, args = args, axioms = axioms} end; fun get_locales thy = Locale.get_locales thy |> map_filter (fn loc => if Experiment.is_experiment thy loc then NONE else SOME (loc, ())); fun get_dependencies prev_thys thy = Locale.dest_dependencies prev_thys thy |> map_filter (fn dep => if Experiment.is_experiment thy (#source dep) orelse Experiment.is_experiment thy (#target dep) then NONE else let val (type_params, params) = Locale.parameters_of thy (#source dep); val typargs = fold (Term.add_tfreesT o #2 o #1) params type_params; val substT = typargs |> map_filter (fn v => let val T = TFree v; val T' = Morphism.typ (#morphism dep) T; in if T = T' then NONE else SOME (v, T') end); val subst = params |> map_filter (fn (v, _) => let val t = Free v; val t' = Morphism.term (#morphism dep) t; in if t aconv t' then NONE else SOME (v, t') end); in SOME (dep, (substT, subst)) end); (* presentation *) fun export_enabled (context: Thy_Info.presentation_context) = Options.bool (#options context) "export_theory"; fun export_body thy name body = if XML.is_empty_body body then () else Export.export thy (Path.binding0 (Path.make ["theory", name])) body; val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => let val rep_tsig = Type.rep_tsig (Sign.tsig_of thy); val consts = Sign.consts_of thy; val thy_ctxt = Proof_Context.init_global thy; val pos_properties = Thy_Info.adjust_pos_properties context; val enabled = export_enabled context; (* strict parents *) val parents = Theory.parents_of thy; val _ = Export.export thy \<^path_binding>\theory/parents\ (XML.Encode.string (cat_lines (map Context.theory_long_name parents))); (* spec rules *) fun spec_rule_content {pos, name, rough_classification, terms, rules} = let val spec = terms @ map Thm.plain_prop_of rules |> Term_Subst.zero_var_indexes_list |> map Logic.unvarify_global; in {props = pos_properties pos, name = name, rough_classification = rough_classification, typargs = rev (fold Term.add_tfrees spec []), args = rev (fold Term.add_frees spec []), terms = map (fn t => (t, Term.type_of t)) (take (length terms) spec), rules = drop (length terms) spec} end; (* entities *) fun make_entity_markup name xname pos serial = let val props = pos_properties pos @ Markup.serial_properties serial; in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end; fun entity_markup space name = let val xname = Name_Space.extern_shortest thy_ctxt space name; val {serial, pos, ...} = Name_Space.the_entry space name; in make_entity_markup name xname pos serial end; fun export_entities export_name get_space decls export = let val parent_spaces = map get_space parents; val space = get_space thy; in - (decls, []) |-> fold (fn (name, decl) => + build (decls |> fold (fn (name, decl) => if exists (fn space => Name_Space.declared space name) parent_spaces then I else (case export name decl of NONE => I | SOME make_body => let val i = #serial (Name_Space.the_entry space name); val body = if enabled then make_body () else []; - in cons (i, XML.Elem (entity_markup space name, body)) end)) + in cons (i, XML.Elem (entity_markup space name, body)) end))) |> sort (int_ord o apply2 #1) |> map #2 |> export_body thy export_name end; (* types *) val encode_type = let open XML.Encode Term_XML.Encode in triple encode_syntax (list string) (option typ) end; val _ = export_entities "types" Sign.type_space (Name_Space.dest_table (#types rep_tsig)) (fn c => (fn Type.LogicalType n => SOME (fn () => encode_type (get_syntax_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE)) | Type.Abbreviation (args, U, false) => SOME (fn () => encode_type (get_syntax_type thy_ctxt c, args, SOME U)) | _ => NONE)); (* consts *) val encode_term = Term_XML.Encode.term consts; val encode_const = let open XML.Encode Term_XML.Encode in pair encode_syntax (pair (list string) (pair typ (pair (option encode_term) bool))) end; val _ = export_entities "consts" Sign.const_space (#constants (Consts.dest consts)) (fn c => fn (T, abbrev) => SOME (fn () => let val syntax = get_syntax_const thy_ctxt c; val U = Logic.unvarifyT_global T; val U0 = Type.strip_sorts U; val trim_abbrev = Proofterm.standard_vars_term Name.context #> map_types Type.strip_sorts; val abbrev' = Option.map trim_abbrev abbrev; val args = map (#1 o dest_TFree) (Consts.typargs consts (c, U0)); val propositional = Object_Logic.is_propositional thy_ctxt (Term.body_type U0); in encode_const (syntax, (args, (U0, (abbrev', propositional)))) end)); (* axioms *) fun standard_prop used extra_shyps raw_prop raw_proof = let val (prop, proof) = Proofterm.standard_vars used (raw_prop, raw_proof); val args = rev (add_frees used prop []); val typargs = rev (add_tfrees used prop []); val used_typargs = fold (Name.declare o #1) typargs used; val sorts = Name.invent used_typargs Name.aT (length extra_shyps) ~~ extra_shyps; in ((sorts @ typargs, args, prop), proof) end; fun standard_prop_of thm = standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm); val encode_prop = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair string typ)) encode_term end; fun encode_axiom used prop = encode_prop (#1 (standard_prop used [] prop NONE)); val _ = export_entities "axioms" Theory.axiom_space (Theory.all_axioms_of thy) (fn _ => fn prop => SOME (fn () => encode_axiom Name.context prop)); (* theorems and proof terms *) val clean_thm = Thm.check_hyps (Context.Theory thy) #> Thm.strip_shyps; val prep_thm = clean_thm #> Thm.unconstrainT #> Thm.strip_shyps; val lookup_thm_id = Global_Theory.lookup_thm_id thy; fun expand_name thm_id (header: Proofterm.thm_header) = if #serial header = #serial thm_id then "" else (case lookup_thm_id (Proofterm.thm_header_id header) of NONE => "" | SOME thm_name => Thm_Name.print thm_name); fun entity_markup_thm (serial, (name, i)) = let val space = Facts.space_of (Global_Theory.facts_of thy); val xname = Name_Space.extern_shortest thy_ctxt space name; val {pos, ...} = Name_Space.the_entry space name; in make_entity_markup (Thm_Name.print (name, i)) (Thm_Name.print (xname, i)) pos serial end; fun encode_thm thm_id raw_thm = let val deps = map (Thm_Name.print o #2) (Thm_Deps.thm_deps thy [raw_thm]); val thm = prep_thm raw_thm; val proof0 = if Proofterm.export_standard_enabled () then Proof_Syntax.standard_proof_of {full = true, expand_name = SOME o expand_name thm_id} thm else if Proofterm.export_enabled () then Thm.reconstruct_proof_of thm else MinProof; val (prop, SOME proof) = standard_prop_of thm (SOME proof0); val _ = Thm.expose_proofs thy [thm]; in (prop, deps, proof) |> let open XML.Encode Term_XML.Encode; val encode_proof = Proofterm.encode_standard_proof consts; in triple encode_prop (list string) encode_proof end end; fun export_thm (thm_id, thm_name) = let val markup = entity_markup_thm (#serial thm_id, thm_name); val body = if enabled then Global_Theory.get_thm_name thy (thm_name, Position.none) |> encode_thm thm_id else []; in XML.Elem (markup, body) end; val _ = export_body thy "thms" (map export_thm (Global_Theory.dest_thm_names thy)); (* type classes *) val encode_class = let open XML.Encode Term_XML.Encode in pair (list (pair string typ)) (list (encode_axiom Name.context)) end; val _ = export_entities "classes" Sign.class_space (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig))))) (fn name => fn () => SOME (fn () => (case try (Axclass.get_info thy) name of NONE => ([], []) | SOME {params, axioms, ...} => (params, map (Thm.plain_prop_of o clean_thm) axioms)) |> encode_class)); (* sort algebra *) val _ = if enabled then let val prop = encode_axiom Name.context o Logic.varify_global; val encode_classrel = let open XML.Encode in list (pair prop (pair string string)) end; val encode_arities = let open XML.Encode Term_XML.Encode in list (pair prop (triple string (list sort) string)) end; val export_classrel = maps (fn (c, cs) => map (pair c) cs) #> map (`Logic.mk_classrel) #> encode_classrel; val export_arities = map (`Logic.mk_arity) #> encode_arities; val {classrel, arities} = Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents) (#2 (#classes rep_tsig)); in if null classrel then () else export_body thy "classrel" (export_classrel classrel); if null arities then () else export_body thy "arities" (export_arities arities) end else (); (* locales *) fun encode_locale used = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair (pair string typ) encode_syntax)) (list (encode_axiom used)) end; val _ = export_entities "locales" Locale.locale_space (get_locales thy) (fn loc => fn () => SOME (fn () => let val {typargs, args, axioms} = locale_content thy loc; val used = fold Name.declare (map #1 typargs @ map (#1 o #1) args) Name.context; in encode_locale used (typargs, args, axioms) end handle ERROR msg => cat_error msg ("The error(s) above occurred in locale " ^ quote (Locale.markup_name thy_ctxt loc)))); (* locale dependencies *) fun encode_locale_dependency (dep: Locale.locale_dependency, subst) = (#source dep, (#target dep, (#prefix dep, subst))) |> let open XML.Encode Term_XML.Encode; val encode_subst = pair (list (pair (pair string sort) typ)) (list (pair (pair string typ) (term consts))); in pair string (pair string (pair (list (pair string bool)) encode_subst)) end; val _ = if enabled then get_dependencies parents thy |> map_index (fn (i, dep) => let val xname = string_of_int (i + 1); val name = Long_Name.implode [Context.theory_name thy, xname]; val markup = make_entity_markup name xname (#pos (#1 dep)) (#serial (#1 dep)); val body = encode_locale_dependency dep; in XML.Elem (markup, body) end) |> export_body thy "locale_dependencies" else (); (* constdefs *) val _ = if enabled then let val constdefs = Defs.dest_constdefs (map Theory.defs_of (Theory.parents_of thy)) (Theory.defs_of thy) |> sort_by #1; val encode = let open XML.Encode in list (pair string string) end; in if null constdefs then () else export_body thy "constdefs" (encode constdefs) end else (); (* spec rules *) val encode_specs = let open XML.Encode Term_XML.Encode in list (fn {props, name, rough_classification, typargs, args, terms, rules} => pair properties (pair string (pair Spec_Rules.encode_rough_classification (pair (list (pair string sort)) (pair (list (pair string typ)) (pair (list (pair encode_term typ)) (list encode_term)))))) (props, (name, (rough_classification, (typargs, (args, (terms, rules))))))) end; val _ = if enabled then (case Spec_Rules.dest_theory thy of [] => () | spec_rules => export_body thy "spec_rules" (encode_specs (map spec_rule_content spec_rules))) else (); in () end); end; diff --git a/src/Pure/context.ML b/src/Pure/context.ML --- a/src/Pure/context.ML +++ b/src/Pure/context.ML @@ -1,768 +1,768 @@ (* Title: Pure/context.ML Author: Markus Wenzel, TU Muenchen Generic theory contexts with unique identity, arbitrarily typed data, monotonic development graph and history support. Generic proof contexts with arbitrarily typed data. Firm naming conventions: thy, thy', thy1, thy2: theory ctxt, ctxt', ctxt1, ctxt2: Proof.context context: Context.generic *) signature BASIC_CONTEXT = sig type theory exception THEORY of string * theory list structure Proof: sig type context end structure Proof_Context: sig val theory_of: Proof.context -> theory val init_global: theory -> Proof.context val get_global: theory -> string -> Proof.context end end; signature CONTEXT = sig include BASIC_CONTEXT (*theory context*) type theory_id val theory_id: theory -> theory_id val timing: bool Unsynchronized.ref val parents_of: theory -> theory list val ancestors_of: theory -> theory list val theory_id_ord: theory_id ord val theory_id_long_name: theory_id -> string val theory_id_name: theory_id -> string val theory_long_name: theory -> string val theory_name: theory -> string val theory_name': {long: bool} -> theory -> string val PureN: string val pretty_thy: theory -> Pretty.T val pretty_abbrev_thy: theory -> Pretty.T val get_theory: {long: bool} -> theory -> string -> theory val eq_thy_id: theory_id * theory_id -> bool val eq_thy: theory * theory -> bool val proper_subthy_id: theory_id * theory_id -> bool val proper_subthy: theory * theory -> bool val subthy_id: theory_id * theory_id -> bool val subthy: theory * theory -> bool val trace_theories: bool Unsynchronized.ref val theories_trace: unit -> {active_positions: Position.T list, active: int, total: int} val join_thys: theory * theory -> theory val begin_thy: string -> theory list -> theory val finish_thy: theory -> theory val theory_data_size: theory -> (Position.T * int) list (*proof context*) val raw_transfer: theory -> Proof.context -> Proof.context (*certificate*) datatype certificate = Certificate of theory | Certificate_Id of theory_id val certificate_theory: certificate -> theory val certificate_theory_id: certificate -> theory_id val eq_certificate: certificate * certificate -> bool val join_certificate: certificate * certificate -> certificate (*generic context*) datatype generic = Theory of theory | Proof of Proof.context val cases: (theory -> 'a) -> (Proof.context -> 'a) -> generic -> 'a val mapping: (theory -> theory) -> (Proof.context -> Proof.context) -> generic -> generic val mapping_result: (theory -> 'a * theory) -> (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic val the_theory: generic -> theory val the_proof: generic -> Proof.context val map_theory: (theory -> theory) -> generic -> generic val map_proof: (Proof.context -> Proof.context) -> generic -> generic val map_theory_result: (theory -> 'a * theory) -> generic -> 'a * generic val map_proof_result: (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic val theory_map: (generic -> generic) -> theory -> theory val proof_map: (generic -> generic) -> Proof.context -> Proof.context val theory_of: generic -> theory (*total*) val proof_of: generic -> Proof.context (*total*) (*thread data*) val get_generic_context: unit -> generic option val put_generic_context: generic option -> unit val setmp_generic_context: generic option -> ('a -> 'b) -> 'a -> 'b val the_generic_context: unit -> generic val the_global_context: unit -> theory val the_local_context: unit -> Proof.context val >> : (generic -> generic) -> unit val >>> : (generic -> 'a * generic) -> 'a end; signature PRIVATE_CONTEXT = sig include CONTEXT structure Theory_Data: sig val declare: Position.T -> Any.T -> (Any.T -> Any.T) -> (theory * theory -> Any.T * Any.T -> Any.T) -> serial val get: serial -> (Any.T -> 'a) -> theory -> 'a val put: serial -> ('a -> Any.T) -> 'a -> theory -> theory end structure Proof_Data: sig val declare: (theory -> Any.T) -> serial val get: serial -> (Any.T -> 'a) -> Proof.context -> 'a val put: serial -> ('a -> Any.T) -> 'a -> Proof.context -> Proof.context end end; structure Context: PRIVATE_CONTEXT = struct (*** theory context ***) (*private copy avoids potential conflict of table exceptions*) structure Datatab = Table(type key = int val ord = int_ord); (** datatype theory **) type stage = int * int Synchronized.var; fun init_stage () : stage = (0, Synchronized.var "Context.stage" 1); fun next_stage ((_, state): stage) = (Synchronized.change_result state (fn n => (n, n + 1)), state); abstype theory_id = Theory_Id of (*identity*) {id: serial, (*identifier*) ids: Inttab.set} * (*cumulative identifiers -- symbolic body content*) (*history*) {name: string, (*official theory name*) stage: stage option} (*index and counter for anonymous updates*) with fun rep_theory_id (Theory_Id args) = args; val make_theory_id = Theory_Id; end; datatype theory = Theory of (*identity*) {theory_id: theory_id, token: Position.T Unsynchronized.ref} * (*ancestry*) {parents: theory list, (*immediate predecessors*) ancestors: theory list} * (*all predecessors -- canonical reverse order*) (*data*) Any.T Datatab.table; (*body content*) exception THEORY of string * theory list; fun rep_theory (Theory args) = args; val theory_identity = #1 o rep_theory; val theory_id = #theory_id o theory_identity; val identity_of_id = #1 o rep_theory_id; val identity_of = identity_of_id o theory_id; val history_of_id = #2 o rep_theory_id; val history_of = history_of_id o theory_id; val ancestry_of = #2 o rep_theory; val data_of = #3 o rep_theory; fun make_identity id ids = {id = id, ids = ids}; fun make_history name = {name = name, stage = SOME (init_stage ())}; fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors}; val theory_id_ord = int_ord o apply2 (#id o identity_of_id); val theory_id_long_name = #name o history_of_id; val theory_id_name = Long_Name.base_name o theory_id_long_name; val theory_long_name = #name o history_of; val theory_name = Long_Name.base_name o theory_long_name; fun theory_name' {long} = if long then theory_long_name else theory_name; val parents_of = #parents o ancestry_of; val ancestors_of = #ancestors o ancestry_of; (* names *) val PureN = "Pure"; fun display_name thy_id = (case history_of_id thy_id of {name, stage = NONE} => name | {name, stage = SOME (i, _)} => name ^ ":" ^ string_of_int i); fun display_names thy = let val name = display_name (theory_id thy); val ancestor_names = map theory_long_name (ancestors_of thy); in rev (name :: ancestor_names) end; val pretty_thy = Pretty.str_list "{" "}" o display_names; val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_thy); fun pretty_abbrev_thy thy = let val names = display_names thy; val n = length names; val abbrev = if n > 5 then "..." :: List.drop (names, n - 5) else names; in Pretty.str_list "{" "}" abbrev end; fun get_theory long thy name = if theory_name' long thy <> name then (case find_first (fn thy' => theory_name' long thy' = name) (ancestors_of thy) of SOME thy' => thy' | NONE => error ("Unknown ancestor theory " ^ quote name)) else if is_none (#stage (history_of thy)) then thy else error ("Unfinished theory " ^ quote name); (* build ids *) fun insert_id id ids = Inttab.update (id, ()) ids; val merge_ids = apply2 (theory_id #> rep_theory_id #> #1) #> (fn ({id = id1, ids = ids1, ...}, {id = id2, ids = ids2, ...}) => Inttab.merge (K true) (ids1, ids2) |> insert_id id1 |> insert_id id2); (* equality and inclusion *) val eq_thy_id = op = o apply2 (#id o identity_of_id); val eq_thy = op = o apply2 (#id o identity_of); val proper_subthy_id = apply2 (rep_theory_id #> #1) #> (fn ({id, ...}, {ids, ...}) => Inttab.defined ids id); val proper_subthy = proper_subthy_id o apply2 theory_id; fun subthy_id p = eq_thy_id p orelse proper_subthy_id p; val subthy = subthy_id o apply2 theory_id; (* consistent ancestors *) fun eq_thy_consistent (thy1, thy2) = eq_thy (thy1, thy2) orelse (theory_name thy1 = theory_name thy2 andalso raise THEORY ("Duplicate theory name", [thy1, thy2])); fun extend_ancestors thy thys = if member eq_thy_consistent thys thy then raise THEORY ("Duplicate theory node", thy :: thys) else thy :: thys; val merge_ancestors = merge eq_thy_consistent; (** theory data **) (* data kinds and access methods *) val timing = Unsynchronized.ref false; local type kind = {pos: Position.T, empty: Any.T, extend: Any.T -> Any.T, merge: theory * theory -> Any.T * Any.T -> Any.T}; val kinds = Synchronized.var "Theory_Data" (Datatab.empty: kind Datatab.table); fun invoke name f k x = (case Datatab.lookup (Synchronized.value kinds) k of SOME kind => if ! timing andalso name <> "" then Timing.cond_timeit true ("Theory_Data." ^ name ^ Position.here (#pos kind)) (fn () => f kind x) else f kind x | NONE => raise Fail "Invalid theory data identifier"); in fun invoke_pos k = invoke "" (K o #pos) k (); fun invoke_empty k = invoke "" (K o #empty) k (); val invoke_extend = invoke "extend" #extend; fun invoke_merge thys = invoke "merge" (fn kind => #merge kind thys); fun declare_theory_data pos empty extend merge = let val k = serial (); val kind = {pos = pos, empty = empty, extend = extend, merge = merge}; val _ = Synchronized.change kinds (Datatab.update (k, kind)); in k end; val extend_data = Datatab.map invoke_extend; fun merge_data thys = Datatab.join (invoke_merge thys) o apply2 extend_data; end; (** build theories **) (* create theory *) val trace_theories = Unsynchronized.ref false; local val theories = Synchronized.var "theory_tokens" ([]: Position.T Unsynchronized.ref option Unsynchronized.ref list); val dummy_token = Unsynchronized.ref Position.none; fun make_token () = if ! trace_theories then let val token = Unsynchronized.ref (Position.thread_data ()); val _ = Synchronized.change theories (cons (Weak.weak (SOME token))); in token end else dummy_token; in fun theories_trace () = let val trace = Synchronized.value theories; val _ = ML_Heap.full_gc (); val active_positions = fold (fn Unsynchronized.ref (SOME pos) => cons (! pos) | _ => I) trace []; in {active_positions = active_positions, active = length active_positions, total = length trace} end; fun create_thy ids history ancestry data = let val theory_id = make_theory_id (make_identity (serial ()) ids, history); val token = make_token (); in Theory ({theory_id = theory_id, token = token}, ancestry, data) end; end; (* primitives *) val pre_pure_thy = create_thy Inttab.empty (make_history PureN) (make_ancestry [] []) Datatab.empty; local fun change_thy finish f thy = let val ({id, ids}, {name, stage}) = rep_theory_id (theory_id thy); val Theory (_, ancestry, data) = thy; val (ancestry', data') = if is_none stage then (make_ancestry [thy] (extend_ancestors thy (ancestors_of thy)), extend_data data) else (ancestry, data); val history' = {name = name, stage = if finish then NONE else Option.map next_stage stage}; val ids' = insert_id id ids; val data'' = f data'; in create_thy ids' history' ancestry' data'' end; in val update_thy = change_thy false; val extend_thy = update_thy I; val finish_thy = change_thy true I #> tap extend_thy; end; (* join: anonymous theory nodes *) local fun bad_join (thy1, thy2) = raise THEORY ("Cannot join theories", [thy1, thy2]); fun join_history thys = apply2 history_of thys |> (fn ({name, stage}, {name = name', stage = stage'}) => if name = name' andalso is_some stage andalso is_some stage' then {name = name, stage = Option.map next_stage stage} else bad_join thys); fun join_ancestry thys = apply2 ancestry_of thys |> (fn (ancestry as {parents, ancestors}, {parents = parents', ancestors = ancestors'}) => if eq_list eq_thy (parents, parents') andalso eq_list eq_thy (ancestors, ancestors') then ancestry else bad_join thys); in fun join_thys thys = let val ids = merge_ids thys; val history = join_history thys; val ancestry = join_ancestry thys; val data = merge_data thys (apply2 data_of thys); in create_thy ids history ancestry data end; end; (* merge: named theory nodes *) local fun merge_thys thys = let val ids = merge_ids thys; val history = make_history ""; val ancestry = make_ancestry [] []; val data = merge_data thys (apply2 data_of thys); in create_thy ids history ancestry data end; fun maximal_thys thys = thys |> filter_out (fn thy => exists (fn thy' => proper_subthy (thy, thy')) thys); in fun begin_thy name imports = if name = "" then error ("Bad theory name: " ^ quote name) else let val parents = maximal_thys (distinct eq_thy imports); val ancestors = Library.foldl merge_ancestors ([], map ancestors_of parents) |> fold extend_ancestors parents; val thy0 = (case parents of [] => error "Missing theory imports" | [thy] => extend_thy thy | thy :: thys => Library.foldl merge_thys (thy, thys)); val ({ids, ...}, _) = rep_theory_id (theory_id thy0); val history = make_history name; val ancestry = make_ancestry parents ancestors; in create_thy ids history ancestry (data_of thy0) |> tap finish_thy end; end; (* theory data *) structure Theory_Data = struct val declare = declare_theory_data; fun get k dest thy = (case Datatab.lookup (data_of thy) k of SOME x => x | NONE => invoke_empty k) |> dest; fun put k mk x = update_thy (Datatab.update (k, mk x)); fun obj_size k thy = Datatab.lookup (data_of thy) k |> Option.map ML_Heap.obj_size; end; fun theory_data_size thy = - (data_of thy, []) |-> Datatab.fold_rev (fn (k, _) => + build (data_of thy |> Datatab.fold_rev (fn (k, _) => (case Theory_Data.obj_size k thy of NONE => I - | SOME n => (cons (invoke_pos k, n)))); + | SOME n => (cons (invoke_pos k, n))))); (*** proof context ***) (* datatype Proof.context *) structure Proof = struct datatype context = Context of Any.T Datatab.table * theory; end; (* proof data kinds *) local val kinds = Synchronized.var "Proof_Data" (Datatab.empty: (theory -> Any.T) Datatab.table); fun init_data thy = Synchronized.value kinds |> Datatab.map (fn _ => fn init => init thy); fun init_new_data thy = Synchronized.value kinds |> Datatab.fold (fn (k, init) => fn data => if Datatab.defined data k then data else Datatab.update (k, init thy) data); fun init_fallback k thy = (case Datatab.lookup (Synchronized.value kinds) k of SOME init => init thy | NONE => raise Fail "Invalid proof data identifier"); in fun raw_transfer thy' (Proof.Context (data, thy)) = let val _ = subthy (thy, thy') orelse error "Cannot transfer proof context: not a super theory"; val data' = init_new_data thy' data; in Proof.Context (data', thy') end; structure Proof_Context = struct fun theory_of (Proof.Context (_, thy)) = thy; fun init_global thy = Proof.Context (init_data thy, thy); fun get_global thy name = init_global (get_theory {long = false} thy name); end; structure Proof_Data = struct fun declare init = let val k = serial (); val _ = Synchronized.change kinds (Datatab.update (k, init)); in k end; fun get k dest (Proof.Context (data, thy)) = (case Datatab.lookup data k of SOME x => x | NONE => init_fallback k thy) |> dest; fun put k mk x (Proof.Context (data, thy)) = Proof.Context (Datatab.update (k, mk x) data, thy); end; end; (*** theory certificate ***) datatype certificate = Certificate of theory | Certificate_Id of theory_id; fun certificate_theory (Certificate thy) = thy | certificate_theory (Certificate_Id thy_id) = error ("No content for theory certificate " ^ display_name thy_id); fun certificate_theory_id (Certificate thy) = theory_id thy | certificate_theory_id (Certificate_Id thy_id) = thy_id; fun eq_certificate (Certificate thy1, Certificate thy2) = eq_thy (thy1, thy2) | eq_certificate (Certificate_Id thy_id1, Certificate_Id thy_id2) = eq_thy_id (thy_id1, thy_id2) | eq_certificate _ = false; fun join_certificate (cert1, cert2) = let val (thy_id1, thy_id2) = apply2 certificate_theory_id (cert1, cert2) in if eq_thy_id (thy_id1, thy_id2) then (case cert1 of Certificate _ => cert1 | _ => cert2) else if proper_subthy_id (thy_id2, thy_id1) then cert1 else if proper_subthy_id (thy_id1, thy_id2) then cert2 else error ("Cannot join unrelated theory certificates " ^ display_name thy_id1 ^ " and " ^ display_name thy_id2) end; (*** generic context ***) datatype generic = Theory of theory | Proof of Proof.context; fun cases f _ (Theory thy) = f thy | cases _ g (Proof prf) = g prf; fun mapping f g = cases (Theory o f) (Proof o g); fun mapping_result f g = cases (apsnd Theory o f) (apsnd Proof o g); val the_theory = cases I (fn _ => error "Ill-typed context: theory expected"); val the_proof = cases (fn _ => error "Ill-typed context: proof expected") I; fun map_theory f = Theory o f o the_theory; fun map_proof f = Proof o f o the_proof; fun map_theory_result f = apsnd Theory o f o the_theory; fun map_proof_result f = apsnd Proof o f o the_proof; fun theory_map f = the_theory o f o Theory; fun proof_map f = the_proof o f o Proof; val theory_of = cases I Proof_Context.theory_of; val proof_of = cases Proof_Context.init_global I; (** thread data **) local val generic_context_var = Thread_Data.var () : generic Thread_Data.var in fun get_generic_context () = Thread_Data.get generic_context_var; val put_generic_context = Thread_Data.put generic_context_var; fun setmp_generic_context opt_context = Thread_Data.setmp generic_context_var opt_context; fun the_generic_context () = (case get_generic_context () of SOME context => context | _ => error "Unknown context"); val the_global_context = theory_of o the_generic_context; val the_local_context = proof_of o the_generic_context; end; fun >>> f = let val (res, context') = f (the_generic_context ()); val _ = put_generic_context (SOME context'); in res end; nonfix >>; fun >> f = >>> (fn context => ((), f context)); val _ = put_generic_context (SOME (Theory pre_pure_thy)); end; structure Basic_Context: BASIC_CONTEXT = Context; open Basic_Context; (*** type-safe interfaces for data declarations ***) (** theory data **) signature THEORY_DATA'_ARGS = sig type T val empty: T val extend: T -> T val merge: theory * theory -> T * T -> T end; signature THEORY_DATA_ARGS = sig type T val empty: T val extend: T -> T val merge: T * T -> T end; signature THEORY_DATA = sig type T val get: theory -> T val put: T -> theory -> theory val map: (T -> T) -> theory -> theory end; functor Theory_Data'(Data: THEORY_DATA'_ARGS): THEORY_DATA = struct type T = Data.T; exception Data of T; val kind = let val pos = Position.thread_data () in Context.Theory_Data.declare pos (Data Data.empty) (fn Data x => let val y = Data.extend x; val _ = if pointer_eq (x, y) then () else raise Fail ("Theory_Data extend needs to be identity" ^ Position.here pos) in Data y end) (fn thys => fn (Data x1, Data x2) => Data (Data.merge thys (x1, x2))) end; val get = Context.Theory_Data.get kind (fn Data x => x); val put = Context.Theory_Data.put kind Data; fun map f thy = put (f (get thy)) thy; end; functor Theory_Data(Data: THEORY_DATA_ARGS): THEORY_DATA = Theory_Data' ( type T = Data.T; val empty = Data.empty; val extend = Data.extend; fun merge _ = Data.merge; ); (** proof data **) signature PROOF_DATA_ARGS = sig type T val init: theory -> T end; signature PROOF_DATA = sig type T val get: Proof.context -> T val put: T -> Proof.context -> Proof.context val map: (T -> T) -> Proof.context -> Proof.context end; functor Proof_Data(Data: PROOF_DATA_ARGS): PROOF_DATA = struct type T = Data.T; exception Data of T; val kind = Context.Proof_Data.declare (Data o Data.init); val get = Context.Proof_Data.get kind (fn Data x => x); val put = Context.Proof_Data.put kind Data; fun map f prf = put (f (get prf)) prf; end; (** generic data **) signature GENERIC_DATA_ARGS = sig type T val empty: T val extend: T -> T val merge: T * T -> T end; signature GENERIC_DATA = sig type T val get: Context.generic -> T val put: T -> Context.generic -> Context.generic val map: (T -> T) -> Context.generic -> Context.generic end; functor Generic_Data(Data: GENERIC_DATA_ARGS): GENERIC_DATA = struct structure Thy_Data = Theory_Data(Data); structure Prf_Data = Proof_Data(type T = Data.T val init = Thy_Data.get); type T = Data.T; fun get (Context.Theory thy) = Thy_Data.get thy | get (Context.Proof prf) = Prf_Data.get prf; fun put x (Context.Theory thy) = Context.Theory (Thy_Data.put x thy) | put x (Context.Proof prf) = Context.Proof (Prf_Data.put x prf); fun map f ctxt = put (f (get ctxt)) ctxt; end; (*hide private interface*) structure Context: CONTEXT = Context; diff --git a/src/Pure/defs.ML b/src/Pure/defs.ML --- a/src/Pure/defs.ML +++ b/src/Pure/defs.ML @@ -1,283 +1,283 @@ (* Title: Pure/defs.ML Author: Makarius Global well-formedness checks for overloaded definitions (mixed constants and types). Recall that constant definitions may be explained syntactically within Pure, but type definitions require particular set-theoretic semantics. *) signature DEFS = sig datatype item_kind = Const | Type type item = item_kind * string type entry = item * typ list val item_kind_ord: item_kind ord val plain_args: typ list -> bool type context = Proof.context * (Name_Space.T * Name_Space.T) val global_context: theory -> context val space: context -> item_kind -> Name_Space.T val pretty_item: context -> item -> Pretty.T val pretty_args: Proof.context -> typ list -> Pretty.T list val pretty_entry: context -> entry -> Pretty.T type T type spec = {def: string option, description: string, pos: Position.T, lhs: typ list, rhs: entry list} val all_specifications_of: T -> (item * spec list) list val specifications_of: T -> item -> spec list val dest: T -> {restricts: (entry * string) list, reducts: (entry * entry list) list} val dest_constdefs: T list -> T -> (string * string) list val empty: T val merge: context -> T * T -> T val define: context -> bool -> string option -> string -> entry -> entry list -> T -> T val get_deps: T -> item -> (typ list * entry list) list end; structure Defs: DEFS = struct (* specification items *) datatype item_kind = Const | Type; type item = item_kind * string; type entry = item * typ list; fun item_kind_ord (Const, Type) = LESS | item_kind_ord (Type, Const) = GREATER | item_kind_ord _ = EQUAL; structure Itemtab = Table(type key = item val ord = prod_ord item_kind_ord fast_string_ord); (* pretty printing *) type context = Proof.context * (Name_Space.T * Name_Space.T); fun global_context thy = (Syntax.init_pretty_global thy, (Sign.const_space thy, Sign.type_space thy)); fun space ((_, spaces): context) kind = if kind = Const then #1 spaces else #2 spaces; fun pretty_item (context as (ctxt, _)) (kind, name) = let val prt_name = Name_Space.pretty ctxt (space context kind) name in if kind = Const then prt_name else Pretty.block [Pretty.keyword1 "type", Pretty.brk 1, prt_name] end; fun pretty_args ctxt args = if null args then [] else [Pretty.list "(" ")" (map (Syntax.pretty_typ ctxt o Logic.unvarifyT_global) args)]; fun pretty_entry context (c, args) = Pretty.block (pretty_item context c :: pretty_args (#1 context) args); (* type arguments *) fun plain_args args = forall Term.is_TVar args andalso not (has_duplicates (op =) args); fun disjoint_args (Ts, Us) = not (Type.could_unifys (Ts, Us)) orelse ((Vartab.build (Type.raw_unifys (Ts, map (Logic.incr_tvar (maxidx_of_typs Ts + 1)) Us)); false) handle Type.TUNIFY => true); fun match_args (Ts, Us) = if Type.could_matches (Ts, Us) then Option.map Envir.subst_type (SOME (Vartab.build (Type.raw_matches (Ts, Us))) handle Type.TYPE_MATCH => NONE) else NONE; (* datatype defs *) type spec = {def: string option, description: string, pos: Position.T, lhs: typ list, rhs: entry list}; type def = {specs: spec Inttab.table, (*source specifications*) restricts: (typ list * string) list, (*global restrictions imposed by incomplete patterns*) reducts: (typ list * entry list) list}; (*specifications as reduction system*) fun make_def (specs, restricts, reducts) = {specs = specs, restricts = restricts, reducts = reducts}: def; fun map_def c f = Itemtab.default (c, make_def (Inttab.empty, [], [])) #> Itemtab.map_entry c (fn {specs, restricts, reducts}: def => make_def (f (specs, restricts, reducts))); datatype T = Defs of def Itemtab.table; fun lookup_list which defs c = (case Itemtab.lookup defs c of SOME (def: def) => which def | NONE => []); fun all_specifications_of (Defs defs) = (map o apsnd) (map snd o Inttab.dest o #specs) (Itemtab.dest defs); fun specifications_of (Defs defs) = lookup_list (map snd o Inttab.dest o #specs) defs; val restricts_of = lookup_list #restricts; val reducts_of = lookup_list #reducts; fun dest (Defs defs) = let val restricts = Itemtab.fold (fn (c, {restricts, ...}) => fold (fn (args, description) => cons ((c, args), description)) restricts) defs []; val reducts = Itemtab.fold (fn (c, {reducts, ...}) => fold (fn (args, deps) => cons ((c, args), deps)) reducts) defs []; in {restricts = restricts, reducts = reducts} end; fun dest_constdefs prevs (Defs defs) = let fun prev_spec c i = prevs |> exists (fn Defs prev_defs => (case Itemtab.lookup prev_defs c of NONE => false | SOME {specs, ...} => Inttab.defined specs i)); in - (defs, []) |-> Itemtab.fold (fn (c, {specs, ...}) => + build (defs |> Itemtab.fold (fn (c, {specs, ...}) => specs |> Inttab.fold (fn (i, spec) => if #1 c = Const andalso is_some (#def spec) andalso not (prev_spec c i) - then cons (#2 c, the (#def spec)) else I)) + then cons (#2 c, the (#def spec)) else I))) end; val empty = Defs Itemtab.empty; (* specifications *) fun disjoint_specs context c (i, {description = a, pos = pos_a, lhs = Ts, ...}: spec) = Inttab.forall (fn (j, {description = b, pos = pos_b, lhs = Us, ...}: spec) => i = j orelse disjoint_args (Ts, Us) orelse error ("Clash of specifications for " ^ Pretty.unformatted_string_of (pretty_item context c) ^ ":\n" ^ " " ^ quote a ^ Position.here pos_a ^ "\n" ^ " " ^ quote b ^ Position.here pos_b)); fun join_specs context c ({specs = specs1, restricts, reducts}, {specs = specs2, ...}: def) = let val specs' = Inttab.fold (fn spec2 => (disjoint_specs context c spec2 specs1; Inttab.update spec2)) specs2 specs1; in make_def (specs', restricts, reducts) end; fun update_specs context c spec = map_def c (fn (specs, restricts, reducts) => (disjoint_specs context c spec specs; (Inttab.update spec specs, restricts, reducts))); (* normalized dependencies: reduction with well-formedness check *) local val prt = Pretty.string_of oo pretty_entry; fun err context (c, Ts) (d, Us) s1 s2 = error (s1 ^ " dependency of " ^ prt context (c, Ts) ^ " -> " ^ prt context (d, Us) ^ s2); fun acyclic context (c, Ts) (d, Us) = c <> d orelse is_none (match_args (Ts, Us)) orelse err context (c, Ts) (d, Us) "Circular" ""; fun reduction context defs const deps = let fun reduct Us (Ts, rhs) = (case match_args (Ts, Us) of NONE => NONE | SOME subst => SOME (map (apsnd (map subst)) rhs)); fun reducts (d, Us) = get_first (reduct Us) (reducts_of defs d); val reds = map (`reducts) deps; val deps' = if forall (is_none o #1) reds then NONE else SOME (fold_rev (fn (NONE, dp) => insert (op =) dp | (SOME dps, _) => fold (insert (op =)) dps) reds []); val _ = forall (acyclic context const) (the_default deps deps'); in deps' end; fun restriction context defs (c, Ts) (d, Us) = plain_args Us orelse (case find_first (fn (Rs, _) => not (disjoint_args (Rs, Us))) (restricts_of defs d) of SOME (Rs, description) => err context (c, Ts) (d, Us) "Malformed" ("\n(restriction " ^ prt context (d, Rs) ^ " from " ^ quote description ^ ")") | NONE => true); in fun normalize context = let fun check_def defs (c, {reducts, ...}: def) = reducts |> forall (fn (Ts, deps) => forall (restriction context defs (c, Ts)) deps); fun check_defs defs = Itemtab.forall (check_def defs) defs; fun norm_update (c, {reducts, ...}: def) (changed, defs) = let val reducts' = reducts |> map (fn (Ts, deps) => (Ts, perhaps (reduction context defs (c, Ts)) deps)); in if reducts = reducts' then (changed, defs) else (true, defs |> map_def c (fn (specs, restricts, _) => (specs, restricts, reducts'))) end; fun norm_loop defs = (case Itemtab.fold norm_update defs (false, defs) of (true, defs') => norm_loop defs' | (false, _) => defs); in norm_loop #> tap check_defs end; fun dependencies context (c, args) restr deps = map_def c (fn (specs, restricts, reducts) => let val restricts' = Library.merge (op =) (restricts, restr); val reducts' = insert (op =) (args, deps) reducts; in (specs, restricts', reducts') end) #> normalize context; end; (* merge *) fun merge context (Defs defs1, Defs defs2) = let fun add_deps (c, args) restr deps defs = if AList.defined (op =) (reducts_of defs c) args then defs else dependencies context (c, args) restr deps defs; fun add_def (c, {restricts, reducts, ...}: def) = fold (fn (args, deps) => add_deps (c, args) restricts deps) reducts; in Defs (Itemtab.join (join_specs context) (defs1, defs2) |> normalize context |> Itemtab.fold add_def defs2) end; (* define *) fun define context unchecked def description (c, args) deps (Defs defs) = let val pos = Position.thread_data (); val restr = if plain_args args orelse (case args of [Term.Type (_, rec_args)] => plain_args rec_args | _ => false) then [] else [(args, description)]; val spec = (serial (), {def = def, description = description, pos = pos, lhs = args, rhs = deps}); val defs' = defs |> update_specs context c spec; in Defs (defs' |> (if unchecked then I else dependencies context (c, args) restr deps)) end; fun get_deps (Defs defs) c = reducts_of defs c; end; diff --git a/src/Pure/drule.ML b/src/Pure/drule.ML --- a/src/Pure/drule.ML +++ b/src/Pure/drule.ML @@ -1,843 +1,843 @@ (* Title: Pure/drule.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Derived rules and other operations on theorems. *) infix 0 RL RLN MRS OF COMP INCR_COMP COMP_INCR; signature BASIC_DRULE = sig val mk_implies: cterm * cterm -> cterm val list_implies: cterm list * cterm -> cterm val strip_imp_prems: cterm -> cterm list val strip_imp_concl: cterm -> cterm val cprems_of: thm -> cterm list val forall_intr_list: cterm list -> thm -> thm val forall_intr_vars: thm -> thm val forall_elim_list: cterm list -> thm -> thm val lift_all: Proof.context -> cterm -> thm -> thm val implies_elim_list: thm -> thm list -> thm val implies_intr_list: cterm list -> thm -> thm val instantiate_normalize: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> thm -> thm val instantiate'_normalize: ctyp option list -> cterm option list -> thm -> thm val infer_instantiate_types: Proof.context -> ((indexname * typ) * cterm) list -> thm -> thm val infer_instantiate: Proof.context -> (indexname * cterm) list -> thm -> thm val infer_instantiate': Proof.context -> cterm option list -> thm -> thm val zero_var_indexes_list: thm list -> thm list val zero_var_indexes: thm -> thm val implies_intr_hyps: thm -> thm val rotate_prems: int -> thm -> thm val rearrange_prems: int list -> thm -> thm val RLN: thm list * (int * thm list) -> thm list val RL: thm list * thm list -> thm list val MRS: thm list * thm -> thm val OF: thm * thm list -> thm val COMP: thm * thm -> thm val INCR_COMP: thm * thm -> thm val COMP_INCR: thm * thm -> thm val size_of_thm: thm -> int val reflexive_thm: thm val symmetric_thm: thm val transitive_thm: thm val extensional: thm -> thm val asm_rl: thm val cut_rl: thm val revcut_rl: thm val thin_rl: thm end; signature DRULE = sig include BASIC_DRULE val outer_params: term -> (string * typ) list val generalize: Symtab.set * Symtab.set -> thm -> thm val list_comb: cterm * cterm list -> cterm val strip_comb: cterm -> cterm * cterm list val beta_conv: cterm -> cterm -> cterm val flexflex_unique: Proof.context option -> thm -> thm val export_without_context: thm -> thm val export_without_context_open: thm -> thm val store_thm: binding -> thm -> thm val store_standard_thm: binding -> thm -> thm val store_thm_open: binding -> thm -> thm val store_standard_thm_open: binding -> thm -> thm val multi_resolve: Proof.context option -> thm list -> thm -> thm Seq.seq val multi_resolves: Proof.context option -> thm list -> thm list -> thm Seq.seq val compose: thm * int * thm -> thm val equals_cong: thm val imp_cong: thm val swap_prems_eq: thm val imp_cong_rule: thm -> thm -> thm val arg_cong_rule: cterm -> thm -> thm val binop_cong_rule: cterm -> thm -> thm -> thm val fun_cong_rule: thm -> cterm -> thm val beta_eta_conversion: cterm -> thm val eta_contraction_rule: thm -> thm val norm_hhf_eq: thm val norm_hhf_eqs: thm list val is_norm_hhf: term -> bool val norm_hhf: theory -> term -> term val norm_hhf_cterm: Proof.context -> cterm -> cterm val protect: cterm -> cterm val protectI: thm val protectD: thm val protect_cong: thm val implies_intr_protected: cterm list -> thm -> thm val termI: thm val mk_term: cterm -> thm val dest_term: thm -> cterm val cterm_rule: (thm -> thm) -> cterm -> cterm val cterm_add_frees: cterm -> cterm list -> cterm list val cterm_add_vars: cterm -> cterm list -> cterm list val dummy_thm: thm val free_dummy_thm: thm val is_sort_constraint: term -> bool val sort_constraintI: thm val sort_constraint_eq: thm val with_subgoal: int -> (thm -> thm) -> thm -> thm val comp_no_flatten: thm * int -> int -> thm -> thm val rename_bvars: (string * string) list -> thm -> thm val rename_bvars': string option list -> thm -> thm val incr_indexes: thm -> thm -> thm val incr_indexes2: thm -> thm -> thm -> thm val triv_forall_equality: thm val distinct_prems_rl: thm val equal_intr_rule: thm val equal_elim_rule1: thm val equal_elim_rule2: thm val remdups_rl: thm val abs_def: thm -> thm end; structure Drule: DRULE = struct (** some cterm->cterm operations: faster than calling cterm_of! **) (* A1\...An\B goes to [A1,...,An], where B is not an implication *) fun strip_imp_prems ct = let val (cA, cB) = Thm.dest_implies ct in cA :: strip_imp_prems cB end handle TERM _ => []; (* A1\...An\B goes to B, where B is not an implication *) fun strip_imp_concl ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => strip_imp_concl (Thm.dest_arg ct) | _ => ct); (*The premises of a theorem, as a cterm list*) val cprems_of = strip_imp_prems o Thm.cprop_of; fun certify t = Thm.global_cterm_of (Context.the_global_context ()) t; val implies = certify Logic.implies; fun mk_implies (A, B) = Thm.apply (Thm.apply implies A) B; (*cterm version of list_implies: [A1,...,An], B goes to \A1;...;An\\B *) fun list_implies([], B) = B | list_implies(A::As, B) = mk_implies (A, list_implies(As,B)); (*cterm version of list_comb: maps (f, [t1,...,tn]) to f(t1,...,tn) *) fun list_comb (f, []) = f | list_comb (f, t::ts) = list_comb (Thm.apply f t, ts); (*cterm version of strip_comb: maps f(t1,...,tn) to (f, [t1,...,tn]) *) fun strip_comb ct = let fun stripc (p as (ct, cts)) = let val (ct1, ct2) = Thm.dest_comb ct in stripc (ct1, ct2 :: cts) end handle CTERM _ => p in stripc (ct, []) end; (*Beta-conversion for cterms, where x is an abstraction. Simply returns the rhs of the meta-equality returned by the beta_conversion rule.*) fun beta_conv x y = Thm.dest_arg (Thm.cprop_of (Thm.beta_conversion false (Thm.apply x y))); (** Standardization of rules **) (*Generalization over a list of variables*) val forall_intr_list = fold_rev Thm.forall_intr; (*Generalization over Vars -- canonical order*) fun forall_intr_vars th = fold Thm.forall_intr (Thm.add_vars th []) th; fun outer_params t = let val vs = Term.strip_all_vars t in Name.variant_list [] (map (Name.clean o #1) vs) ~~ map #2 vs end; (*lift vars wrt. outermost goal parameters -- reverses the effect of gen_all modulo higher-order unification*) fun lift_all ctxt raw_goal raw_th = let val thy = Proof_Context.theory_of ctxt; val goal = Thm.transfer_cterm thy raw_goal; val th = Thm.transfer thy raw_th; val maxidx = Thm.maxidx_of th; val ps = outer_params (Thm.term_of goal) |> map (fn (x, T) => Var ((x, maxidx + 1), Logic.incr_tvar (maxidx + 1) T)); val Ts = map Term.fastype_of ps; val inst = Term_Subst.Vars.build (th |> (Thm.fold_terms o Term.fold_aterms) (fn t => fn inst => (case t of Var (xi, T) => if Term_Subst.Vars.defined inst (xi, T) then inst else let val ct = Thm.cterm_of ctxt (Term.list_comb (Var (xi, Ts ---> T), ps)) in Term_Subst.Vars.update ((xi, T), ct) inst end | _ => inst))); in th |> Thm.instantiate ([], Term_Subst.Vars.dest inst) |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) ps end; (*direct generalization*) fun generalize names th = Thm.generalize names (Thm.maxidx_of th + 1) th; (*specialization over a list of cterms*) val forall_elim_list = fold Thm.forall_elim; (*maps A1,...,An |- B to \A1;...;An\ \ B*) val implies_intr_list = fold_rev Thm.implies_intr; (*maps \A1;...;An\ \ B and [A1,...,An] to B*) fun implies_elim_list impth ths = fold Thm.elim_implies ths impth; (*Reset Var indexes to zero, renaming to preserve distinctness*) fun zero_var_indexes_list [] = [] | zero_var_indexes_list ths = let val (instT, inst) = Term_Subst.zero_var_indexes_inst Name.context (map Thm.full_prop_of ths); val tvars = fold Thm.add_tvars ths []; fun the_tvar v = the (find_first (fn cT => v = dest_TVar (Thm.typ_of cT)) tvars); val instT' = - (instT, []) |-> Term_Subst.TVars.fold (fn (v, TVar (b, _)) => - cons (v, Thm.rename_tvar b (the_tvar v))); + build (instT |> Term_Subst.TVars.fold (fn (v, TVar (b, _)) => + cons (v, Thm.rename_tvar b (the_tvar v)))); val vars = fold (Thm.add_vars o Thm.instantiate (instT', [])) ths []; fun the_var v = the (find_first (fn ct => v = dest_Var (Thm.term_of ct)) vars); val inst' = - (inst, []) |-> Term_Subst.Vars.fold (fn (v, Var (b, _)) => - cons (v, Thm.var (b, Thm.ctyp_of_cterm (the_var v)))); + build (inst |> Term_Subst.Vars.fold (fn (v, Var (b, _)) => + cons (v, Thm.var (b, Thm.ctyp_of_cterm (the_var v))))); in map (Thm.adjust_maxidx_thm ~1 o Thm.instantiate (instT', inst')) ths end; val zero_var_indexes = singleton zero_var_indexes_list; (** Standard form of object-rule: no hypotheses, flexflex constraints, Frees, or outer quantifiers; all generality expressed by Vars of index 0.**) (*Discharge all hypotheses.*) fun implies_intr_hyps th = fold Thm.implies_intr (Thm.chyps_of th) th; (*Squash a theorem's flexflex constraints provided it can be done uniquely. This step can lose information.*) fun flexflex_unique opt_ctxt th = if null (Thm.tpairs_of th) then th else (case distinct Thm.eq_thm (Seq.list_of (Thm.flexflex_rule opt_ctxt th)) of [th] => th | [] => raise THM ("flexflex_unique: impossible constraints", 0, [th]) | _ => raise THM ("flexflex_unique: multiple unifiers", 0, [th])); (* old-style export without context *) val export_without_context_open = implies_intr_hyps #> Thm.forall_intr_frees #> `Thm.maxidx_of #-> (fn maxidx => Thm.forall_elim_vars (maxidx + 1) #> Thm.strip_shyps #> zero_var_indexes #> Thm.varifyT_global); val export_without_context = flexflex_unique NONE #> export_without_context_open #> Thm.close_derivation \<^here>; (*Rotates a rule's premises to the left by k*) fun rotate_prems 0 = I | rotate_prems k = Thm.permute_prems 0 k; fun with_subgoal i f = rotate_prems (i - 1) #> f #> rotate_prems (1 - i); (*Permute prems, where the i-th position in the argument list (counting from 0) gives the position within the original thm to be transferred to position i. Any remaining trailing positions are left unchanged.*) val rearrange_prems = let fun rearr new [] thm = thm | rearr new (p :: ps) thm = rearr (new + 1) (map (fn q => if new <= q andalso q < p then q + 1 else q) ps) (Thm.permute_prems (new + 1) (new - p) (Thm.permute_prems new (p - new) thm)) in rearr 0 end; (*Resolution: multiple arguments, multiple results*) local fun res opt_ctxt th i rule = (Thm.biresolution opt_ctxt false [(false, th)] i rule handle THM _ => Seq.empty) |> Seq.map Thm.solve_constraints; fun multi_res _ _ [] rule = Seq.single rule | multi_res opt_ctxt i (th :: ths) rule = Seq.maps (res opt_ctxt th i) (multi_res opt_ctxt (i + 1) ths rule); in fun multi_resolve opt_ctxt = multi_res opt_ctxt 1; fun multi_resolves opt_ctxt facts rules = Seq.maps (multi_resolve opt_ctxt facts) (Seq.of_list rules); end; (*For joining lists of rules*) fun thas RLN (i, thbs) = let val resolve = Thm.biresolution NONE false (map (pair false) thas) i fun resb thb = Seq.list_of (resolve thb) handle THM _ => [] in maps resb thbs |> map Thm.solve_constraints end; fun thas RL thbs = thas RLN (1, thbs); (*Isar-style multi-resolution*) fun bottom_rl OF rls = (case Seq.chop 2 (multi_resolve NONE rls bottom_rl) of ([th], _) => Thm.solve_constraints th | ([], _) => raise THM ("OF: no unifiers", 0, bottom_rl :: rls) | _ => raise THM ("OF: multiple unifiers", 0, bottom_rl :: rls)); (*Resolve a list of rules against bottom_rl from right to left; makes proof trees*) fun rls MRS bottom_rl = bottom_rl OF rls; (*compose Q and \...,Qi,Q(i+1),...\ \ R to \...,Q(i+1),...\ \ R with no lifting or renaming! Q may contain \ or meta-quantifiers ALWAYS deletes premise i *) fun compose (tha, i, thb) = Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, tha, 0) i thb |> Seq.list_of |> distinct Thm.eq_thm |> (fn [th] => Thm.solve_constraints th | _ => raise THM ("compose: unique result expected", i, [tha, thb])); (** theorem equality **) (*Useful "distance" function for BEST_FIRST*) val size_of_thm = size_of_term o Thm.full_prop_of; (*** Meta-Rewriting Rules ***) val read_prop = certify o Simple_Syntax.read_prop; fun store_thm name th = Context.>>> (Context.map_theory_result (Global_Theory.store_thm (name, th))); fun store_thm_open name th = Context.>>> (Context.map_theory_result (Global_Theory.store_thm_open (name, th))); fun store_standard_thm name th = store_thm name (export_without_context th); fun store_standard_thm_open name th = store_thm_open name (export_without_context_open th); val reflexive_thm = let val cx = certify (Var(("x",0),TVar(("'a",0),[]))) in store_standard_thm_open (Binding.make ("reflexive", \<^here>)) (Thm.reflexive cx) end; val symmetric_thm = let val xy = read_prop "x::'a \ y::'a"; val thm = Thm.implies_intr xy (Thm.symmetric (Thm.assume xy)); in store_standard_thm_open (Binding.make ("symmetric", \<^here>)) thm end; val transitive_thm = let val xy = read_prop "x::'a \ y::'a"; val yz = read_prop "y::'a \ z::'a"; val xythm = Thm.assume xy; val yzthm = Thm.assume yz; val thm = Thm.implies_intr yz (Thm.transitive xythm yzthm); in store_standard_thm_open (Binding.make ("transitive", \<^here>)) thm end; fun extensional eq = let val eq' = Thm.abstract_rule "x" (Thm.dest_arg (fst (Thm.dest_equals (Thm.cprop_of eq)))) eq in Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of eq')) eq' end; val equals_cong = store_standard_thm_open (Binding.make ("equals_cong", \<^here>)) (Thm.reflexive (read_prop "x::'a \ y::'a")); val imp_cong = let val ABC = read_prop "A \ B::prop \ C::prop" val AB = read_prop "A \ B" val AC = read_prop "A \ C" val A = read_prop "A" in store_standard_thm_open (Binding.make ("imp_cong", \<^here>)) (Thm.implies_intr ABC (Thm.equal_intr (Thm.implies_intr AB (Thm.implies_intr A (Thm.equal_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.implies_elim (Thm.assume AB) (Thm.assume A))))) (Thm.implies_intr AC (Thm.implies_intr A (Thm.equal_elim (Thm.symmetric (Thm.implies_elim (Thm.assume ABC) (Thm.assume A))) (Thm.implies_elim (Thm.assume AC) (Thm.assume A))))))) end; val swap_prems_eq = let val ABC = read_prop "A \ B \ C" val BAC = read_prop "B \ A \ C" val A = read_prop "A" val B = read_prop "B" in store_standard_thm_open (Binding.make ("swap_prems_eq", \<^here>)) (Thm.equal_intr (Thm.implies_intr ABC (Thm.implies_intr B (Thm.implies_intr A (Thm.implies_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.assume B))))) (Thm.implies_intr BAC (Thm.implies_intr A (Thm.implies_intr B (Thm.implies_elim (Thm.implies_elim (Thm.assume BAC) (Thm.assume B)) (Thm.assume A)))))) end; val imp_cong_rule = Thm.combination o Thm.combination (Thm.reflexive implies); fun arg_cong_rule ct th = Thm.combination (Thm.reflexive ct) th; (*AP_TERM in LCF/HOL*) fun fun_cong_rule th ct = Thm.combination th (Thm.reflexive ct); (*AP_THM in LCF/HOL*) fun binop_cong_rule ct th1 th2 = Thm.combination (arg_cong_rule ct th1) th2; fun beta_eta_conversion ct = let val thm = Thm.beta_conversion true ct in Thm.transitive thm (Thm.eta_conversion (Thm.rhs_of thm)) end; (*Contract all eta-redexes in the theorem, lest they give rise to needless abstractions*) fun eta_contraction_rule th = Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of th)) th; (* abs_def *) (* f ?x1 ... ?xn \ u -------------------- f \ \x1 ... xn. u *) local fun contract_lhs th = Thm.transitive (Thm.symmetric (beta_eta_conversion (#1 (Thm.dest_equals (Thm.cprop_of th))))) th; fun var_args ct = (case try Thm.dest_comb ct of SOME (f, arg) => (case Thm.term_of arg of Var ((x, _), _) => update (eq_snd (op aconvc)) (x, arg) (var_args f) | _ => []) | NONE => []); in fun abs_def th = let val th' = contract_lhs th; val args = var_args (Thm.lhs_of th'); in contract_lhs (fold (uncurry Thm.abstract_rule) args th') end; end; (*** Some useful meta-theorems ***) (*The rule V/V, obtains assumption solving for eresolve_tac*) val asm_rl = store_standard_thm_open (Binding.make ("asm_rl", \<^here>)) (Thm.trivial (read_prop "?psi")); (*Meta-level cut rule: \V \ W; V\ \ W *) val cut_rl = store_standard_thm_open (Binding.make ("cut_rl", \<^here>)) (Thm.trivial (read_prop "?psi \ ?theta")); (*Generalized elim rule for one conclusion; cut_rl with reversed premises: \PROP V; PROP V \ PROP W\ \ PROP W *) val revcut_rl = let val V = read_prop "V"; val VW = read_prop "V \ W"; in store_standard_thm_open (Binding.make ("revcut_rl", \<^here>)) (Thm.implies_intr V (Thm.implies_intr VW (Thm.implies_elim (Thm.assume VW) (Thm.assume V)))) end; (*for deleting an unwanted assumption*) val thin_rl = let val V = read_prop "V"; val W = read_prop "W"; val thm = Thm.implies_intr V (Thm.implies_intr W (Thm.assume W)); in store_standard_thm_open (Binding.make ("thin_rl", \<^here>)) thm end; (* (\x. PROP ?V) \ PROP ?V Allows removal of redundant parameters*) val triv_forall_equality = let val V = read_prop "V"; val QV = read_prop "\x::'a. V"; val x = certify (Free ("x", Term.aT [])); in store_standard_thm_open (Binding.make ("triv_forall_equality", \<^here>)) (Thm.equal_intr (Thm.implies_intr QV (Thm.forall_elim x (Thm.assume QV))) (Thm.implies_intr V (Thm.forall_intr x (Thm.assume V)))) end; (* (PROP ?Phi \ PROP ?Phi \ PROP ?Psi) \ (PROP ?Phi \ PROP ?Psi) *) val distinct_prems_rl = let val AAB = read_prop "Phi \ Phi \ Psi"; val A = read_prop "Phi"; in store_standard_thm_open (Binding.make ("distinct_prems_rl", \<^here>)) (implies_intr_list [AAB, A] (implies_elim_list (Thm.assume AAB) [Thm.assume A, Thm.assume A])) end; (* \PROP ?phi \ PROP ?psi; PROP ?psi \ PROP ?phi\ \ PROP ?phi \ PROP ?psi Introduction rule for \ as a meta-theorem. *) val equal_intr_rule = let val PQ = read_prop "phi \ psi"; val QP = read_prop "psi \ phi"; in store_standard_thm_open (Binding.make ("equal_intr_rule", \<^here>)) (Thm.implies_intr PQ (Thm.implies_intr QP (Thm.equal_intr (Thm.assume PQ) (Thm.assume QP)))) end; (* PROP ?phi \ PROP ?psi \ PROP ?phi \ PROP ?psi *) val equal_elim_rule1 = let val eq = read_prop "phi::prop \ psi::prop"; val P = read_prop "phi"; in store_standard_thm_open (Binding.make ("equal_elim_rule1", \<^here>)) (Thm.equal_elim (Thm.assume eq) (Thm.assume P) |> implies_intr_list [eq, P]) end; (* PROP ?psi \ PROP ?phi \ PROP ?phi \ PROP ?psi *) val equal_elim_rule2 = store_standard_thm_open (Binding.make ("equal_elim_rule2", \<^here>)) (symmetric_thm RS equal_elim_rule1); (* PROP ?phi \ PROP ?phi \ PROP ?psi \ PROP ?psi *) val remdups_rl = let val P = read_prop "phi"; val Q = read_prop "psi"; val thm = implies_intr_list [P, P, Q] (Thm.assume Q); in store_standard_thm_open (Binding.make ("remdups_rl", \<^here>)) thm end; (** embedded terms and types **) local val A = certify (Free ("A", propT)); val axiom = Thm.unvarify_axiom (Context.the_global_context ()); val prop_def = axiom "Pure.prop_def"; val term_def = axiom "Pure.term_def"; val sort_constraint_def = axiom "Pure.sort_constraint_def"; val C = Thm.lhs_of sort_constraint_def; val T = Thm.dest_arg C; val CA = mk_implies (C, A); in (* protect *) val protect = Thm.apply (certify Logic.protectC); val protectI = store_standard_thm (Binding.concealed (Binding.make ("protectI", \<^here>))) (Thm.equal_elim (Thm.symmetric prop_def) (Thm.assume A)); val protectD = store_standard_thm (Binding.concealed (Binding.make ("protectD", \<^here>))) (Thm.equal_elim prop_def (Thm.assume (protect A))); val protect_cong = store_standard_thm_open (Binding.make ("protect_cong", \<^here>)) (Thm.reflexive (protect A)); fun implies_intr_protected asms th = let val asms' = map protect asms in implies_elim_list (implies_intr_list asms th) (map (fn asm' => Thm.assume asm' RS protectD) asms') |> implies_intr_list asms' end; (* term *) val termI = store_standard_thm (Binding.concealed (Binding.make ("termI", \<^here>))) (Thm.equal_elim (Thm.symmetric term_def) (Thm.forall_intr A (Thm.trivial A))); fun mk_term ct = let val cT = Thm.ctyp_of_cterm ct; val T = Thm.typ_of cT; in Thm.instantiate ([((("'a", 0), []), cT)], [((("x", 0), T), ct)]) termI end; fun dest_term th = let val cprop = strip_imp_concl (Thm.cprop_of th) in if can Logic.dest_term (Thm.term_of cprop) then Thm.dest_arg cprop else raise THM ("dest_term", 0, [th]) end; fun cterm_rule f = dest_term o f o mk_term; val cterm_add_frees = Thm.add_frees o mk_term; val cterm_add_vars = Thm.add_vars o mk_term; val dummy_thm = mk_term (certify Term.dummy_prop); val free_dummy_thm = Thm.tag_free_dummy dummy_thm; (* sort_constraint *) fun is_sort_constraint (Const ("Pure.sort_constraint", _) $ Const ("Pure.type", _)) = true | is_sort_constraint _ = false; val sort_constraintI = store_standard_thm (Binding.concealed (Binding.make ("sort_constraintI", \<^here>))) (Thm.equal_elim (Thm.symmetric sort_constraint_def) (mk_term T)); val sort_constraint_eq = store_standard_thm (Binding.concealed (Binding.make ("sort_constraint_eq", \<^here>))) (Thm.equal_intr (Thm.implies_intr CA (Thm.implies_elim (Thm.assume CA) (Thm.unvarify_global (Context.the_global_context ()) sort_constraintI))) (implies_intr_list [A, C] (Thm.assume A))); end; (* HHF normalization *) (* (PROP ?phi \ (\x. PROP ?psi x)) \ (\x. PROP ?phi \ PROP ?psi x) *) val norm_hhf_eq = let val aT = TFree ("'a", []); val x = Free ("x", aT); val phi = Free ("phi", propT); val psi = Free ("psi", aT --> propT); val cx = certify x; val cphi = certify phi; val lhs = certify (Logic.mk_implies (phi, Logic.all x (psi $ x))); val rhs = certify (Logic.all x (Logic.mk_implies (phi, psi $ x))); in Thm.equal_intr (Thm.implies_elim (Thm.assume lhs) (Thm.assume cphi) |> Thm.forall_elim cx |> Thm.implies_intr cphi |> Thm.forall_intr cx |> Thm.implies_intr lhs) (Thm.implies_elim (Thm.assume rhs |> Thm.forall_elim cx) (Thm.assume cphi) |> Thm.forall_intr cx |> Thm.implies_intr cphi |> Thm.implies_intr rhs) |> store_standard_thm_open (Binding.make ("norm_hhf_eq", \<^here>)) end; val norm_hhf_prop = Logic.dest_equals (Thm.prop_of norm_hhf_eq); val norm_hhf_eqs = [norm_hhf_eq, sort_constraint_eq]; fun is_norm_hhf (Const ("Pure.sort_constraint", _)) = false | is_norm_hhf (Const ("Pure.imp", _) $ _ $ (Const ("Pure.all", _) $ _)) = false | is_norm_hhf (Abs _ $ _) = false | is_norm_hhf (t $ u) = is_norm_hhf t andalso is_norm_hhf u | is_norm_hhf (Abs (_, _, t)) = is_norm_hhf t | is_norm_hhf _ = true; fun norm_hhf thy t = if is_norm_hhf t then t else Pattern.rewrite_term thy [norm_hhf_prop] [] t; fun norm_hhf_cterm ctxt raw_ct = let val thy = Proof_Context.theory_of ctxt; val ct = Thm.transfer_cterm thy raw_ct; val t = Thm.term_of ct; in if is_norm_hhf t then ct else Thm.cterm_of ctxt (norm_hhf thy t) end; (* var indexes *) fun incr_indexes th = Thm.incr_indexes (Thm.maxidx_of th + 1); fun incr_indexes2 th1 th2 = Thm.incr_indexes (Int.max (Thm.maxidx_of th1, Thm.maxidx_of th2) + 1); local (*compose Q and \Q1,Q2,...,Qk\ \ R to \Q2,...,Qk\ \ R getting unique result*) fun comp incremented th1 th2 = Thm.bicompose NONE {flatten = true, match = false, incremented = incremented} (false, th1, 0) 1 th2 |> Seq.list_of |> distinct Thm.eq_thm |> (fn [th] => Thm.solve_constraints th | _ => raise THM ("COMP", 1, [th1, th2])); in fun th1 COMP th2 = comp false th1 th2; fun th1 INCR_COMP th2 = comp true (incr_indexes th2 th1) th2; fun th1 COMP_INCR th2 = comp true th1 (incr_indexes th1 th2); end; fun comp_no_flatten (th, n) i rule = (case distinct Thm.eq_thm (Seq.list_of (Thm.bicompose NONE {flatten = false, match = false, incremented = true} (false, th, n) i (incr_indexes th rule))) of [th'] => Thm.solve_constraints th' | [] => raise THM ("comp_no_flatten", i, [th, rule]) | _ => raise THM ("comp_no_flatten: unique result expected", i, [th, rule])); (** variations on Thm.instantiate **) fun instantiate_normalize instpair th = Thm.adjust_maxidx_thm ~1 (Thm.instantiate instpair th COMP_INCR asm_rl); fun instantiate'_normalize Ts ts th = Thm.adjust_maxidx_thm ~1 (Thm.instantiate' Ts ts th COMP_INCR asm_rl); (*instantiation with type-inference for variables*) fun infer_instantiate_types _ [] th = th | infer_instantiate_types ctxt args raw_th = let val thy = Proof_Context.theory_of ctxt; val th = Thm.transfer thy raw_th; fun infer ((xi, T), cu) (tyenv, maxidx) = let val _ = Thm.ctyp_of ctxt T; val _ = Thm.transfer_cterm thy cu; val U = Thm.typ_of_cterm cu; val maxidx' = maxidx |> Integer.max (#2 xi) |> Term.maxidx_typ T |> Integer.max (Thm.maxidx_of_cterm cu); val (tyenv', maxidx'') = Sign.typ_unify thy (T, U) (tyenv, maxidx') handle Type.TUNIFY => let val t = Var (xi, T); val u = Thm.term_of cu; in raise THM ("infer_instantiate_types: type " ^ Syntax.string_of_typ ctxt (Envir.norm_type tyenv T) ^ " of variable " ^ Syntax.string_of_term ctxt (Term.map_types (Envir.norm_type tyenv) t) ^ "\ncannot be unified with type " ^ Syntax.string_of_typ ctxt (Envir.norm_type tyenv U) ^ " of term " ^ Syntax.string_of_term ctxt (Term.map_types (Envir.norm_type tyenv) u), 0, [th]) end; in (tyenv', maxidx'') end; val (tyenv, _) = fold infer args (Vartab.empty, 0); val instT = Vartab.fold (fn (xi, (S, T)) => cons ((xi, S), Thm.ctyp_of ctxt (Envir.norm_type tyenv T))) tyenv []; val inst = args |> map (fn ((xi, T), cu) => ((xi, Envir.norm_type tyenv T), Thm.instantiate_cterm (instT, []) (Thm.transfer_cterm thy cu))); in instantiate_normalize (instT, inst) th end handle CTERM (msg, _) => raise THM (msg, 0, [raw_th]) | TERM (msg, _) => raise THM (msg, 0, [raw_th]) | TYPE (msg, _, _) => raise THM (msg, 0, [raw_th]); fun infer_instantiate _ [] th = th | infer_instantiate ctxt args th = let val vars = Term.add_vars (Thm.full_prop_of th) []; val dups = duplicates (eq_fst op =) vars; val _ = null dups orelse raise THM ("infer_instantiate: inconsistent types for variables " ^ commas_quote (map (Syntax.string_of_term (Config.put show_types true ctxt) o Var) dups), 0, [th]); val args' = args |> map_filter (fn (xi, cu) => AList.lookup (op =) vars xi |> Option.map (fn T => ((xi, T), cu))); in infer_instantiate_types ctxt args' th end; fun infer_instantiate' ctxt args th = let val vars = rev (Term.add_vars (Thm.full_prop_of th) []); val args' = zip_options vars args handle ListPair.UnequalLengths => raise THM ("infer_instantiate': more instantiations than variables in thm", 0, [th]); in infer_instantiate_types ctxt args' th end; (** renaming of bound variables **) (* replace bound variables x_i in thm by y_i *) (* where vs = [(x_1, y_1), ..., (x_n, y_n)] *) fun rename_bvars [] thm = thm | rename_bvars vs thm = let fun rename (Abs (x, T, t)) = Abs (AList.lookup (op =) vs x |> the_default x, T, rename t) | rename (t $ u) = rename t $ rename u | rename a = a; in Thm.renamed_prop (rename (Thm.prop_of thm)) thm end; (* renaming in left-to-right order *) fun rename_bvars' xs thm = let fun rename [] t = ([], t) | rename (x' :: xs) (Abs (x, T, t)) = let val (xs', t') = rename xs t in (xs', Abs (the_default x x', T, t')) end | rename xs (t $ u) = let val (xs', t') = rename xs t; val (xs'', u') = rename xs' u; in (xs'', t' $ u') end | rename xs a = (xs, a); in (case rename xs (Thm.prop_of thm) of ([], prop') => Thm.renamed_prop prop' thm | _ => error "More names than abstractions in theorem") end; end; structure Basic_Drule: BASIC_DRULE = Drule; open Basic_Drule; diff --git a/src/Pure/facts.ML b/src/Pure/facts.ML --- a/src/Pure/facts.ML +++ b/src/Pure/facts.ML @@ -1,299 +1,300 @@ (* Title: Pure/facts.ML Author: Makarius Environment of named facts, optionally indexed by proposition. *) signature FACTS = sig val err_selection: string * Position.T -> int -> thm list -> 'a val err_single: string * Position.T -> thm list -> 'a val the_single: string * Position.T -> thm list -> thm datatype interval = FromTo of int * int | From of int | Single of int datatype ref = Named of (string * Position.T) * interval list option | Fact of string val named: string -> ref val ref_name: ref -> string val ref_pos: ref -> Position.T val map_ref_name: (string -> string) -> ref -> ref val string_of_selection: interval list option -> string val string_of_ref: ref -> string val select: ref -> thm list -> thm list val selections: string * thm list -> (ref * thm) list type T val empty: T val space_of: T -> Name_Space.T val alias: Name_Space.naming -> binding -> string -> T -> T val is_concealed: T -> string -> bool val check: Context.generic -> T -> xstring * Position.T -> string val intern: T -> xstring -> string val extern: Proof.context -> T -> string -> xstring val markup_extern: Proof.context -> T -> string -> Markup.T * xstring val defined: T -> string -> bool val is_dynamic: T -> string -> bool val lookup: Context.generic -> T -> string -> {dynamic: bool, thms: thm list} option val retrieve: Context.generic -> T -> xstring * Position.T -> {name: string, dynamic: bool, thms: thm list} val fold_static: (string * thm list -> 'a -> 'a) -> T -> 'a -> 'a val dest_static: bool -> T list -> T -> (string * thm list) list val dest_all: Context.generic -> bool -> T list -> T -> (string * thm list) list val props: T -> (thm * Position.T) list val could_unify: T -> term -> (thm * Position.T) list val merge: T * T -> T val add_static: Context.generic -> {strict: bool, index: bool} -> binding * thm list lazy -> T -> string * T val add_dynamic: Context.generic -> binding * (Context.generic -> thm list) -> T -> string * T val del: string -> T -> T val hide: bool -> string -> T -> T end; structure Facts: FACTS = struct (** fact references **) fun length_msg (thms: thm list) pos = " (length " ^ string_of_int (length thms) ^ ")" ^ Position.here pos; fun err_selection (name, pos) i thms = error ("Bad fact selection " ^ quote (name ^ enclose "(" ")" (string_of_int i)) ^ length_msg thms pos); fun err_single (name, pos) thms = error ("Expected singleton fact " ^ quote name ^ length_msg thms pos); fun the_single _ [thm] = thm | the_single name_pos thms = err_single name_pos thms; (* datatype interval *) datatype interval = FromTo of int * int | From of int | Single of int; fun string_of_interval (FromTo (i, j)) = string_of_int i ^ "-" ^ string_of_int j | string_of_interval (From i) = string_of_int i ^ "-" | string_of_interval (Single i) = string_of_int i; fun interval n iv = let fun err () = raise Fail ("Bad interval specification " ^ string_of_interval iv) in (case iv of FromTo (i, j) => if i <= j then i upto j else err () | From i => if i <= n then i upto n else err () | Single i => [i]) end; (* datatype ref *) datatype ref = Named of (string * Position.T) * interval list option | Fact of string; fun named name = Named ((name, Position.none), NONE); fun ref_name (Named ((name, _), _)) = name | ref_name (Fact _) = raise Fail "Illegal literal fact"; fun ref_pos (Named ((_, pos), _)) = pos | ref_pos (Fact _) = Position.none; fun map_ref_name f (Named ((name, pos), is)) = Named ((f name, pos), is) | map_ref_name _ r = r; fun string_of_selection NONE = "" | string_of_selection (SOME is) = enclose "(" ")" (commas (map string_of_interval is)); fun string_of_ref (Named ((name, _), sel)) = name ^ string_of_selection sel | string_of_ref (Fact _) = raise Fail "Illegal literal fact"; (* select *) fun select (Fact _) ths = ths | select (Named (_, NONE)) ths = ths | select (Named ((name, pos), SOME ivs)) ths = let val n = length ths; fun err msg = error (msg ^ " for fact " ^ quote name ^ length_msg ths pos); fun sel i = if i > 0 andalso i <= n then nth ths (i - 1) else err_selection (name, pos) i ths; val is = maps (interval n) ivs handle Fail msg => err msg; in map sel is end; (* selections *) fun selections (name, [th]) = [(Named ((name, Position.none), NONE), th)] | selections (name, ths) = map2 (fn i => fn th => (Named ((name, Position.none), SOME [Single i]), th)) (1 upto length ths) ths; (** fact environment **) (* datatypes *) datatype fact = Static of thm list lazy | Dynamic of Context.generic -> thm list; datatype T = Facts of {facts: fact Name_Space.table, props: (thm * Position.T) Net.net}; fun make_facts facts props = Facts {facts = facts, props = props}; val empty = make_facts (Name_Space.empty_table Markup.factN) Net.empty; (* named facts *) fun facts_of (Facts {facts, ...}) = facts; val space_of = Name_Space.space_of_table o facts_of; fun alias naming binding name (Facts {facts, props}) = make_facts (Name_Space.alias_table naming binding name facts) props; val is_concealed = Name_Space.is_concealed o space_of; fun check context facts (xname, pos) = let val (name, fact) = Name_Space.check context (facts_of facts) (xname, pos); val _ = (case fact of Static _ => () | Dynamic _ => Context_Position.report_generic context pos (Markup.dynamic_fact name)); in name end; val intern = Name_Space.intern o space_of; fun extern ctxt = Name_Space.extern ctxt o space_of; fun markup_extern ctxt = Name_Space.markup_extern ctxt o space_of; (* retrieve *) val defined = Name_Space.defined o facts_of; fun is_dynamic facts name = (case Name_Space.lookup (facts_of facts) name of SOME (Dynamic _) => true | _ => false); fun lookup context facts name = (case Name_Space.lookup (facts_of facts) name of NONE => NONE | SOME (Static ths) => SOME {dynamic = false, thms = Lazy.force ths} | SOME (Dynamic f) => SOME {dynamic = true, thms = f context}); fun retrieve context facts (xname, pos) = let val name = check context facts (xname, pos); val {dynamic, thms} = (case lookup context facts name of SOME res => (if #dynamic res then Context_Position.report_generic context pos (Markup.dynamic_fact name) else (); res) | NONE => error ("Unknown fact " ^ quote name ^ Position.here pos)); in {name = name, dynamic = dynamic, thms = map (Thm.transfer'' context) thms} end; (* content *) local fun fold_static_lazy f = Name_Space.fold_table (fn (a, Static ths) => f (a, ths) | _ => I) o facts_of; fun consolidate facts = let val unfinished = - (facts, []) |-> fold_static_lazy (fn (_, ths) => if Lazy.is_finished ths then I else cons ths); + build (facts |> fold_static_lazy (fn (_, ths) => + if Lazy.is_finished ths then I else cons ths)); val _ = Lazy.consolidate unfinished; in facts end; fun included verbose prev_facts facts name = not (exists (fn prev => defined prev name) prev_facts orelse not verbose andalso is_concealed facts name); in fun fold_static f facts = fold_static_lazy (f o apsnd Lazy.force) (consolidate facts); fun dest_static verbose prev_facts facts = fold_static (fn (a, ths) => included verbose prev_facts facts a ? cons (a, ths)) facts [] |> sort_by #1; fun dest_all context verbose prev_facts facts = (facts_of (consolidate facts), []) |-> Name_Space.fold_table (fn (a, fact) => let val ths = (case fact of Static ths => Lazy.force ths | Dynamic f => f context) in included verbose prev_facts facts a ? cons (a, ths) end) |> sort_by #1; end; (* indexed props *) val prop_ord = Term_Ord.term_ord o apply2 (Thm.full_prop_of o fst); fun props (Facts {props, ...}) = sort_distinct prop_ord (Net.content props); fun could_unify (Facts {props, ...}) = Net.unify_term props; (* merge facts *) fun merge (Facts {facts = facts1, props = props1}, Facts {facts = facts2, props = props2}) = let val facts' = Name_Space.merge_tables (facts1, facts2); val props' = if Net.is_empty props2 then props1 else if Net.is_empty props1 then props2 else Net.merge (is_equal o prop_ord) (props1, props2); (*beware of non-canonical merge*) in make_facts facts' props' end; (* add entries *) fun add_prop pos th = Net.insert_term (K false) (Thm.full_prop_of th, (th, pos)); fun add_static context {strict, index} (b, ths) (Facts {facts, props}) = let val ths' = ths |> index ? Lazy.force_value |> Lazy.map_finished (map Thm.trim_context); val (name, facts') = if Binding.is_empty b then ("", facts) else Name_Space.define context strict (b, Static ths') facts; val props' = if index then props |> fold (add_prop (Binding.pos_of (Binding.default_pos b))) (Lazy.force ths') else props; in (name, make_facts facts' props') end; fun add_dynamic context (b, f) (Facts {facts, props}) = let val (name, facts') = Name_Space.define context true (b, Dynamic f) facts; in (name, make_facts facts' props) end; (* remove entries *) fun del name (Facts {facts, props}) = make_facts (Name_Space.del_table name facts) props; fun hide fully name (Facts {facts, props}) = make_facts (Name_Space.hide_table fully name facts) props; end; diff --git a/src/Pure/library.ML b/src/Pure/library.ML --- a/src/Pure/library.ML +++ b/src/Pure/library.ML @@ -1,1101 +1,1106 @@ (* Title: Pure/library.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Markus Wenzel, TU Muenchen Basic library: functions, pairs, booleans, lists, integers, strings, lists as sets, orders, current directory, misc. See also General/basics.ML for the most fundamental concepts. *) infixr 0 ||| infix 2 ? infix 3 o oo ooo oooo infix 4 ~~ upto downto infix orf andf signature BASIC_LIBRARY = sig (*functions*) val undefined: 'a -> 'b val I: 'a -> 'a val K: 'a -> 'b -> 'a val curry: ('a * 'b -> 'c) -> 'a -> 'b -> 'c val uncurry: ('a -> 'b -> 'c) -> 'a * 'b -> 'c val ? : bool * ('a -> 'a) -> 'a -> 'a val oo: ('a -> 'b) * ('c -> 'd -> 'a) -> 'c -> 'd -> 'b val ooo: ('a -> 'b) * ('c -> 'd -> 'e -> 'a) -> 'c -> 'd -> 'e -> 'b val oooo: ('a -> 'b) * ('c -> 'd -> 'e -> 'f -> 'a) -> 'c -> 'd -> 'e -> 'f -> 'b val funpow: int -> ('a -> 'a) -> 'a -> 'a val funpow_yield: int -> ('a -> 'b * 'a) -> 'a -> 'b list * 'a (*pairs*) val pair: 'a -> 'b -> 'a * 'b val rpair: 'a -> 'b -> 'b * 'a val fst: 'a * 'b -> 'a val snd: 'a * 'b -> 'b val eq_fst: ('a * 'c -> bool) -> ('a * 'b) * ('c * 'd) -> bool val eq_snd: ('b * 'd -> bool) -> ('a * 'b) * ('c * 'd) -> bool val eq_pair: ('a * 'c -> bool) -> ('b * 'd -> bool) -> ('a * 'b) * ('c * 'd) -> bool val swap: 'a * 'b -> 'b * 'a val apfst: ('a -> 'b) -> 'a * 'c -> 'b * 'c val apsnd: ('a -> 'b) -> 'c * 'a -> 'c * 'b val apply2: ('a -> 'b) -> 'a * 'a -> 'b * 'b (*booleans*) val equal: ''a -> ''a -> bool val not_equal: ''a -> ''a -> bool val orf: ('a -> bool) * ('a -> bool) -> 'a -> bool val andf: ('a -> bool) * ('a -> bool) -> 'a -> bool val exists: ('a -> bool) -> 'a list -> bool val forall: ('a -> bool) -> 'a list -> bool (*lists*) + val build: ('a list -> 'a list) -> 'a list + val build_rev: ('a list -> 'a list) -> 'a list val single: 'a -> 'a list val the_single: 'a list -> 'a val singleton: ('a list -> 'b list) -> 'a -> 'b val yield_singleton: ('a list -> 'c -> 'b list * 'c) -> 'a -> 'c -> 'b * 'c val perhaps_apply: ('a -> 'a option) list -> 'a -> 'a option val perhaps_loop: ('a -> 'a option) -> 'a -> 'a option val foldl1: ('a * 'a -> 'a) -> 'a list -> 'a val foldr1: ('a * 'a -> 'a) -> 'a list -> 'a val eq_list: ('a * 'a -> bool) -> 'a list * 'a list -> bool val maps: ('a -> 'b list) -> 'a list -> 'b list val filter: ('a -> bool) -> 'a list -> 'a list val filter_out: ('a -> bool) -> 'a list -> 'a list val map_filter: ('a -> 'b option) -> 'a list -> 'b list val take: int -> 'a list -> 'a list val drop: int -> 'a list -> 'a list val chop: int -> 'a list -> 'a list * 'a list val chop_groups: int -> 'a list -> 'a list list val nth: 'a list -> int -> 'a val nth_list: 'a list list -> int -> 'a list val nth_map: int -> ('a -> 'a) -> 'a list -> 'a list val nth_drop: int -> 'a list -> 'a list val map_index: (int * 'a -> 'b) -> 'a list -> 'b list val fold_index: (int * 'a -> 'b -> 'b) -> 'a list -> 'b -> 'b val map_range: (int -> 'a) -> int -> 'a list val fold_range: (int -> 'a -> 'a) -> int -> 'a -> 'a val split_last: 'a list -> 'a list * 'a val find_first: ('a -> bool) -> 'a list -> 'a option val find_index: ('a -> bool) -> 'a list -> int val get_first: ('a -> 'b option) -> 'a list -> 'b option val get_index: ('a -> 'b option) -> 'a list -> (int * 'b) option val flat: 'a list list -> 'a list val unflat: 'a list list -> 'b list -> 'b list list val grouped: int -> (('a list -> 'b list) -> 'c list list -> 'd list list) -> ('a -> 'b) -> 'c list -> 'd list val burrow: ('a list -> 'b list) -> 'a list list -> 'b list list val burrow_options: ('a list -> 'b list) -> 'a option list -> 'b option list val fold_burrow: ('a list -> 'c -> 'b list * 'd) -> 'a list list -> 'c -> 'b list list * 'd val separate: 'a -> 'a list -> 'a list val surround: 'a -> 'a list -> 'a list val replicate: int -> 'a -> 'a list val map_product: ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list val fold_product: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c val map2: ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list val fold2: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c val map_split: ('a -> 'b * 'c) -> 'a list -> 'b list * 'c list val zip_options: 'a list -> 'b option list -> ('a * 'b) list val ~~ : 'a list * 'b list -> ('a * 'b) list val split_list: ('a * 'b) list -> 'a list * 'b list val burrow_fst: ('a list -> 'b list) -> ('a * 'c) list -> ('b * 'c) list val take_prefix: ('a -> bool) -> 'a list -> 'a list val drop_prefix: ('a -> bool) -> 'a list -> 'a list val chop_prefix: ('a -> bool) -> 'a list -> 'a list * 'a list val take_suffix: ('a -> bool) -> 'a list -> 'a list val drop_suffix: ('a -> bool) -> 'a list -> 'a list val chop_suffix: ('a -> bool) -> 'a list -> 'a list * 'a list val is_prefix: ('a * 'a -> bool) -> 'a list -> 'a list -> bool val chop_common_prefix: ('a * 'b -> bool) -> 'a list * 'b list -> 'a list * ('a list * 'b list) val prefixes1: 'a list -> 'a list list val prefixes: 'a list -> 'a list list val suffixes1: 'a list -> 'a list list val suffixes: 'a list -> 'a list list val trim: ('a -> bool) -> 'a list -> 'a list (*integers*) val upto: int * int -> int list val downto: int * int -> int list val hex_digit: int -> string val radixpand: int * int -> int list val radixstring: int * string * int -> string val string_of_int: int -> string val signed_string_of_int: int -> string val string_of_indexname: string * int -> string val read_radix_int: int -> string list -> int * string list val read_int: string list -> int * string list val oct_char: string -> string (*strings*) val nth_string: string -> int -> string val fold_string: (string -> 'a -> 'a) -> string -> 'a -> 'a val exists_string: (string -> bool) -> string -> bool val forall_string: (string -> bool) -> string -> bool val member_string: string -> string -> bool val first_field: string -> string -> (string * string) option val enclose: string -> string -> string -> string val unenclose: string -> string val quote: string -> string val cartouche: string -> string val space_implode: string -> string list -> string val commas: string list -> string val commas_quote: string list -> string val cat_lines: string list -> string val space_explode: string -> string -> string list val split_lines: string -> string list val plain_words: string -> string val prefix_lines: string -> string -> string val prefix: string -> string -> string val suffix: string -> string -> string val unprefix: string -> string -> string val unsuffix: string -> string -> string val trim_line: string -> string val trim_split_lines: string -> string list val normalize_lines: string -> string val replicate_string: int -> string -> string val translate_string: (string -> string) -> string -> string val encode_lines: string -> string val decode_lines: string -> string val align_right: string -> int -> string -> string val match_string: string -> string -> bool (*reals*) val string_of_real: real -> string val signed_string_of_real: real -> string (*lists as sets -- see also Pure/General/ord_list.ML*) val member: ('b * 'a -> bool) -> 'a list -> 'b -> bool val insert: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list val remove: ('b * 'a -> bool) -> 'b -> 'a list -> 'a list val update: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list val union: ('a * 'a -> bool) -> 'a list -> 'a list -> 'a list val subtract: ('b * 'a -> bool) -> 'b list -> 'a list -> 'a list val inter: ('a * 'b -> bool) -> 'b list -> 'a list -> 'a list val merge: ('a * 'a -> bool) -> 'a list * 'a list -> 'a list val subset: ('a * 'b -> bool) -> 'a list * 'b list -> bool val eq_set: ('a * 'a -> bool) -> 'a list * 'a list -> bool val distinct: ('a * 'a -> bool) -> 'a list -> 'a list val duplicates: ('a * 'a -> bool) -> 'a list -> 'a list val has_duplicates: ('a * 'a -> bool) -> 'a list -> bool val map_transpose: ('a list -> 'b) -> 'a list list -> 'b list (*lists as multisets*) val remove1: ('b * 'a -> bool) -> 'b -> 'a list -> 'a list val combine: ('a * 'a -> bool) -> 'a list -> 'a list -> 'a list val submultiset: ('a * 'b -> bool) -> 'a list * 'b list -> bool (*orders*) type 'a ord = 'a * 'a -> order val is_equal: order -> bool val is_less: order -> bool val is_less_equal: order -> bool val is_greater: order -> bool val is_greater_equal: order -> bool val rev_order: order -> order val make_ord: ('a * 'a -> bool) -> 'a ord val pointer_eq_ord: ('a * 'a -> order) -> 'a * 'a -> order val bool_ord: bool ord val int_ord: int ord val string_ord: string ord val size_ord: string ord val fast_string_ord: string ord val option_ord: ('a * 'b -> order) -> 'a option * 'b option -> order val ||| : ('a -> order) * ('a -> order) -> 'a -> order val prod_ord: ('a * 'b -> order) -> ('c * 'd -> order) -> ('a * 'c) * ('b * 'd) -> order val dict_ord: ('a * 'b -> order) -> 'a list * 'b list -> order val length_ord: 'a list * 'b list -> order val list_ord: ('a * 'b -> order) -> 'a list * 'b list -> order val sort: 'a ord -> 'a list -> 'a list val sort_distinct: 'a ord -> 'a list -> 'a list val sort_strings: string list -> string list val sort_by: ('a -> string) -> 'a list -> 'a list val tag_list: int -> 'a list -> (int * 'a) list val untag_list: (int * 'a) list -> 'a list val order_list: (int * 'a) list -> 'a list (*misc*) val divide_and_conquer: ('a -> 'a list * ('b list -> 'b)) -> 'a -> 'b val divide_and_conquer': ('a -> 'b -> ('a list * ('c list * 'b -> 'c * 'b)) * 'b) -> 'a -> 'b -> 'c * 'b val partition_eq: ('a * 'a -> bool) -> 'a list -> 'a list list val partition_list: (int -> 'a -> bool) -> int -> int -> 'a list -> 'a list list type serial = int val serial: unit -> serial val serial_string: unit -> string eqtype stamp val stamp: unit -> stamp structure Any: sig type T = exn end val getenv: string -> string val getenv_strict: string -> string end; signature LIBRARY = sig include BASIC_LIBRARY val foldl: ('a * 'b -> 'a) -> 'a * 'b list -> 'a val foldr: ('a * 'b -> 'b) -> 'a list * 'b -> 'b end; structure Library: LIBRARY = struct (* functions *) fun undefined _ = raise Match; fun I x = x; fun K x = fn _ => x; fun curry f x y = f (x, y); fun uncurry f (x, y) = f x y; (*conditional application*) fun b ? f = fn x => if b then f x else x; (*composition with multiple args*) fun (f oo g) x y = f (g x y); fun (f ooo g) x y z = f (g x y z); fun (f oooo g) x y z w = f (g x y z w); (*function exponentiation: f (... (f x) ...) with n applications of f*) fun funpow (0: int) _ x = x | funpow n f x = funpow (n - 1) f (f x); fun funpow_yield (0 : int) _ x = ([], x) | funpow_yield n f x = x |> f ||>> funpow_yield (n - 1) f |>> op ::; (* pairs *) fun pair x y = (x, y); fun rpair x y = (y, x); fun fst (x, y) = x; fun snd (x, y) = y; fun eq_fst eq ((x1, _), (x2, _)) = eq (x1, x2); fun eq_snd eq ((_, y1), (_, y2)) = eq (y1, y2); fun eq_pair eqx eqy ((x1, y1), (x2, y2)) = eqx (x1, x2) andalso eqy (y1, y2); fun swap (x, y) = (y, x); fun apfst f (x, y) = (f x, y); fun apsnd f (x, y) = (x, f y); fun apply2 f (x, y) = (f x, f y); (* booleans *) (*polymorphic equality*) fun equal x y = x = y; fun not_equal x y = x <> y; (*combining predicates*) fun p orf q = fn x => p x orelse q x; fun p andf q = fn x => p x andalso q x; val exists = List.exists; val forall = List.all; (** lists **) +fun build (f: 'a list -> 'a list) = f []; +fun build_rev f = build f |> rev; + fun single x = [x]; fun the_single [x] = x | the_single _ = raise List.Empty; fun singleton f x = the_single (f [x]); fun yield_singleton f x = f [x] #>> the_single; fun perhaps_apply funs arg = let fun app [] res = res | app (f :: fs) (changed, x) = (case f x of NONE => app fs (changed, x) | SOME x' => app fs (true, x')); in (case app funs (false, arg) of (false, _) => NONE | (true, arg') => SOME arg') end; fun perhaps_loop f arg = let fun loop (changed, x) = (case f x of NONE => (changed, x) | SOME x' => loop (true, x')); in (case loop (false, arg) of (false, _) => NONE | (true, arg') => SOME arg') end; (* fold -- old versions *) (*the following versions of fold are designed to fit nicely with infixes*) (* (op @) (e, [x1, ..., xn]) ===> ((e @ x1) @ x2) ... @ xn for operators that associate to the left (TAIL RECURSIVE)*) fun foldl (f: 'a * 'b -> 'a) : 'a * 'b list -> 'a = let fun itl (e, []) = e | itl (e, a::l) = itl (f(e, a), l) in itl end; (* (op @) ([x1, ..., xn], e) ===> x1 @ (x2 ... @ (xn @ e)) for operators that associate to the right (not tail recursive)*) fun foldr f (l, e) = let fun itr [] = e | itr (a::l) = f(a, itr l) in itr l end; (* (op @) [x1, ..., xn] ===> ((x1 @ x2) @ x3) ... @ xn for operators that associate to the left (TAIL RECURSIVE)*) fun foldl1 f [] = raise List.Empty | foldl1 f (x :: xs) = foldl f (x, xs); (* (op @) [x1, ..., xn] ===> x1 @ (x2 ... @ (x[n-1] @ xn)) for n > 0, operators that associate to the right (not tail recursive)*) fun foldr1 f [] = raise List.Empty | foldr1 f l = let fun itr [x] = x | itr (x::l) = f(x, itr l) in itr l end; (* basic list functions *) fun eq_list eq (list1, list2) = pointer_eq (list1, list2) orelse let fun eq_lst (x :: xs, y :: ys) = eq (x, y) andalso eq_lst (xs, ys) | eq_lst _ = true; in length list1 = length list2 andalso eq_lst (list1, list2) end; fun maps f [] = [] | maps f (x :: xs) = f x @ maps f xs; val filter = List.filter; fun filter_out f = filter (not o f); val map_filter = List.mapPartial; fun take (0: int) xs = [] | take _ [] = [] | take n (x :: xs) = x :: take (n - 1) xs; fun drop (0: int) xs = xs | drop _ [] = [] | drop n (x :: xs) = drop (n - 1) xs; fun chop (0: int) xs = ([], xs) | chop _ [] = ([], []) | chop n (x :: xs) = chop (n - 1) xs |>> cons x; fun chop_groups n list = (case chop (Int.max (n, 1)) list of ([], _) => [] | (g, rest) => g :: chop_groups n rest); (*return nth element of a list, where 0 designates the first element; raise Subscript if list too short*) fun nth xs i = List.nth (xs, i); fun nth_list xss i = nth xss i handle General.Subscript => []; fun nth_map 0 f (x :: xs) = f x :: xs | nth_map n f (x :: xs) = x :: nth_map (n - 1) f xs | nth_map (_: int) _ [] = raise Subscript; fun nth_drop n xs = List.take (xs, n) @ List.drop (xs, n + 1); fun map_index f = let fun map_aux (_: int) [] = [] | map_aux i (x :: xs) = f (i, x) :: map_aux (i + 1) xs in map_aux 0 end; fun fold_index f = let fun fold_aux (_: int) [] y = y | fold_aux i (x :: xs) y = fold_aux (i + 1) xs (f (i, x) y) in fold_aux 0 end; fun map_range f i = let fun map_aux (k: int) = if k < i then f k :: map_aux (k + 1) else [] in map_aux 0 end; fun fold_range f i = let fun fold_aux (k: int) y = if k < i then fold_aux (k + 1) (f k y) else y in fold_aux 0 end; (*rear decomposition*) fun split_last [] = raise List.Empty | split_last [x] = ([], x) | split_last (x :: xs) = apfst (cons x) (split_last xs); (*find first element satisfying predicate*) val find_first = List.find; (*find position of first element satisfying a predicate*) fun find_index pred = let fun find (_: int) [] = ~1 | find n (x :: xs) = if pred x then n else find (n + 1) xs; in find 0 end; (*get first element by lookup function*) fun get_first _ [] = NONE | get_first f (x :: xs) = (case f x of NONE => get_first f xs | some => some); fun get_index f = let fun get_aux (_: int) [] = NONE | get_aux i (x :: xs) = (case f x of NONE => get_aux (i + 1) xs | SOME y => SOME (i, y)) in get_aux 0 end; val flat = List.concat; fun unflat (xs :: xss) ys = let val (ps, qs) = chop (length xs) ys in ps :: unflat xss qs end | unflat [] [] = [] | unflat _ _ = raise ListPair.UnequalLengths; fun grouped n comb f = chop_groups n #> comb (map f) #> flat; fun burrow f xss = unflat xss (f (flat xss)); fun burrow_options f os = map (try hd) (burrow f (map the_list os)); fun fold_burrow f xss s = apfst (unflat xss) (f (flat xss) s); (*separate s [x1, x2, ..., xn] ===> [x1, s, x2, s, ..., s, xn]*) fun separate s (x :: (xs as _ :: _)) = x :: s :: separate s xs | separate _ xs = xs; fun surround s (x :: xs) = s :: x :: surround s xs | surround s [] = [s]; (*make the list [x, x, ..., x] of length n*) fun replicate (n: int) x = let fun rep (0, xs) = xs | rep (n, xs) = rep (n - 1, x :: xs) in if n < 0 then raise Subscript else rep (n, []) end; (* direct product *) fun map_product f _ [] = [] | map_product f [] _ = [] | map_product f (x :: xs) ys = map (f x) ys @ map_product f xs ys; fun fold_product f _ [] z = z | fold_product f [] _ z = z | fold_product f (x :: xs) ys z = z |> fold (f x) ys |> fold_product f xs ys; (* lists of pairs *) fun map2 _ [] [] = [] | map2 f (x :: xs) (y :: ys) = f x y :: map2 f xs ys | map2 _ _ _ = raise ListPair.UnequalLengths; fun fold2 _ [] [] z = z | fold2 f (x :: xs) (y :: ys) z = fold2 f xs ys (f x y z) | fold2 _ _ _ _ = raise ListPair.UnequalLengths; fun map_split _ [] = ([], []) | map_split f (x :: xs) = let val (y, w) = f x; val (ys, ws) = map_split f xs; in (y :: ys, w :: ws) end; fun zip_options (x :: xs) (SOME y :: ys) = (x, y) :: zip_options xs ys | zip_options (_ :: xs) (NONE :: ys) = zip_options xs ys | zip_options _ [] = [] | zip_options [] _ = raise ListPair.UnequalLengths; (*combine two lists forming a list of pairs: [x1, ..., xn] ~~ [y1, ..., yn] ===> [(x1, y1), ..., (xn, yn)]*) fun [] ~~ [] = [] | (x :: xs) ~~ (y :: ys) = (x, y) :: (xs ~~ ys) | _ ~~ _ = raise ListPair.UnequalLengths; (*inverse of ~~; the old 'split': [(x1, y1), ..., (xn, yn)] ===> ([x1, ..., xn], [y1, ..., yn])*) val split_list = ListPair.unzip; fun burrow_fst f xs = split_list xs |>> f |> op ~~; (* take, drop, chop, trim according to predicate *) fun take_prefix pred list = let fun take res (x :: xs) = if pred x then take (x :: res) xs else rev res | take res [] = rev res; in take [] list end; fun drop_prefix pred list = let fun drop (x :: xs) = if pred x then drop xs else x :: xs | drop [] = []; in drop list end; fun chop_prefix pred list = let val prfx = take_prefix pred list; val sffx = drop (length prfx) list; in (prfx, sffx) end; fun take_suffix pred list = let fun take res (x :: xs) = if pred x then take (x :: res) xs else res | take res [] = res; in take [] (rev list) end; fun drop_suffix pred list = let fun drop (x :: xs) = if pred x then drop xs else rev (x :: xs) | drop [] = []; in drop (rev list) end; fun chop_suffix pred list = let val prfx = drop_suffix pred list; val sffx = drop (length prfx) list; in (prfx, sffx) end; fun trim pred = drop_prefix pred #> drop_suffix pred; (* prefixes, suffixes *) fun is_prefix _ [] _ = true | is_prefix eq (x :: xs) (y :: ys) = eq (x, y) andalso is_prefix eq xs ys | is_prefix eq _ _ = false; fun chop_common_prefix eq ([], ys) = ([], ([], ys)) | chop_common_prefix eq (xs, []) = ([], (xs, [])) | chop_common_prefix eq (xs as x :: xs', ys as y :: ys') = if eq (x, y) then let val (ps', xys'') = chop_common_prefix eq (xs', ys') in (x :: ps', xys'') end else ([], (xs, ys)); fun prefixes1 [] = [] | prefixes1 (x :: xs) = map (cons x) ([] :: prefixes1 xs); fun prefixes xs = [] :: prefixes1 xs; fun suffixes1 xs = map rev (prefixes1 (rev xs)); fun suffixes xs = [] :: suffixes1 xs; (** integers **) (* lists of integers *) (*make the list [from, from + 1, ..., to]*) fun ((i: int) upto j) = if i > j then [] else i :: (i + 1 upto j); (*make the list [from, from - 1, ..., to]*) fun ((i: int) downto j) = if i < j then [] else i :: (i - 1 downto j); (* convert integers to strings *) (*hexadecimal*) fun hex_digit i = if i < 10 then chr (Char.ord #"0" + i) else chr (Char.ord #"a" + i - 10); (*expand the number in the given base; example: radixpand (2, 8) gives [1, 0, 0, 0]*) fun radixpand (base, num) : int list = let fun radix (n, tail) = if n < base then n :: tail else radix (n div base, (n mod base) :: tail) in radix (num, []) end; (*expands a number into a string of characters starting from "zerochar"; example: radixstring (2, "0", 8) gives "1000"*) fun radixstring (base, zerochar, num) = let val offset = ord zerochar; fun chrof n = chr (offset + n) in implode (map chrof (radixpand (base, num))) end; local val zero = Char.ord #"0"; val small_int = 10000: int; val small_int_table = Vector.tabulate (small_int, Int.toString); in fun string_of_int i = if i < 0 then Int.toString i else if i < 10 then chr (zero + i) else if i < small_int then Vector.sub (small_int_table, i) else Int.toString i; end; fun signed_string_of_int i = if i < 0 then "-" ^ string_of_int (~ i) else string_of_int i; fun string_of_indexname (a, 0) = a | string_of_indexname (a, i) = a ^ "_" ^ string_of_int i; (* read integers *) fun read_radix_int radix cs = let val zero = Char.ord #"0"; val limit = zero + radix; fun scan (num, []) = (num, []) | scan (num, c :: cs) = if zero <= ord c andalso ord c < limit then scan (radix * num + (ord c - zero), cs) else (num, c :: cs); in scan (0, cs) end; val read_int = read_radix_int 10; fun oct_char s = chr (#1 (read_radix_int 8 (raw_explode s))); (** strings **) (* functions tuned for strings, avoiding explode *) fun nth_string str i = (case try String.substring (str, i, 1) of SOME s => s | NONE => raise Subscript); fun fold_string f str x0 = let val n = size str; fun iter (x, i) = if i < n then iter (f (String.substring (str, i, 1)) x, i + 1) else x; in iter (x0, 0) end; fun exists_string pred str = let val n = size str; fun ex i = i < n andalso (pred (String.substring (str, i, 1)) orelse ex (i + 1)); in ex 0 end; fun forall_string pred = not o exists_string (not o pred); fun member_string str s = exists_string (fn s' => s = s') str; fun first_field sep str = let val n = size sep; val len = size str; fun find i = if i + n > len then NONE else if String.substring (str, i, n) = sep then SOME i else find (i + 1); in (case find 0 of NONE => NONE | SOME i => SOME (String.substring (str, 0, i), String.extract (str, i + n, NONE))) end; (*enclose in brackets*) fun enclose lpar rpar str = lpar ^ str ^ rpar; fun unenclose str = String.substring (str, 1, size str - 2); (*simple quoting (does not escape special chars)*) val quote = enclose "\"" "\""; val cartouche = enclose "\" "\"; val space_implode = String.concatWith; val commas = space_implode ", "; val commas_quote = commas o map quote; val cat_lines = space_implode "\n"; (*space_explode "." "h.e..l.lo" = ["h", "e", "", "l", "lo"]*) fun space_explode _ "" = [] | space_explode sep s = String.fields (fn c => str c = sep) s; val split_lines = space_explode "\n"; fun plain_words s = space_explode "_" s |> space_implode " "; fun prefix_lines "" txt = txt | prefix_lines prfx txt = txt |> split_lines |> map (fn s => prfx ^ s) |> cat_lines; fun prefix prfx s = prfx ^ s; fun suffix sffx s = s ^ sffx; fun unprefix prfx s = if String.isPrefix prfx s then String.substring (s, size prfx, size s - size prfx) else raise Fail "unprefix"; fun unsuffix sffx s = if String.isSuffix sffx s then String.substring (s, 0, size s - size sffx) else raise Fail "unsuffix"; fun trim_line s = if String.isSuffix "\r\n" s then String.substring (s, 0, size s - 2) else if String.isSuffix "\r" s orelse String.isSuffix "\n" s then String.substring (s, 0, size s - 1) else s; val trim_split_lines = trim_line #> split_lines #> map trim_line; fun normalize_lines str = if exists_string (fn s => s = "\r") str then split_lines str |> map trim_line |> cat_lines else str; fun replicate_string (0: int) _ = "" | replicate_string 1 a = a | replicate_string k a = if k mod 2 = 0 then replicate_string (k div 2) (a ^ a) else replicate_string (k div 2) (a ^ a) ^ a; fun translate_string f = String.translate (f o String.str); val encode_lines = translate_string (fn "\n" => "\v" | c => c); val decode_lines = translate_string (fn "\v" => "\n" | c => c); fun align_right c k s = let val _ = if size c <> 1 orelse size s > k then raise Fail "align_right" else () in replicate_string (k - size s) c ^ s end; (*crude matching of str against simple glob pat*) fun match_string pat str = let fun match [] _ = true | match (p :: ps) s = size p <= size s andalso (case try (unprefix p) s of SOME s' => match ps s' | NONE => match (p :: ps) (String.substring (s, 1, size s - 1))); in match (space_explode "*" pat) str end; (** reals **) val string_of_real = Real.fmt (StringCvt.GEN NONE); fun signed_string_of_real x = if x < 0.0 then "-" ^ string_of_real (~ x) else string_of_real x; (** lists as sets -- see also Pure/General/ord_list.ML **) (* canonical operations *) fun member eq list x = let fun memb [] = false | memb (y :: ys) = eq (x, y) orelse memb ys; in memb list end; fun insert eq x xs = if member eq xs x then xs else x :: xs; fun remove eq x xs = if member eq xs x then filter_out (fn y => eq (x, y)) xs else xs; fun update eq x xs = cons x (remove eq x xs); fun inter eq xs = filter (member eq xs); fun union eq = fold (insert eq); fun subtract eq = fold (remove eq); fun merge eq (xs, ys) = if pointer_eq (xs, ys) then xs else if null xs then ys else fold_rev (insert eq) ys xs; (* subset and set equality *) fun subset eq (xs, ys) = forall (member eq ys) xs; fun eq_set eq (xs, ys) = eq_list eq (xs, ys) orelse (subset eq (xs, ys) andalso subset (eq o swap) (ys, xs)); (*makes a list of the distinct members of the input; preserves order, takes first of equal elements*) fun distinct eq lst = let fun dist (rev_seen, []) = rev rev_seen | dist (rev_seen, x :: xs) = if member eq rev_seen x then dist (rev_seen, xs) else dist (x :: rev_seen, xs); in dist ([], lst) end; (*returns a list containing all repeated elements exactly once; preserves order, takes first of equal elements*) fun duplicates eq lst = let fun dups (rev_dups, []) = rev rev_dups | dups (rev_dups, x :: xs) = if member eq rev_dups x orelse not (member eq xs x) then dups (rev_dups, xs) else dups (x :: rev_dups, xs); in dups ([], lst) end; fun has_duplicates eq = let fun dups [] = false | dups (x :: xs) = member eq xs x orelse dups xs; in dups end; (* matrices *) fun map_transpose f xss = let val n = (case distinct (op =) (map length xss) of [] => 0 | [n] => n | _ => raise ListPair.UnequalLengths); in map_range (fn m => f (map (fn xs => nth xs m) xss)) n end; (** lists as multisets **) fun remove1 eq x [] = [] | remove1 eq x (y :: ys) = if eq (x, y) then ys else y :: remove1 eq x ys; fun combine eq xs ys = fold (remove1 eq) ys xs @ ys; fun submultiset _ ([], _) = true | submultiset eq (x :: xs, ys) = member eq ys x andalso submultiset eq (xs, remove1 eq x ys); (** orders **) type 'a ord = 'a * 'a -> order; fun is_equal ord = ord = EQUAL; fun is_less ord = ord = LESS; fun is_less_equal ord = ord = LESS orelse ord = EQUAL; fun is_greater ord = ord = GREATER; fun is_greater_equal ord = ord = GREATER orelse ord = EQUAL; fun rev_order LESS = GREATER | rev_order EQUAL = EQUAL | rev_order GREATER = LESS; (*compose orders*) fun (a_ord ||| b_ord) p = (case a_ord p of EQUAL => b_ord p | ord => ord); (*assume rel is a linear strict order*) fun make_ord rel (x, y) = if rel (x, y) then LESS else if rel (y, x) then GREATER else EQUAL; fun pointer_eq_ord ord (x, y) = if pointer_eq (x, y) then EQUAL else ord (x, y); fun bool_ord (false, true) = LESS | bool_ord (true, false) = GREATER | bool_ord _ = EQUAL; val int_ord = Int.compare; val string_ord = String.compare; val size_ord = int_ord o apply2 size; val fast_string_ord = pointer_eq_ord (size_ord ||| string_ord); fun option_ord ord (SOME x, SOME y) = ord (x, y) | option_ord _ (NONE, NONE) = EQUAL | option_ord _ (NONE, SOME _) = LESS | option_ord _ (SOME _, NONE) = GREATER; (*lexicographic product*) fun prod_ord a_ord b_ord ((x, y), (x', y')) = (case a_ord (x, x') of EQUAL => b_ord (y, y') | ord => ord); (*dictionary order -- in general NOT well-founded!*) fun dict_ord elem_ord (x :: xs, y :: ys) = (case elem_ord (x, y) of EQUAL => dict_ord elem_ord (xs, ys) | ord => ord) | dict_ord _ ([], []) = EQUAL | dict_ord _ ([], _ :: _) = LESS | dict_ord _ (_ :: _, []) = GREATER; (*lexicographic product of lists*) fun length_ord (xs, ys) = int_ord (length xs, length ys); fun list_ord elem_ord = length_ord ||| dict_ord elem_ord; (* sorting *) (*stable mergesort -- preserves order of equal elements*) fun mergesort unique ord = let fun merge (xs as x :: xs') (ys as y :: ys') = (case ord (x, y) of LESS => x :: merge xs' ys | EQUAL => if unique then merge xs ys' else x :: merge xs' ys | GREATER => y :: merge xs ys') | merge [] ys = ys | merge xs [] = xs; fun merge_all [xs] = xs | merge_all xss = merge_all (merge_pairs xss) and merge_pairs (xs :: ys :: xss) = merge xs ys :: merge_pairs xss | merge_pairs xss = xss; fun runs (x :: y :: xs) = (case ord (x, y) of LESS => ascending y [x] xs | EQUAL => if unique then runs (x :: xs) else ascending y [x] xs | GREATER => descending y [x] xs) | runs xs = [xs] and ascending x xs (zs as y :: ys) = (case ord (x, y) of LESS => ascending y (x :: xs) ys | EQUAL => if unique then ascending x xs ys else ascending y (x :: xs) ys | GREATER => rev (x :: xs) :: runs zs) | ascending x xs [] = [rev (x :: xs)] and descending x xs (zs as y :: ys) = (case ord (x, y) of GREATER => descending y (x :: xs) ys | EQUAL => if unique then descending x xs ys else (x :: xs) :: runs zs | LESS => (x :: xs) :: runs zs) | descending x xs [] = [x :: xs]; in merge_all o runs end; fun sort ord = mergesort false ord; fun sort_distinct ord = mergesort true ord; val sort_strings = sort string_ord; fun sort_by key xs = sort (string_ord o apply2 key) xs; (* items tagged by integer index *) (*insert tags*) fun tag_list k [] = [] | tag_list k (x :: xs) = (k:int, x) :: tag_list (k + 1) xs; (*remove tags and suppress duplicates -- list is assumed sorted!*) fun untag_list [] = [] | untag_list [(k: int, x)] = [x] | untag_list ((k, x) :: (rest as (k', x') :: _)) = if k = k' then untag_list rest else x :: untag_list rest; (*return list elements in original order*) fun order_list list = untag_list (sort (int_ord o apply2 fst) list); (** misc **) fun divide_and_conquer decomp x = let val (ys, recomb) = decomp x in recomb (map (divide_and_conquer decomp) ys) end; fun divide_and_conquer' decomp x s = let val ((ys, recomb), s') = decomp x s in recomb (fold_map (divide_and_conquer' decomp) ys s') end; (*Partition a list into buckets [ bi, b(i+1), ..., bj ] putting x in bk if p(k)(x) holds. Preserve order of elements if possible.*) fun partition_list p i j = let fun part (k: int) xs = if k > j then (case xs of [] => [] | _ => raise Fail "partition_list") else let val (ns, rest) = List.partition (p k) xs in ns :: part (k + 1) rest end; in part (i: int) end; fun partition_eq (eq: 'a * 'a -> bool) = let fun part [] = [] | part (x :: ys) = let val (xs, xs') = List.partition (fn y => eq (x, y)) ys in (x :: xs) :: part xs' end; in part end; (* serial numbers and abstract stamps *) type serial = int; val serial = Counter.make (); val serial_string = string_of_int o serial; datatype stamp = Stamp of serial; fun stamp () = Stamp (serial ()); (* values of any type *) (*note that the builtin exception datatype may be extended by new constructors at any time*) structure Any = struct type T = exn end; (* getenv *) fun getenv x = (case OS.Process.getEnv x of NONE => "" | SOME y => y); fun getenv_strict x = (case getenv x of "" => error ("Undefined Isabelle environment variable: " ^ quote x) | y => y); end; structure Basic_Library: BASIC_LIBRARY = Library; open Basic_Library; diff --git a/src/Pure/proofterm.ML b/src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML +++ b/src/Pure/proofterm.ML @@ -1,2341 +1,2341 @@ (* Title: Pure/proofterm.ML Author: Stefan Berghofer, TU Muenchen LF style proof terms. *) infix 8 % %% %>; signature PROOFTERM = sig type thm_header = {serial: serial, pos: Position.T list, theory_name: string, name: string, prop: term, types: typ list option} type thm_body type thm_node datatype proof = MinProof | PBound of int | Abst of string * typ option * proof | AbsP of string * term option * proof | % of proof * term option | %% of proof * proof | Hyp of term | PAxm of string * term * typ list option | PClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} type oracle = (string * Position.T) * term option type thm = serial * thm_node exception MIN_PROOF of unit val proof_of: proof_body -> proof val join_proof: proof_body future -> proof val map_proof_of: (proof -> proof) -> proof_body -> proof_body val thm_header: serial -> Position.T list -> string -> string -> term -> typ list option -> thm_header val thm_body: proof_body -> thm_body val thm_body_proof_raw: thm_body -> proof val thm_body_proof_open: thm_body -> proof val thm_node_theory_name: thm_node -> string val thm_node_name: thm_node -> string val thm_node_prop: thm_node -> term val thm_node_body: thm_node -> proof_body future val thm_node_thms: thm_node -> thm list val join_thms: thm list -> proof_body list val make_thm: thm_header -> thm_body -> thm val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a val fold_body_thms: ({serial: serial, name: string, prop: term, body: proof_body} -> 'a -> 'a) -> proof_body list -> 'a -> 'a val oracle_ord: oracle ord val thm_ord: thm ord val unions_oracles: oracle Ord_List.T list -> oracle Ord_List.T val unions_thms: thm Ord_List.T list -> thm Ord_List.T val no_proof_body: proof -> proof_body val no_thm_names: proof -> proof val no_thm_proofs: proof -> proof val no_body_proofs: proof -> proof val encode: Consts.T -> proof XML.Encode.T val encode_body: Consts.T -> proof_body XML.Encode.T val encode_standard_term: Consts.T -> term XML.Encode.T val encode_standard_proof: Consts.T -> proof XML.Encode.T val decode: Consts.T -> proof XML.Decode.T val decode_body: Consts.T -> proof_body XML.Decode.T val %> : proof * term -> proof (*primitive operations*) val proofs: int Unsynchronized.ref val proofs_enabled: unit -> bool val atomic_proof: proof -> bool val compact_proof: proof -> bool val proof_combt: proof * term list -> proof val proof_combt': proof * term option list -> proof val proof_combP: proof * proof list -> proof val strip_combt: proof -> proof * term option list val strip_combP: proof -> proof * proof list val strip_thm_body: proof_body -> proof_body val map_proof_same: term Same.operation -> typ Same.operation -> (typ * class -> proof) -> proof Same.operation val map_proof_terms_same: term Same.operation -> typ Same.operation -> proof Same.operation val map_proof_types_same: typ Same.operation -> proof Same.operation val map_proof_terms: (term -> term) -> (typ -> typ) -> proof -> proof val map_proof_types: (typ -> typ) -> proof -> proof val fold_proof_terms: (term -> 'a -> 'a) -> proof -> 'a -> 'a val fold_proof_terms_types: (term -> 'a -> 'a) -> (typ -> 'a -> 'a) -> proof -> 'a -> 'a val maxidx_proof: proof -> int -> int val size_of_proof: proof -> int val change_types: typ list option -> proof -> proof val prf_abstract_over: term -> proof -> proof val prf_incr_bv: int -> int -> int -> int -> proof -> proof val incr_pboundvars: int -> int -> proof -> proof val prf_loose_bvar1: proof -> int -> bool val prf_loose_Pbvar1: proof -> int -> bool val prf_add_loose_bnos: int -> int -> proof -> int list * int list -> int list * int list val norm_proof: Envir.env -> proof -> proof val norm_proof': Envir.env -> proof -> proof val prf_subst_bounds: term list -> proof -> proof val prf_subst_pbounds: proof list -> proof -> proof val freeze_thaw_prf: proof -> proof * (proof -> proof) (*proof terms for specific inference rules*) val trivial_proof: proof val implies_intr_proof: term -> proof -> proof val implies_intr_proof': term -> proof -> proof val forall_intr_proof: string * term -> typ option -> proof -> proof val forall_intr_proof': term -> proof -> proof val varify_proof: term -> Term_Subst.TFrees.set -> proof -> proof val legacy_freezeT: term -> proof -> proof val rotate_proof: term list -> term -> (string * typ) list -> term list -> int -> proof -> proof val permute_prems_proof: term list -> int -> int -> proof -> proof val generalize_proof: Symtab.set * Symtab.set -> int -> term -> proof -> proof val instantiate: typ Term_Subst.TVars.table * term Term_Subst.Vars.table -> proof -> proof val lift_proof: term -> int -> term -> proof -> proof val incr_indexes: int -> proof -> proof val assumption_proof: term list -> term -> int -> proof -> proof val bicompose_proof: bool -> term list -> term list -> term option -> term list -> int -> int -> proof -> proof -> proof val equality_axms: (string * term) list val reflexive_axm: proof val symmetric_axm: proof val transitive_axm: proof val equal_intr_axm: proof val equal_elim_axm: proof val abstract_rule_axm: proof val combination_axm: proof val reflexive_proof: proof val symmetric_proof: proof -> proof val transitive_proof: typ -> term -> proof -> proof -> proof val equal_intr_proof: term -> term -> proof -> proof -> proof val equal_elim_proof: term -> term -> proof -> proof -> proof val abstract_rule_proof: string * term -> proof -> proof val combination_proof: term -> term -> term -> term -> proof -> proof -> proof val strip_shyps_proof: Sorts.algebra -> (typ * sort) list -> (typ * sort) list -> sort list -> proof -> proof val of_sort_proof: Sorts.algebra -> (class * class -> proof) -> (string * class list list * class -> proof) -> (typ * class -> proof) -> typ * sort -> proof list val axm_proof: string -> term -> proof val oracle_proof: string -> term -> proof val shrink_proof: proof -> proof (*rewriting on proof terms*) val add_prf_rrule: proof * proof -> theory -> theory val add_prf_rproc: (typ list -> term option list -> proof -> (proof * proof) option) -> theory -> theory val set_preproc: (theory -> proof -> proof) -> theory -> theory val apply_preproc: theory -> proof -> proof val forall_intr_variables_term: term -> term val forall_intr_variables: term -> proof -> proof val no_skel: proof val normal_skel: proof val rewrite_proof: theory -> (proof * proof) list * (typ list -> term option list -> proof -> (proof * proof) option) list -> proof -> proof val rewrite_proof_notypes: (proof * proof) list * (typ list -> term option list -> proof -> (proof * proof) option) list -> proof -> proof val rew_proof: theory -> proof -> proof val reconstruct_proof: theory -> term -> proof -> proof val prop_of': term list -> proof -> term val prop_of: proof -> term val expand_name_empty: thm_header -> string option val expand_proof: theory -> (thm_header -> string option) -> proof -> proof val standard_vars: Name.context -> term * proof option -> term * proof option val standard_vars_term: Name.context -> term -> term val add_standard_vars: proof -> (string * typ) list -> (string * typ) list val add_standard_vars_term: term -> (string * typ) list -> (string * typ) list val export_enabled: unit -> bool val export_standard_enabled: unit -> bool val export_proof_boxes_required: theory -> bool val export_proof_boxes: proof_body list -> unit val fulfill_norm_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body val thm_proof: theory -> (class * class -> proof) -> (string * class list list * class -> proof) -> string * Position.T -> sort list -> term list -> term -> (serial * proof_body future) list -> proof_body -> thm * proof val unconstrain_thm_proof: theory -> (class * class -> proof) -> (string * class list list * class -> proof) -> sort list -> term -> (serial * proof_body future) list -> proof_body -> thm * proof val get_identity: sort list -> term list -> term -> proof -> {serial: serial, theory_name: string, name: string} option val get_approximative_name: sort list -> term list -> term -> proof -> string type thm_id = {serial: serial, theory_name: string} val make_thm_id: serial * string -> thm_id val thm_header_id: thm_header -> thm_id val thm_id: thm -> thm_id val get_id: sort list -> term list -> term -> proof -> thm_id option val this_id: thm_id option -> thm_id -> bool val proof_boxes: {excluded: thm_id -> bool, included: thm_id -> bool} -> proof list -> (thm_header * proof) list (*exception MIN_PROOF*) end structure Proofterm : PROOFTERM = struct (** datatype proof **) type thm_header = {serial: serial, pos: Position.T list, theory_name: string, name: string, prop: term, types: typ list option}; datatype proof = MinProof | PBound of int | Abst of string * typ option * proof | AbsP of string * term option * proof | op % of proof * term option | op %% of proof * proof | Hyp of term | PAxm of string * term * typ list option | PClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} and thm_body = Thm_Body of {open_proof: proof -> proof, body: proof_body future} and thm_node = Thm_Node of {theory_name: string, name: string, prop: term, body: proof_body future, export: unit lazy, consolidate: unit lazy}; type oracle = (string * Position.T) * term option; val oracle_ord: oracle ord = prod_ord (prod_ord fast_string_ord Position.ord) (option_ord Term_Ord.fast_term_ord); type thm = serial * thm_node; val thm_ord: thm ord = fn ((i, _), (j, _)) => int_ord (j, i); exception MIN_PROOF of unit; fun proof_of (PBody {proof, ...}) = proof; val join_proof = Future.join #> proof_of; fun map_proof_of f (PBody {oracles, thms, proof}) = PBody {oracles = oracles, thms = thms, proof = f proof}; fun thm_header serial pos theory_name name prop types : thm_header = {serial = serial, pos = pos, theory_name = theory_name, name = name, prop = prop, types = types}; fun thm_body body = Thm_Body {open_proof = I, body = Future.value body}; fun thm_body_proof_raw (Thm_Body {body, ...}) = join_proof body; fun thm_body_proof_open (Thm_Body {open_proof, body, ...}) = open_proof (join_proof body); fun rep_thm_node (Thm_Node args) = args; val thm_node_theory_name = #theory_name o rep_thm_node; val thm_node_name = #name o rep_thm_node; val thm_node_prop = #prop o rep_thm_node; val thm_node_body = #body o rep_thm_node; val thm_node_thms = thm_node_body #> Future.join #> (fn PBody {thms, ...} => thms); val thm_node_export = #export o rep_thm_node; val thm_node_consolidate = #consolidate o rep_thm_node; fun join_thms (thms: thm list) = Future.joins (map (thm_node_body o #2) thms); val consolidate_bodies = maps (fn PBody {thms, ...} => map (thm_node_consolidate o #2) thms) #> Lazy.consolidate #> map Lazy.force #> ignore; fun make_thm_node theory_name name prop body export = let val consolidate = Lazy.lazy_name "Proofterm.make_thm_node" (fn () => let val PBody {thms, ...} = Future.join body in consolidate_bodies (join_thms thms) end); in Thm_Node {theory_name = theory_name, name = name, prop = prop, body = body, export = export, consolidate = consolidate} end; val no_export = Lazy.value (); fun make_thm ({serial, theory_name, name, prop, ...}: thm_header) (Thm_Body {body, ...}) = (serial, make_thm_node theory_name name prop body no_export); (* proof atoms *) fun fold_proof_atoms all f = let fun app (Abst (_, _, prf)) = app prf | app (AbsP (_, _, prf)) = app prf | app (prf % _) = app prf | app (prf1 %% prf2) = app prf1 #> app prf2 | app (prf as PThm ({serial = i, ...}, Thm_Body {body, ...})) = (fn (x, seen) => if Inttab.defined seen i then (x, seen) else let val (x', seen') = (if all then app (join_proof body) else I) (x, Inttab.update (i, ()) seen) in (f prf x', seen') end) | app prf = (fn (x, seen) => (f prf x, seen)); in fn prfs => fn x => #1 (fold app prfs (x, Inttab.empty)) end; fun fold_body_thms f = let fun app (PBody {thms, ...}) = tap join_thms thms |> fold (fn (i, thm_node) => fn (x, seen) => if Inttab.defined seen i then (x, seen) else let val name = thm_node_name thm_node; val prop = thm_node_prop thm_node; val body = Future.join (thm_node_body thm_node); val (x', seen') = app body (x, Inttab.update (i, ()) seen); in (f {serial = i, name = name, prop = prop, body = body} x', seen') end); in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end; (* proof body *) val unions_oracles = Ord_List.unions oracle_ord; val unions_thms = Ord_List.unions thm_ord; fun no_proof_body proof = PBody {oracles = [], thms = [], proof = proof}; val no_thm_body = thm_body (no_proof_body MinProof); fun no_thm_names (Abst (x, T, prf)) = Abst (x, T, no_thm_names prf) | no_thm_names (AbsP (x, t, prf)) = AbsP (x, t, no_thm_names prf) | no_thm_names (prf % t) = no_thm_names prf % t | no_thm_names (prf1 %% prf2) = no_thm_names prf1 %% no_thm_names prf2 | no_thm_names (PThm ({serial, pos, theory_name, name = _, prop, types}, thm_body)) = PThm (thm_header serial pos theory_name "" prop types, thm_body) | no_thm_names a = a; fun no_thm_proofs (Abst (x, T, prf)) = Abst (x, T, no_thm_proofs prf) | no_thm_proofs (AbsP (x, t, prf)) = AbsP (x, t, no_thm_proofs prf) | no_thm_proofs (prf % t) = no_thm_proofs prf % t | no_thm_proofs (prf1 %% prf2) = no_thm_proofs prf1 %% no_thm_proofs prf2 | no_thm_proofs (PThm (header, _)) = PThm (header, no_thm_body) | no_thm_proofs a = a; fun no_body_proofs (Abst (x, T, prf)) = Abst (x, T, no_body_proofs prf) | no_body_proofs (AbsP (x, t, prf)) = AbsP (x, t, no_body_proofs prf) | no_body_proofs (prf % t) = no_body_proofs prf % t | no_body_proofs (prf1 %% prf2) = no_body_proofs prf1 %% no_body_proofs prf2 | no_body_proofs (PThm (header, Thm_Body {open_proof, body})) = let val body' = Future.value (no_proof_body (join_proof body)); val thm_body' = Thm_Body {open_proof = open_proof, body = body'}; in PThm (header, thm_body') end | no_body_proofs a = a; (** XML data representation **) (* encode *) local open XML.Encode Term_XML.Encode; fun proof consts prf = prf |> variant [fn MinProof => ([], []), fn PBound a => ([], int a), fn Abst (a, b, c) => ([a], pair (option typ) (proof consts) (b, c)), fn AbsP (a, b, c) => ([a], pair (option (term consts)) (proof consts) (b, c)), fn a % b => ([], pair (proof consts) (option (term consts)) (a, b)), fn a %% b => ([], pair (proof consts) (proof consts) (a, b)), fn Hyp a => ([], term consts a), fn PAxm (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PClass (a, b) => ([b], typ a), fn Oracle (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PThm ({serial, pos, theory_name, name, prop, types}, Thm_Body {open_proof, body, ...}) => ([int_atom serial, theory_name, name], pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) (map Position.properties_of pos, (prop, (types, map_proof_of open_proof (Future.join body)))))] and proof_body consts (PBody {oracles, thms, proof = prf}) = triple (list (pair (pair string (properties o Position.properties_of)) (option (term consts)))) (list (thm consts)) (proof consts) (oracles, thms, prf) and thm consts (a, thm_node) = pair int (pair string (pair string (pair (term consts) (proof_body consts)))) (a, (thm_node_theory_name thm_node, (thm_node_name thm_node, (thm_node_prop thm_node, (Future.join (thm_node_body thm_node)))))); fun standard_term consts t = t |> variant [fn Const (a, b) => ([a], list typ (Consts.typargs consts (a, b))), fn Free (a, _) => ([a], []), fn Var (a, _) => (indexname a, []), fn Bound a => ([], int a), fn Abs (a, b, c) => ([a], pair typ (standard_term consts) (b, c)), fn op $ a => ([], pair (standard_term consts) (standard_term consts) a)]; fun standard_proof consts prf = prf |> variant [fn MinProof => ([], []), fn PBound a => ([], int a), fn Abst (a, SOME b, c) => ([a], pair typ (standard_proof consts) (b, c)), fn AbsP (a, SOME b, c) => ([a], pair (standard_term consts) (standard_proof consts) (b, c)), fn a % SOME b => ([], pair (standard_proof consts) (standard_term consts) (a, b)), fn a %% b => ([], pair (standard_proof consts) (standard_proof consts) (a, b)), fn Hyp a => ([], standard_term consts a), fn PAxm (name, _, SOME Ts) => ([name], list typ Ts), fn PClass (T, c) => ([c], typ T), fn Oracle (name, prop, SOME Ts) => ([name], pair (standard_term consts) (list typ) (prop, Ts)), fn PThm ({serial, theory_name, name, types = SOME Ts, ...}, _) => ([int_atom serial, theory_name, name], list typ Ts)]; in val encode = proof; val encode_body = proof_body; val encode_standard_term = standard_term; val encode_standard_proof = standard_proof; end; (* decode *) local open XML.Decode Term_XML.Decode; fun proof consts prf = prf |> variant [fn ([], []) => MinProof, fn ([], a) => PBound (int a), fn ([a], b) => let val (c, d) = pair (option typ) (proof consts) b in Abst (a, c, d) end, fn ([a], b) => let val (c, d) = pair (option (term consts)) (proof consts) b in AbsP (a, c, d) end, fn ([], a) => op % (pair (proof consts) (option (term consts)) a), fn ([], a) => op %% (pair (proof consts) (proof consts) a), fn ([], a) => Hyp (term consts a), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in PAxm (a, c, d) end, fn ([b], a) => PClass (typ a, b), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in Oracle (a, c, d) end, fn ([a, b, c], d) => let val ((e, (f, (g, h)))) = pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) d; val header = thm_header (int_atom a) (map Position.of_properties e) b c f g; in PThm (header, thm_body h) end] and proof_body consts x = let val (a, b, c) = triple (list (pair (pair string (Position.of_properties o properties)) (option (term consts)))) (list (thm consts)) (proof consts) x; in PBody {oracles = a, thms = b, proof = c} end and thm consts x = let val (a, (b, (c, (d, e)))) = pair int (pair string (pair string (pair (term consts) (proof_body consts)))) x in (a, make_thm_node b c d (Future.value e) no_export) end; in val decode = proof; val decode_body = proof_body; end; (** proof objects with different levels of detail **) val proofs = Unsynchronized.ref 2; fun proofs_enabled () = ! proofs >= 2; fun atomic_proof prf = (case prf of Abst _ => false | AbsP _ => false | op % _ => false | op %% _ => false | MinProof => false | _ => true); fun compact_proof (prf % _) = compact_proof prf | compact_proof (prf1 %% prf2) = atomic_proof prf2 andalso compact_proof prf1 | compact_proof prf = atomic_proof prf; fun (prf %> t) = prf % SOME t; val proof_combt = Library.foldl (op %>); val proof_combt' = Library.foldl (op %); val proof_combP = Library.foldl (op %%); fun strip_combt prf = let fun stripc (prf % t, ts) = stripc (prf, t::ts) | stripc x = x in stripc (prf, []) end; fun strip_combP prf = let fun stripc (prf %% prf', prfs) = stripc (prf, prf'::prfs) | stripc x = x in stripc (prf, []) end; fun strip_thm_body (body as PBody {proof, ...}) = (case fst (strip_combt (fst (strip_combP proof))) of PThm (_, Thm_Body {body = body', ...}) => Future.join body' | _ => body); val mk_Abst = fold_rev (fn (x, _: typ) => fn prf => Abst (x, NONE, prf)); val mk_AbsP = fold_rev (fn _: term => fn prf => AbsP ("H", NONE, prf)); fun map_proof_same term typ ofclass = let val typs = Same.map typ; fun proof (Abst (s, T, prf)) = (Abst (s, Same.map_option typ T, Same.commit proof prf) handle Same.SAME => Abst (s, T, proof prf)) | proof (AbsP (s, t, prf)) = (AbsP (s, Same.map_option term t, Same.commit proof prf) handle Same.SAME => AbsP (s, t, proof prf)) | proof (prf % t) = (proof prf % Same.commit (Same.map_option term) t handle Same.SAME => prf % Same.map_option term t) | proof (prf1 %% prf2) = (proof prf1 %% Same.commit proof prf2 handle Same.SAME => prf1 %% proof prf2) | proof (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (typs Ts)) | proof (PClass T_c) = ofclass T_c | proof (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (typs Ts)) | proof (PThm ({serial, pos, theory_name, name, prop, types = SOME Ts}, thm_body)) = PThm (thm_header serial pos theory_name name prop (SOME (typs Ts)), thm_body) | proof _ = raise Same.SAME; in proof end; fun map_proof_terms_same term typ = map_proof_same term typ (fn (T, c) => PClass (typ T, c)); fun map_proof_types_same typ = map_proof_terms_same (Term_Subst.map_types_same typ) typ; fun same eq f x = let val x' = f x in if eq (x, x') then raise Same.SAME else x' end; fun map_proof_terms f g = Same.commit (map_proof_terms_same (same (op =) f) (same (op =) g)); fun map_proof_types f = Same.commit (map_proof_types_same (same (op =) f)); fun fold_proof_terms f (Abst (_, _, prf)) = fold_proof_terms f prf | fold_proof_terms f (AbsP (_, SOME t, prf)) = f t #> fold_proof_terms f prf | fold_proof_terms f (AbsP (_, NONE, prf)) = fold_proof_terms f prf | fold_proof_terms f (prf % SOME t) = fold_proof_terms f prf #> f t | fold_proof_terms f (prf % NONE) = fold_proof_terms f prf | fold_proof_terms f (prf1 %% prf2) = fold_proof_terms f prf1 #> fold_proof_terms f prf2 | fold_proof_terms _ _ = I; fun fold_proof_terms_types f g (Abst (_, SOME T, prf)) = g T #> fold_proof_terms_types f g prf | fold_proof_terms_types f g (Abst (_, NONE, prf)) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (AbsP (_, SOME t, prf)) = f t #> fold_proof_terms_types f g prf | fold_proof_terms_types f g (AbsP (_, NONE, prf)) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (prf % SOME t) = fold_proof_terms_types f g prf #> f t | fold_proof_terms_types f g (prf % NONE) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (prf1 %% prf2) = fold_proof_terms_types f g prf1 #> fold_proof_terms_types f g prf2 | fold_proof_terms_types _ g (PAxm (_, _, SOME Ts)) = fold g Ts | fold_proof_terms_types _ g (PClass (T, _)) = g T | fold_proof_terms_types _ g (Oracle (_, _, SOME Ts)) = fold g Ts | fold_proof_terms_types _ g (PThm ({types = SOME Ts, ...}, _)) = fold g Ts | fold_proof_terms_types _ _ _ = I; fun maxidx_proof prf = fold_proof_terms_types Term.maxidx_term Term.maxidx_typ prf; fun size_of_proof (Abst (_, _, prf)) = 1 + size_of_proof prf | size_of_proof (AbsP (_, _, prf)) = 1 + size_of_proof prf | size_of_proof (prf % _) = 1 + size_of_proof prf | size_of_proof (prf1 %% prf2) = size_of_proof prf1 + size_of_proof prf2 | size_of_proof _ = 1; fun change_types types (PAxm (name, prop, _)) = PAxm (name, prop, types) | change_types (SOME [T]) (PClass (_, c)) = PClass (T, c) | change_types types (Oracle (name, prop, _)) = Oracle (name, prop, types) | change_types types (PThm ({serial, pos, theory_name, name, prop, types = _}, thm_body)) = PThm (thm_header serial pos theory_name name prop types, thm_body) | change_types _ prf = prf; (* utilities *) fun strip_abs (_::Ts) (Abs (_, _, t)) = strip_abs Ts t | strip_abs _ t = t; fun mk_abs Ts t = Library.foldl (fn (t', T) => Abs ("", T, t')) (t, Ts); (*Abstraction of a proof term over its occurrences of v, which must contain no loose bound variables. The resulting proof term is ready to become the body of an Abst.*) fun prf_abstract_over v = let fun abst' lev u = if v aconv u then Bound lev else (case u of Abs (a, T, t) => Abs (a, T, abst' (lev + 1) t) | f $ t => (abst' lev f $ absth' lev t handle Same.SAME => f $ abst' lev t) | _ => raise Same.SAME) and absth' lev t = (abst' lev t handle Same.SAME => t); fun abst lev (AbsP (a, t, prf)) = (AbsP (a, Same.map_option (abst' lev) t, absth lev prf) handle Same.SAME => AbsP (a, t, abst lev prf)) | abst lev (Abst (a, T, prf)) = Abst (a, T, abst (lev + 1) prf) | abst lev (prf1 %% prf2) = (abst lev prf1 %% absth lev prf2 handle Same.SAME => prf1 %% abst lev prf2) | abst lev (prf % t) = (abst lev prf % Option.map (absth' lev) t handle Same.SAME => prf % Same.map_option (abst' lev) t) | abst _ _ = raise Same.SAME and absth lev prf = (abst lev prf handle Same.SAME => prf); in absth 0 end; (*increments a proof term's non-local bound variables required when moving a proof term within abstractions inc is increment for bound variables lev is level at which a bound variable is considered 'loose'*) fun incr_bv' inct tlev t = incr_bv (inct, tlev, t); fun prf_incr_bv' incP _ Plev _ (PBound i) = if i >= Plev then PBound (i+incP) else raise Same.SAME | prf_incr_bv' incP inct Plev tlev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (same (op =) (incr_bv' inct tlev)) t, prf_incr_bv incP inct (Plev+1) tlev body) handle Same.SAME => AbsP (a, t, prf_incr_bv' incP inct (Plev+1) tlev body)) | prf_incr_bv' incP inct Plev tlev (Abst (a, T, body)) = Abst (a, T, prf_incr_bv' incP inct Plev (tlev+1) body) | prf_incr_bv' incP inct Plev tlev (prf %% prf') = (prf_incr_bv' incP inct Plev tlev prf %% prf_incr_bv incP inct Plev tlev prf' handle Same.SAME => prf %% prf_incr_bv' incP inct Plev tlev prf') | prf_incr_bv' incP inct Plev tlev (prf % t) = (prf_incr_bv' incP inct Plev tlev prf % Option.map (incr_bv' inct tlev) t handle Same.SAME => prf % Same.map_option (same (op =) (incr_bv' inct tlev)) t) | prf_incr_bv' _ _ _ _ _ = raise Same.SAME and prf_incr_bv incP inct Plev tlev prf = (prf_incr_bv' incP inct Plev tlev prf handle Same.SAME => prf); fun incr_pboundvars 0 0 prf = prf | incr_pboundvars incP inct prf = prf_incr_bv incP inct 0 0 prf; fun prf_loose_bvar1 (prf1 %% prf2) k = prf_loose_bvar1 prf1 k orelse prf_loose_bvar1 prf2 k | prf_loose_bvar1 (prf % SOME t) k = prf_loose_bvar1 prf k orelse loose_bvar1 (t, k) | prf_loose_bvar1 (_ % NONE) _ = true | prf_loose_bvar1 (AbsP (_, SOME t, prf)) k = loose_bvar1 (t, k) orelse prf_loose_bvar1 prf k | prf_loose_bvar1 (AbsP (_, NONE, _)) _ = true | prf_loose_bvar1 (Abst (_, _, prf)) k = prf_loose_bvar1 prf (k+1) | prf_loose_bvar1 _ _ = false; fun prf_loose_Pbvar1 (PBound i) k = i = k | prf_loose_Pbvar1 (prf1 %% prf2) k = prf_loose_Pbvar1 prf1 k orelse prf_loose_Pbvar1 prf2 k | prf_loose_Pbvar1 (prf % _) k = prf_loose_Pbvar1 prf k | prf_loose_Pbvar1 (AbsP (_, _, prf)) k = prf_loose_Pbvar1 prf (k+1) | prf_loose_Pbvar1 (Abst (_, _, prf)) k = prf_loose_Pbvar1 prf k | prf_loose_Pbvar1 _ _ = false; fun prf_add_loose_bnos plev _ (PBound i) (is, js) = if i < plev then (is, js) else (insert (op =) (i-plev) is, js) | prf_add_loose_bnos plev tlev (prf1 %% prf2) p = prf_add_loose_bnos plev tlev prf2 (prf_add_loose_bnos plev tlev prf1 p) | prf_add_loose_bnos plev tlev (prf % opt) (is, js) = prf_add_loose_bnos plev tlev prf (case opt of NONE => (is, insert (op =) ~1 js) | SOME t => (is, add_loose_bnos (t, tlev, js))) | prf_add_loose_bnos plev tlev (AbsP (_, opt, prf)) (is, js) = prf_add_loose_bnos (plev+1) tlev prf (case opt of NONE => (is, insert (op =) ~1 js) | SOME t => (is, add_loose_bnos (t, tlev, js))) | prf_add_loose_bnos plev tlev (Abst (_, _, prf)) p = prf_add_loose_bnos plev (tlev+1) prf p | prf_add_loose_bnos _ _ _ _ = ([], []); (* substitutions *) local fun conflicting_tvarsT envT = Term.fold_atyps (fn T => fn instT => (case T of TVar (v as (_, S)) => if Term_Subst.TVars.defined instT v orelse can (Type.lookup envT) v then instT else Term_Subst.TVars.update (v, Logic.dummy_tfree S) instT | _ => instT)); fun conflicting_tvars env = Term.fold_aterms (fn t => fn inst => (case t of Var (v as (_, T)) => if Term_Subst.Vars.defined inst v orelse can (Envir.lookup env) v then inst else Term_Subst.Vars.update (v, Free ("dummy", T)) inst | _ => inst)); fun del_conflicting_tvars envT ty = Term_Subst.instantiateT (Term_Subst.TVars.build (conflicting_tvarsT envT ty)) ty; fun del_conflicting_vars env tm = let val instT = Term_Subst.TVars.build (tm |> Term.fold_types (conflicting_tvarsT (Envir.type_env env))); val inst = Term_Subst.Vars.build (tm |> conflicting_tvars env); in Term_Subst.instantiate (instT, inst) tm end; in fun norm_proof env = let val envT = Envir.type_env env; fun msg s = warning ("type conflict in norm_proof:\n" ^ s); fun htype f t = f env t handle TYPE (s, _, _) => (msg s; f env (del_conflicting_vars env t)); fun htypeT f T = f envT T handle TYPE (s, _, _) => (msg s; f envT (del_conflicting_tvars envT T)); fun htypeTs f Ts = f envT Ts handle TYPE (s, _, _) => (msg s; f envT (map (del_conflicting_tvars envT) Ts)); fun norm (Abst (s, T, prf)) = (Abst (s, Same.map_option (htypeT Envir.norm_type_same) T, Same.commit norm prf) handle Same.SAME => Abst (s, T, norm prf)) | norm (AbsP (s, t, prf)) = (AbsP (s, Same.map_option (htype Envir.norm_term_same) t, Same.commit norm prf) handle Same.SAME => AbsP (s, t, norm prf)) | norm (prf % t) = (norm prf % Option.map (htype Envir.norm_term) t handle Same.SAME => prf % Same.map_option (htype Envir.norm_term_same) t) | norm (prf1 %% prf2) = (norm prf1 %% Same.commit norm prf2 handle Same.SAME => prf1 %% norm prf2) | norm (PAxm (s, prop, Ts)) = PAxm (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) | norm (PClass (T, c)) = PClass (htypeT Envir.norm_type_same T, c) | norm (Oracle (s, prop, Ts)) = Oracle (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) | norm (PThm ({serial = i, pos = p, theory_name, name = a, prop = t, types = Ts}, thm_body)) = PThm (thm_header i p theory_name a t (Same.map_option (htypeTs Envir.norm_types_same) Ts), thm_body) | norm _ = raise Same.SAME; in Same.commit norm end; end; (* remove some types in proof term (to save space) *) fun remove_types (Abs (s, _, t)) = Abs (s, dummyT, remove_types t) | remove_types (t $ u) = remove_types t $ remove_types u | remove_types (Const (s, _)) = Const (s, dummyT) | remove_types t = t; fun remove_types_env (Envir.Envir {maxidx, tenv, tyenv}) = Envir.Envir {maxidx = maxidx, tenv = Vartab.map (K (apsnd remove_types)) tenv, tyenv = tyenv}; fun norm_proof' env prf = norm_proof (remove_types_env env) prf; (* substitution of bound variables *) fun prf_subst_bounds args prf = let val n = length args; fun subst' lev (Bound i) = (if i Bound (i-n)) (*loose: change it*) | subst' lev (Abs (a, T, body)) = Abs (a, T, subst' (lev+1) body) | subst' lev (f $ t) = (subst' lev f $ substh' lev t handle Same.SAME => f $ subst' lev t) | subst' _ _ = raise Same.SAME and substh' lev t = (subst' lev t handle Same.SAME => t); fun subst lev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (subst' lev) t, substh lev body) handle Same.SAME => AbsP (a, t, subst lev body)) | subst lev (Abst (a, T, body)) = Abst (a, T, subst (lev+1) body) | subst lev (prf %% prf') = (subst lev prf %% substh lev prf' handle Same.SAME => prf %% subst lev prf') | subst lev (prf % t) = (subst lev prf % Option.map (substh' lev) t handle Same.SAME => prf % Same.map_option (subst' lev) t) | subst _ _ = raise Same.SAME and substh lev prf = (subst lev prf handle Same.SAME => prf); in (case args of [] => prf | _ => substh 0 prf) end; fun prf_subst_pbounds args prf = let val n = length args; fun subst (PBound i) Plev tlev = (if i < Plev then raise Same.SAME (*var is locally bound*) else incr_pboundvars Plev tlev (nth args (i-Plev)) handle General.Subscript => PBound (i-n) (*loose: change it*)) | subst (AbsP (a, t, body)) Plev tlev = AbsP (a, t, subst body (Plev+1) tlev) | subst (Abst (a, T, body)) Plev tlev = Abst (a, T, subst body Plev (tlev+1)) | subst (prf %% prf') Plev tlev = (subst prf Plev tlev %% substh prf' Plev tlev handle Same.SAME => prf %% subst prf' Plev tlev) | subst (prf % t) Plev tlev = subst prf Plev tlev % t | subst _ _ _ = raise Same.SAME and substh prf Plev tlev = (subst prf Plev tlev handle Same.SAME => prf) in (case args of [] => prf | _ => substh prf 0 0) end; (* freezing and thawing of variables in proof terms *) local fun frzT names = map_type_tvar (fn (ixn, S) => TFree (the (AList.lookup (op =) names ixn), S)); fun thawT names = map_type_tfree (fn (a, S) => (case AList.lookup (op =) names a of NONE => TFree (a, S) | SOME ixn => TVar (ixn, S))); fun freeze names names' (t $ u) = freeze names names' t $ freeze names names' u | freeze names names' (Abs (s, T, t)) = Abs (s, frzT names' T, freeze names names' t) | freeze _ names' (Const (s, T)) = Const (s, frzT names' T) | freeze _ names' (Free (s, T)) = Free (s, frzT names' T) | freeze names names' (Var (ixn, T)) = Free (the (AList.lookup (op =) names ixn), frzT names' T) | freeze _ _ t = t; fun thaw names names' (t $ u) = thaw names names' t $ thaw names names' u | thaw names names' (Abs (s, T, t)) = Abs (s, thawT names' T, thaw names names' t) | thaw _ names' (Const (s, T)) = Const (s, thawT names' T) | thaw names names' (Free (s, T)) = let val T' = thawT names' T in (case AList.lookup (op =) names s of NONE => Free (s, T') | SOME ixn => Var (ixn, T')) end | thaw _ names' (Var (ixn, T)) = Var (ixn, thawT names' T) | thaw _ _ t = t; in fun freeze_thaw_prf prf = let val (fs, Tfs, vs, Tvs) = fold_proof_terms_types (fn t => fn (fs, Tfs, vs, Tvs) => (Term.add_free_names t fs, Term.add_tfree_names t Tfs, Term.add_var_names t vs, Term.add_tvar_names t Tvs)) (fn T => fn (fs, Tfs, vs, Tvs) => (fs, Term.add_tfree_namesT T Tfs, vs, Term.add_tvar_namesT T Tvs)) prf ([], [], [], []); val names = vs ~~ Name.variant_list fs (map fst vs); val names' = Tvs ~~ Name.variant_list Tfs (map fst Tvs); val rnames = map swap names; val rnames' = map swap names'; in (map_proof_terms (freeze names names') (frzT names') prf, map_proof_terms (thaw rnames rnames') (thawT rnames')) end; end; (** inference rules **) (* trivial implication *) val trivial_proof = AbsP ("H", NONE, PBound 0); (* implication introduction *) fun gen_implies_intr_proof f h prf = let fun abshyp i (Hyp t) = if h aconv t then PBound i else raise Same.SAME | abshyp i (Abst (s, T, prf)) = Abst (s, T, abshyp i prf) | abshyp i (AbsP (s, t, prf)) = AbsP (s, t, abshyp (i + 1) prf) | abshyp i (prf % t) = abshyp i prf % t | abshyp i (prf1 %% prf2) = (abshyp i prf1 %% abshyph i prf2 handle Same.SAME => prf1 %% abshyp i prf2) | abshyp _ _ = raise Same.SAME and abshyph i prf = (abshyp i prf handle Same.SAME => prf); in AbsP ("H", f h, abshyph 0 prf) end; val implies_intr_proof = gen_implies_intr_proof (K NONE); val implies_intr_proof' = gen_implies_intr_proof SOME; (* forall introduction *) fun forall_intr_proof (a, v) opt_T prf = Abst (a, opt_T, prf_abstract_over v prf); fun forall_intr_proof' v prf = let val (a, T) = (case v of Var ((a, _), T) => (a, T) | Free (a, T) => (a, T)) in forall_intr_proof (a, v) (SOME T) prf end; (* varify *) fun varify_proof t fixed prf = let val fs = - (t, []) |-> (Term.fold_types o Term.fold_atyps) - (fn TFree v => if Term_Subst.TFrees.defined fixed v then I else insert (op =) v | _ => I); + build (t |> (Term.fold_types o Term.fold_atyps) + (fn TFree v => if Term_Subst.TFrees.defined fixed v then I else insert (op =) v | _ => I)); val used = Name.context |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t; val fmap = fs ~~ #1 (fold_map Name.variant (map fst fs) used); fun thaw (a, S) = (case AList.lookup (op =) fmap (a, S) of NONE => TFree (a, S) | SOME b => TVar ((b, 0), S)); in map_proof_terms (map_types (map_type_tfree thaw)) (map_type_tfree thaw) prf end; local fun new_name ix (pairs, used) = let val v = singleton (Name.variant_list used) (string_of_indexname ix) in ((ix, v) :: pairs, v :: used) end; fun freeze_one alist (ix, sort) = (case AList.lookup (op =) alist ix of NONE => TVar (ix, sort) | SOME name => TFree (name, sort)); in fun legacy_freezeT t prf = let val used = Term.add_tfree_names t []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvars t [])) ([], used); in (case alist of [] => prf (*nothing to do!*) | _ => let val frzT = map_type_tvar (freeze_one alist) in map_proof_terms (map_types frzT) frzT prf end) end; end; (* rotate assumptions *) fun rotate_proof Bs Bi' params asms m prf = let val i = length asms; val j = length Bs; in mk_AbsP (Bs @ [Bi']) (proof_combP (prf, map PBound (j downto 1) @ [mk_Abst params (mk_AbsP asms (proof_combP (proof_combt (PBound i, map Bound ((length params - 1) downto 0)), map PBound (((i-m-1) downto 0) @ ((i-1) downto (i-m))))))])) end; (* permute premises *) fun permute_prems_proof prems' j k prf = let val n = length prems' in mk_AbsP prems' (proof_combP (prf, map PBound ((n-1 downto n-j) @ (k-1 downto 0) @ (n-j-1 downto k)))) end; (* generalization *) fun generalize_proof (tfrees, frees) idx prop prf = let val gen = if Symtab.is_empty frees then [] else fold_aterms (fn Free (x, T) => Symtab.defined frees x ? insert (op =) (x, T) | _ => I) (Term_Subst.generalize (tfrees, Symtab.empty) idx prop) []; in prf |> Same.commit (map_proof_terms_same (Term_Subst.generalize_same (tfrees, Symtab.empty) idx) (Term_Subst.generalizeT_same tfrees idx)) |> fold (fn (x, T) => forall_intr_proof (x, Free (x, T)) NONE) gen |> fold_rev (fn (x, T) => fn prf' => prf' %> Var (Name.clean_index (x, idx), T)) gen end; (* instantiation *) fun instantiate (instT, inst) = Same.commit (map_proof_terms_same (Term_Subst.instantiate_same (instT, Term_Subst.Vars.map (K remove_types) inst)) (Term_Subst.instantiateT_same instT)); (* lifting *) fun lift_proof Bi inc prop prf = let fun lift'' Us Ts t = strip_abs Ts (Logic.incr_indexes ([], Us, inc) (mk_abs Ts t)); fun lift' Us Ts (Abst (s, T, prf)) = (Abst (s, Same.map_option (Logic.incr_tvar_same inc) T, lifth' Us (dummyT::Ts) prf) handle Same.SAME => Abst (s, T, lift' Us (dummyT::Ts) prf)) | lift' Us Ts (AbsP (s, t, prf)) = (AbsP (s, Same.map_option (same (op =) (lift'' Us Ts)) t, lifth' Us Ts prf) handle Same.SAME => AbsP (s, t, lift' Us Ts prf)) | lift' Us Ts (prf % t) = (lift' Us Ts prf % Option.map (lift'' Us Ts) t handle Same.SAME => prf % Same.map_option (same (op =) (lift'' Us Ts)) t) | lift' Us Ts (prf1 %% prf2) = (lift' Us Ts prf1 %% lifth' Us Ts prf2 handle Same.SAME => prf1 %% lift' Us Ts prf2) | lift' _ _ (PAxm (s, prop, Ts)) = PAxm (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PClass (T, c)) = PClass (Logic.incr_tvar_same inc T, c) | lift' _ _ (Oracle (s, prop, Ts)) = Oracle (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PThm ({serial = i, pos = p, theory_name, name = s, prop, types = Ts}, thm_body)) = PThm (thm_header i p theory_name s prop ((Same.map_option o Same.map) (Logic.incr_tvar inc) Ts), thm_body) | lift' _ _ _ = raise Same.SAME and lifth' Us Ts prf = (lift' Us Ts prf handle Same.SAME => prf); val ps = map (Logic.lift_all inc Bi) (Logic.strip_imp_prems prop); val k = length ps; fun mk_app b (i, j, prf) = if b then (i-1, j, prf %% PBound i) else (i, j-1, prf %> Bound j); fun lift Us bs i j (Const ("Pure.imp", _) $ A $ B) = AbsP ("H", NONE (*A*), lift Us (true::bs) (i+1) j B) | lift Us bs i j (Const ("Pure.all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), lift (T::Us) (false::bs) i (j+1) t) | lift Us bs i j _ = proof_combP (lifth' (rev Us) [] prf, map (fn k => (#3 (fold_rev mk_app bs (i-1, j-1, PBound k)))) (i + k - 1 downto i)); in mk_AbsP ps (lift [] [] 0 0 Bi) end; fun incr_indexes i = Same.commit (map_proof_terms_same (Logic.incr_indexes_same ([], [], i)) (Logic.incr_tvar_same i)); (* proof by assumption *) fun mk_asm_prf t i m = let fun imp_prf _ i 0 = PBound i | imp_prf (Const ("Pure.imp", _) $ A $ B) i m = AbsP ("H", NONE (*A*), imp_prf B (i+1) (m-1)) | imp_prf _ i _ = PBound i; fun all_prf (Const ("Pure.all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), all_prf t) | all_prf t = imp_prf t (~i) m in all_prf t end; fun assumption_proof Bs Bi n prf = mk_AbsP Bs (proof_combP (prf, map PBound (length Bs - 1 downto 0) @ [mk_asm_prf Bi n ~1])); (* composition of object rule with proof state *) fun flatten_params_proof i j n (Const ("Pure.imp", _) $ A $ B, k) = AbsP ("H", NONE (*A*), flatten_params_proof (i+1) j n (B, k)) | flatten_params_proof i j n (Const ("Pure.all", _) $ Abs (a, T, t), k) = Abst (a, NONE (*T*), flatten_params_proof i (j+1) n (t, k)) | flatten_params_proof i j n (_, k) = proof_combP (proof_combt (PBound (k+i), map Bound (j-1 downto 0)), map PBound (remove (op =) (i-n) (i-1 downto 0))); fun bicompose_proof flatten Bs As A oldAs n m rprf sprf = let val lb = length Bs; val la = length As; in mk_AbsP (Bs @ As) (proof_combP (sprf, map PBound (lb + la - 1 downto la)) %% proof_combP (rprf, (if n>0 then [mk_asm_prf (the A) n m] else []) @ map (if flatten then flatten_params_proof 0 0 n else PBound o snd) (oldAs ~~ (la - 1 downto 0)))) end; (** type classes **) fun strip_shyps_proof algebra present witnessed extra prf = let val replacements = present @ witnessed @ map (`Logic.dummy_tfree) extra; fun get_replacement S = replacements |> get_first (fn (T', S') => if Sorts.sort_le algebra (S', S) then SOME T' else NONE); fun replace T = if exists (fn (T', _) => T' = T) present then raise Same.SAME else (case get_replacement (Type.sort_of_atyp T) of SOME T' => T' | NONE => raise Fail "strip_shyps_proof: bad type variable in proof term"); in Same.commit (map_proof_types_same (Term_Subst.map_atypsT_same replace)) prf end; fun of_sort_proof algebra classrel_proof arity_proof hyps = Sorts.of_sort_derivation algebra {class_relation = fn _ => fn _ => fn (prf, c1) => fn c2 => if c1 = c2 then prf else classrel_proof (c1, c2) %% prf, type_constructor = fn (a, _) => fn dom => fn c => let val Ss = map (map snd) dom and prfs = maps (map fst) dom in proof_combP (arity_proof (a, Ss, c), prfs) end, type_variable = fn typ => map (fn c => (hyps (typ, c), c)) (Type.sort_of_atyp typ)}; (** axioms and theorems **) val add_type_variables = (fold_types o fold_atyps) (insert (op =)); fun type_variables_of t = rev (add_type_variables t []); val add_variables = fold_aterms (fn a => (is_Var a orelse is_Free a) ? insert (op =) a); fun variables_of t = rev (add_variables t []); fun test_args _ [] = true | test_args is (Bound i :: ts) = not (member (op =) is i) andalso test_args (i :: is) ts | test_args _ _ = false; fun is_fun (Type ("fun", _)) = true | is_fun (TVar _) = true | is_fun _ = false; fun vars_of t = map Var (rev (Term.add_vars t [])); fun add_funvars Ts (vs, t) = if is_fun (fastype_of1 (Ts, t)) then union (op =) vs (map_filter (fn Var (ixn, T) => if is_fun T then SOME ixn else NONE | _ => NONE) (vars_of t)) else vs; fun add_npvars q p Ts (vs, Const ("Pure.imp", _) $ t $ u) = add_npvars q p Ts (add_npvars q (not p) Ts (vs, t), u) | add_npvars q p Ts (vs, Const ("Pure.all", Type (_, [Type (_, [T, _]), _])) $ t) = add_npvars q p Ts (vs, if p andalso q then betapply (t, Var (("",0), T)) else t) | add_npvars q p Ts (vs, Abs (_, T, t)) = add_npvars q p (T::Ts) (vs, t) | add_npvars _ _ Ts (vs, t) = add_npvars' Ts (vs, t) and add_npvars' Ts (vs, t) = (case strip_comb t of (Var (ixn, _), ts) => if test_args [] ts then vs else Library.foldl (add_npvars' Ts) (AList.update (op =) (ixn, Library.foldl (add_funvars Ts) ((these ooo AList.lookup) (op =) vs ixn, ts)) vs, ts) | (Abs (_, T, u), ts) => Library.foldl (add_npvars' (T::Ts)) (vs, u :: ts) | (_, ts) => Library.foldl (add_npvars' Ts) (vs, ts)); fun prop_vars (Const ("Pure.imp", _) $ P $ Q) = union (op =) (prop_vars P) (prop_vars Q) | prop_vars (Const ("Pure.all", _) $ Abs (_, _, t)) = prop_vars t | prop_vars t = (case strip_comb t of (Var (ixn, _), _) => [ixn] | _ => []); fun is_proj t = let fun is_p i t = (case strip_comb t of (Bound _, []) => false | (Bound j, ts) => j >= i orelse exists (is_p i) ts | (Abs (_, _, u), _) => is_p (i+1) u | (_, ts) => exists (is_p i) ts) in (case strip_abs_body t of Bound _ => true | t' => is_p 0 t') end; fun prop_args prop = let val needed_vars = union (op =) (Library.foldl (uncurry (union (op =))) ([], map (uncurry (insert (op =))) (add_npvars true true [] ([], prop)))) (prop_vars prop); in variables_of prop |> map (fn var as Var (ixn, _) => if member (op =) needed_vars ixn then SOME var else NONE | free => SOME free) end; fun const_proof mk name prop = let val args = prop_args prop; val ({outer_constraints, ...}, prop1) = Logic.unconstrainT [] prop; val head = mk (name, prop1, NONE); in proof_combP (proof_combt' (head, args), map PClass outer_constraints) end; val axm_proof = const_proof PAxm; val oracle_proof = const_proof Oracle; val shrink_proof = let fun shrink ls lev (prf as Abst (a, T, body)) = let val (b, is, ch, body') = shrink ls (lev+1) body in (b, is, ch, if ch then Abst (a, T, body') else prf) end | shrink ls lev (prf as AbsP (a, t, body)) = let val (b, is, ch, body') = shrink (lev::ls) lev body in (b orelse member (op =) is 0, map_filter (fn 0 => NONE | i => SOME (i-1)) is, ch, if ch then AbsP (a, t, body') else prf) end | shrink ls lev prf = let val (is, ch, _, prf') = shrink' ls lev [] [] prf in (false, is, ch, prf') end and shrink' ls lev ts prfs (prf as prf1 %% prf2) = let val p as (_, is', ch', prf') = shrink ls lev prf2; val (is, ch, ts', prf'') = shrink' ls lev ts (p::prfs) prf1 in (union (op =) is is', ch orelse ch', ts', if ch orelse ch' then prf'' %% prf' else prf) end | shrink' ls lev ts prfs (prf as prf1 % t) = let val (is, ch, (ch', t')::ts', prf') = shrink' ls lev (t::ts) prfs prf1 in (is, ch orelse ch', ts', if ch orelse ch' then prf' % t' else prf) end | shrink' ls lev ts prfs (prf as PBound i) = (if exists (fn SOME (Bound j) => lev-j <= nth ls i | _ => true) ts orelse has_duplicates (op =) (Library.foldl (fn (js, SOME (Bound j)) => j :: js | (js, _) => js) ([], ts)) orelse exists #1 prfs then [i] else [], false, map (pair false) ts, prf) | shrink' _ _ ts _ (Hyp t) = ([], false, map (pair false) ts, Hyp t) | shrink' _ _ ts _ (prf as MinProof) = ([], false, map (pair false) ts, prf) | shrink' _ _ ts _ (prf as PClass _) = ([], false, map (pair false) ts, prf) | shrink' _ _ ts prfs prf = let val prop = (case prf of PAxm (_, prop, _) => prop | Oracle (_, prop, _) => prop | PThm ({prop, ...}, _) => prop | _ => raise Fail "shrink: proof not in normal form"); val vs = vars_of prop; val (ts', ts'') = chop (length vs) ts; val insts = take (length ts') (map (fst o dest_Var) vs) ~~ ts'; val nvs = Library.foldl (fn (ixns', (ixn, ixns)) => insert (op =) ixn (case AList.lookup (op =) insts ixn of SOME (SOME t) => if is_proj t then union (op =) ixns ixns' else ixns' | _ => union (op =) ixns ixns')) (needed prop ts'' prfs, add_npvars false true [] ([], prop)); val insts' = map (fn (ixn, x as SOME _) => if member (op =) nvs ixn then (false, x) else (true, NONE) | (_, x) => (false, x)) insts in ([], false, insts' @ map (pair false) ts'', prf) end and needed (Const ("Pure.imp", _) $ t $ u) ts ((b, _, _, _)::prfs) = union (op =) (if b then map (fst o dest_Var) (vars_of t) else []) (needed u ts prfs) | needed (Var (ixn, _)) (_::_) _ = [ixn] | needed _ _ _ = []; in fn prf => #4 (shrink [] 0 prf) end; (** axioms for equality **) val aT = TFree ("'a", []); val bT = TFree ("'b", []); val x = Free ("x", aT); val y = Free ("y", aT); val z = Free ("z", aT); val A = Free ("A", propT); val B = Free ("B", propT); val f = Free ("f", aT --> bT); val g = Free ("g", aT --> bT); val equality_axms = [("reflexive", Logic.mk_equals (x, x)), ("symmetric", Logic.mk_implies (Logic.mk_equals (x, y), Logic.mk_equals (y, x))), ("transitive", Logic.list_implies ([Logic.mk_equals (x, y), Logic.mk_equals (y, z)], Logic.mk_equals (x, z))), ("equal_intr", Logic.list_implies ([Logic.mk_implies (A, B), Logic.mk_implies (B, A)], Logic.mk_equals (A, B))), ("equal_elim", Logic.list_implies ([Logic.mk_equals (A, B), A], B)), ("abstract_rule", Logic.mk_implies (Logic.all x (Logic.mk_equals (f $ x, g $ x)), Logic.mk_equals (lambda x (f $ x), lambda x (g $ x)))), ("combination", Logic.list_implies ([Logic.mk_equals (f, g), Logic.mk_equals (x, y)], Logic.mk_equals (f $ x, g $ y)))]; val [reflexive_axm, symmetric_axm, transitive_axm, equal_intr_axm, equal_elim_axm, abstract_rule_axm, combination_axm] = map (fn (s, t) => PAxm ("Pure." ^ s, Logic.varify_global t, NONE)) equality_axms; val reflexive_proof = reflexive_axm % NONE; val is_reflexive_proof = fn PAxm ("Pure.reflexive", _, _) % _ => true | _ => false; fun symmetric_proof prf = if is_reflexive_proof prf then prf else symmetric_axm % NONE % NONE %% prf; fun transitive_proof U u prf1 prf2 = if is_reflexive_proof prf1 then prf2 else if is_reflexive_proof prf2 then prf1 else if U = propT then transitive_axm % NONE % SOME (remove_types u) % NONE %% prf1 %% prf2 else transitive_axm % NONE % NONE % NONE %% prf1 %% prf2; fun equal_intr_proof A B prf1 prf2 = equal_intr_axm %> remove_types A %> remove_types B %% prf1 %% prf2; fun equal_elim_proof A B prf1 prf2 = equal_elim_axm %> remove_types A %> remove_types B %% prf1 %% prf2; fun abstract_rule_proof (a, x) prf = abstract_rule_axm % NONE % NONE %% forall_intr_proof (a, x) NONE prf; fun check_comb (PAxm ("Pure.combination", _, _) % f % _ % _ % _ %% prf %% _) = is_some f orelse check_comb prf | check_comb (PAxm ("Pure.transitive", _, _) % _ % _ % _ %% prf1 %% prf2) = check_comb prf1 andalso check_comb prf2 | check_comb (PAxm ("Pure.symmetric", _, _) % _ % _ %% prf) = check_comb prf | check_comb _ = false; fun combination_proof f g t u prf1 prf2 = let val f = Envir.beta_norm f; val g = Envir.beta_norm g; val prf = if check_comb prf1 then combination_axm % NONE % NONE else (case prf1 of PAxm ("Pure.reflexive", _, _) % _ => combination_axm %> remove_types f % NONE | _ => combination_axm %> remove_types f %> remove_types g) in prf % (case head_of f of Abs _ => SOME (remove_types t) | Var _ => SOME (remove_types t) | _ => NONE) % (case head_of g of Abs _ => SOME (remove_types u) | Var _ => SOME (remove_types u) | _ => NONE) %% prf1 %% prf2 end; (** rewriting on proof terms **) (* simple first order matching functions for terms and proofs (see pattern.ML) *) exception PMatch; fun flt (i: int) = filter (fn n => n < i); fun fomatch Ts tymatch j instsp p = let fun mtch (instsp as (tyinsts, insts)) = fn (Var (ixn, T), t) => if j>0 andalso not (null (flt j (loose_bnos t))) then raise PMatch else (tymatch (tyinsts, fn () => (T, fastype_of1 (Ts, t))), (ixn, t) :: insts) | (Free (a, T), Free (b, U)) => if a=b then (tymatch (tyinsts, K (T, U)), insts) else raise PMatch | (Const (a, T), Const (b, U)) => if a=b then (tymatch (tyinsts, K (T, U)), insts) else raise PMatch | (f $ t, g $ u) => mtch (mtch instsp (f, g)) (t, u) | (Bound i, Bound j) => if i=j then instsp else raise PMatch | _ => raise PMatch in mtch instsp (apply2 Envir.beta_eta_contract p) end; fun match_proof Ts tymatch = let fun optmatch _ inst (NONE, _) = inst | optmatch _ _ (SOME _, NONE) = raise PMatch | optmatch mtch inst (SOME x, SOME y) = mtch inst (x, y) fun matcht Ts j (pinst, tinst) (t, u) = (pinst, fomatch Ts tymatch j tinst (t, Envir.beta_norm u)); fun matchT (pinst, (tyinsts, insts)) p = (pinst, (tymatch (tyinsts, K p), insts)); fun matchTs inst (Ts, Us) = Library.foldl (uncurry matchT) (inst, Ts ~~ Us); fun mtch Ts i j (pinst, tinst) (Hyp (Var (ixn, _)), prf) = if i = 0 andalso j = 0 then ((ixn, prf) :: pinst, tinst) else (case apfst (flt i) (apsnd (flt j) (prf_add_loose_bnos 0 0 prf ([], []))) of ([], []) => ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst) | ([], _) => if j = 0 then ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst) else raise PMatch | _ => raise PMatch) | mtch Ts i j inst (prf1 % opt1, prf2 % opt2) = optmatch (matcht Ts j) (mtch Ts i j inst (prf1, prf2)) (opt1, opt2) | mtch Ts i j inst (prf1 %% prf2, prf1' %% prf2') = mtch Ts i j (mtch Ts i j inst (prf1, prf1')) (prf2, prf2') | mtch Ts i j inst (Abst (_, opT, prf1), Abst (_, opU, prf2)) = mtch (the_default dummyT opU :: Ts) i (j+1) (optmatch matchT inst (opT, opU)) (prf1, prf2) | mtch Ts i j inst (prf1, Abst (_, opU, prf2)) = mtch (the_default dummyT opU :: Ts) i (j+1) inst (incr_pboundvars 0 1 prf1 %> Bound 0, prf2) | mtch Ts i j inst (AbsP (_, opt, prf1), AbsP (_, opu, prf2)) = mtch Ts (i+1) j (optmatch (matcht Ts j) inst (opt, opu)) (prf1, prf2) | mtch Ts i j inst (prf1, AbsP (_, _, prf2)) = mtch Ts (i+1) j inst (incr_pboundvars 1 0 prf1 %% PBound 0, prf2) | mtch Ts i j inst (PAxm (s1, _, opTs), PAxm (s2, _, opUs)) = if s1 = s2 then optmatch matchTs inst (opTs, opUs) else raise PMatch | mtch Ts i j inst (PClass (T1, c1), PClass (T2, c2)) = if c1 = c2 then matchT inst (T1, T2) else raise PMatch | mtch Ts i j inst (PThm ({name = name1, prop = prop1, types = types1, ...}, _), PThm ({name = name2, prop = prop2, types = types2, ...}, _)) = if name1 = name2 andalso prop1 = prop2 then optmatch matchTs inst (types1, types2) else raise PMatch | mtch _ _ _ inst (PBound i, PBound j) = if i = j then inst else raise PMatch | mtch _ _ _ _ _ = raise PMatch in mtch Ts 0 0 end; fun prf_subst (pinst, (tyinsts, insts)) = let val substT = Envir.subst_type_same tyinsts; val substTs = Same.map substT; fun subst' lev (Var (xi, _)) = (case AList.lookup (op =) insts xi of NONE => raise Same.SAME | SOME u => incr_boundvars lev u) | subst' _ (Const (s, T)) = Const (s, substT T) | subst' _ (Free (s, T)) = Free (s, substT T) | subst' lev (Abs (a, T, body)) = (Abs (a, substT T, Same.commit (subst' (lev + 1)) body) handle Same.SAME => Abs (a, T, subst' (lev + 1) body)) | subst' lev (f $ t) = (subst' lev f $ Same.commit (subst' lev) t handle Same.SAME => f $ subst' lev t) | subst' _ _ = raise Same.SAME; fun subst plev tlev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (subst' tlev) t, Same.commit (subst (plev + 1) tlev) body) handle Same.SAME => AbsP (a, t, subst (plev + 1) tlev body)) | subst plev tlev (Abst (a, T, body)) = (Abst (a, Same.map_option substT T, Same.commit (subst plev (tlev + 1)) body) handle Same.SAME => Abst (a, T, subst plev (tlev + 1) body)) | subst plev tlev (prf %% prf') = (subst plev tlev prf %% Same.commit (subst plev tlev) prf' handle Same.SAME => prf %% subst plev tlev prf') | subst plev tlev (prf % t) = (subst plev tlev prf % Same.commit (Same.map_option (subst' tlev)) t handle Same.SAME => prf % Same.map_option (subst' tlev) t) | subst plev tlev (Hyp (Var (xi, _))) = (case AList.lookup (op =) pinst xi of NONE => raise Same.SAME | SOME prf' => incr_pboundvars plev tlev prf') | subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Same.map_option substTs Ts) | subst _ _ (PClass (T, c)) = PClass (substT T, c) | subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Same.map_option substTs Ts) | subst _ _ (PThm ({serial = i, pos = p, theory_name, name = id, prop, types}, thm_body)) = PThm (thm_header i p theory_name id prop (Same.map_option substTs types), thm_body) | subst _ _ _ = raise Same.SAME; in fn t => subst 0 0 t handle Same.SAME => t end; (*A fast unification filter: true unless the two terms cannot be unified. Terms must be NORMAL. Treats all Vars as distinct. *) fun could_unify prf1 prf2 = let fun matchrands (prf1 %% prf2) (prf1' %% prf2') = could_unify prf2 prf2' andalso matchrands prf1 prf1' | matchrands (prf % SOME t) (prf' % SOME t') = Term.could_unify (t, t') andalso matchrands prf prf' | matchrands (prf % _) (prf' % _) = matchrands prf prf' | matchrands _ _ = true fun head_of (prf %% _) = head_of prf | head_of (prf % _) = head_of prf | head_of prf = prf in case (head_of prf1, head_of prf2) of (_, Hyp (Var _)) => true | (Hyp (Var _), _) => true | (PAxm (a, _, _), PAxm (b, _, _)) => a = b andalso matchrands prf1 prf2 | (PClass (_, c), PClass (_, d)) => c = d andalso matchrands prf1 prf2 | (PThm ({name = a, prop = propa, ...}, _), PThm ({name = b, prop = propb, ...}, _)) => a = b andalso propa = propb andalso matchrands prf1 prf2 | (PBound i, PBound j) => i = j andalso matchrands prf1 prf2 | (AbsP _, _) => true (*because of possible eta equality*) | (Abst _, _) => true | (_, AbsP _) => true | (_, Abst _) => true | _ => false end; (* rewrite proof *) val no_skel = PBound 0; val normal_skel = Hyp (Var ((Name.uu, 0), propT)); fun rewrite_prf tymatch (rules, procs) prf = let fun rew _ _ (Abst (_, _, body) % SOME t) = SOME (prf_subst_bounds [t] body, no_skel) | rew _ _ (AbsP (_, _, body) %% prf) = SOME (prf_subst_pbounds [prf] body, no_skel) | rew Ts hs prf = (case get_first (fn r => r Ts hs prf) procs of NONE => get_first (fn (prf1, prf2) => SOME (prf_subst (match_proof Ts tymatch ([], (Vartab.empty, [])) (prf1, prf)) prf2, prf2) handle PMatch => NONE) (filter (could_unify prf o fst) rules) | some => some); fun rew0 Ts hs (prf as AbsP (_, _, prf' %% PBound 0)) = if prf_loose_Pbvar1 prf' 0 then rew Ts hs prf else let val prf'' = incr_pboundvars (~1) 0 prf' in SOME (the_default (prf'', no_skel) (rew Ts hs prf'')) end | rew0 Ts hs (prf as Abst (_, _, prf' % SOME (Bound 0))) = if prf_loose_bvar1 prf' 0 then rew Ts hs prf else let val prf'' = incr_pboundvars 0 (~1) prf' in SOME (the_default (prf'', no_skel) (rew Ts hs prf'')) end | rew0 Ts hs prf = rew Ts hs prf; fun rew1 _ _ (Hyp (Var _)) _ = NONE | rew1 Ts hs skel prf = (case rew2 Ts hs skel prf of SOME prf1 => (case rew0 Ts hs prf1 of SOME (prf2, skel') => SOME (the_default prf2 (rew1 Ts hs skel' prf2)) | NONE => SOME prf1) | NONE => (case rew0 Ts hs prf of SOME (prf1, skel') => SOME (the_default prf1 (rew1 Ts hs skel' prf1)) | NONE => NONE)) and rew2 Ts hs skel (prf % SOME t) = (case prf of Abst (_, _, body) => let val prf' = prf_subst_bounds [t] body in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end | _ => (case rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf of SOME prf' => SOME (prf' % SOME t) | NONE => NONE)) | rew2 Ts hs skel (prf % NONE) = Option.map (fn prf' => prf' % NONE) (rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf) | rew2 Ts hs skel (prf1 %% prf2) = (case prf1 of AbsP (_, _, body) => let val prf' = prf_subst_pbounds [prf2] body in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end | _ => let val (skel1, skel2) = (case skel of skel1 %% skel2 => (skel1, skel2) | _ => (no_skel, no_skel)) in (case rew1 Ts hs skel1 prf1 of SOME prf1' => (case rew1 Ts hs skel2 prf2 of SOME prf2' => SOME (prf1' %% prf2') | NONE => SOME (prf1' %% prf2)) | NONE => (case rew1 Ts hs skel2 prf2 of SOME prf2' => SOME (prf1 %% prf2') | NONE => NONE)) end) | rew2 Ts hs skel (Abst (s, T, prf)) = (case rew1 (the_default dummyT T :: Ts) hs (case skel of Abst (_, _, skel') => skel' | _ => no_skel) prf of SOME prf' => SOME (Abst (s, T, prf')) | NONE => NONE) | rew2 Ts hs skel (AbsP (s, t, prf)) = (case rew1 Ts (t :: hs) (case skel of AbsP (_, _, skel') => skel' | _ => no_skel) prf of SOME prf' => SOME (AbsP (s, t, prf')) | NONE => NONE) | rew2 _ _ _ _ = NONE; in the_default prf (rew1 [] [] no_skel prf) end; fun rewrite_proof thy = rewrite_prf (fn (tyenv, f) => Sign.typ_match thy (f ()) tyenv handle Type.TYPE_MATCH => raise PMatch); fun rewrite_proof_notypes rews = rewrite_prf fst rews; (* theory data *) structure Data = Theory_Data ( type T = ((stamp * (proof * proof)) list * (stamp * (typ list -> term option list -> proof -> (proof * proof) option)) list) * (theory -> proof -> proof) option; val empty = (([], []), NONE); val extend = I; fun merge (((rules1, procs1), preproc1), ((rules2, procs2), preproc2)) : T = ((AList.merge (op =) (K true) (rules1, rules2), AList.merge (op =) (K true) (procs1, procs2)), merge_options (preproc1, preproc2)); ); fun get_rew_data thy = let val (rules, procs) = #1 (Data.get thy) in (map #2 rules, map #2 procs) end; fun rew_proof thy = rewrite_prf fst (get_rew_data thy); fun add_prf_rrule r = (Data.map o apfst o apfst) (cons (stamp (), r)); fun add_prf_rproc p = (Data.map o apfst o apsnd) (cons (stamp (), p)); fun set_preproc f = (Data.map o apsnd) (K (SOME f)); fun apply_preproc thy = (case #2 (Data.get thy) of NONE => I | SOME f => f thy); (** reconstruction of partial proof terms **) fun forall_intr_variables_term prop = fold_rev Logic.all (variables_of prop) prop; fun forall_intr_variables prop prf = fold_rev forall_intr_proof' (variables_of prop) prf; local fun app_types shift prop Ts prf = let val inst = type_variables_of prop ~~ Ts; fun subst_same A = (case AList.lookup (op =) inst A of SOME T => T | NONE => raise Same.SAME); val subst_type_same = Term_Subst.map_atypsT_same (fn TVar ((a, i), S) => subst_same (TVar ((a, i - shift), S)) | A => subst_same A); in Same.commit (map_proof_types_same subst_type_same) prf end; fun guess_name (PThm ({name, ...}, _)) = name | guess_name (prf %% Hyp _) = guess_name prf | guess_name (prf %% PClass _) = guess_name prf | guess_name (prf % NONE) = guess_name prf | guess_name (prf % SOME (Var _)) = guess_name prf | guess_name _ = ""; (* generate constraints for proof term *) fun mk_var env Ts T = let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T) in (list_comb (v, map Bound (length Ts - 1 downto 0)), env') end; fun mk_tvar S (Envir.Envir {maxidx, tenv, tyenv}) = (TVar (("'t", maxidx + 1), S), Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}); val mk_abs = fold (fn T => fn u => Abs ("", T, u)); fun unifyT thy env T U = let val Envir.Envir {maxidx, tenv, tyenv} = env; val (tyenv', maxidx') = Sign.typ_unify thy (T, U) (tyenv, maxidx); in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end; fun chaseT env (T as TVar v) = (case Type.lookup (Envir.type_env env) v of NONE => T | SOME T' => chaseT env T') | chaseT _ T = T; fun infer_type thy (env as Envir.Envir {maxidx, tenv, tyenv}) _ vTs (t as Const (s, T)) = if T = dummyT then (case Sign.const_type thy s of NONE => error ("reconstruct_proof: No such constant: " ^ quote s) | SOME T => let val T' = Type.strip_sorts (Logic.incr_tvar (maxidx + 1) T) in (Const (s, T'), T', vTs, Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}) end) else (t, T, vTs, env) | infer_type _ env _ vTs (t as Free (s, T)) = if T = dummyT then (case Symtab.lookup vTs s of NONE => let val (T, env') = mk_tvar [] env in (Free (s, T), T, Symtab.update_new (s, T) vTs, env') end | SOME T => (Free (s, T), T, vTs, env)) else (t, T, vTs, env) | infer_type _ _ _ _ (Var _) = error "reconstruct_proof: internal error" | infer_type thy env Ts vTs (Abs (s, T, t)) = let val (T', env') = if T = dummyT then mk_tvar [] env else (T, env); val (t', U, vTs', env'') = infer_type thy env' (T' :: Ts) vTs t in (Abs (s, T', t'), T' --> U, vTs', env'') end | infer_type thy env Ts vTs (t $ u) = let val (t', T, vTs1, env1) = infer_type thy env Ts vTs t; val (u', U, vTs2, env2) = infer_type thy env1 Ts vTs1 u; in (case chaseT env2 T of Type ("fun", [U', V]) => (t' $ u', V, vTs2, unifyT thy env2 U U') | _ => let val (V, env3) = mk_tvar [] env2 in (t' $ u', V, vTs2, unifyT thy env3 T (U --> V)) end) end | infer_type _ env Ts vTs (t as Bound i) = ((t, nth Ts i, vTs, env) handle General.Subscript => error ("infer_type: bad variable index " ^ string_of_int i)); fun cantunify thy (t, u) = error ("Non-unifiable terms:\n" ^ Syntax.string_of_term_global thy t ^ "\n\n" ^ Syntax.string_of_term_global thy u); fun decompose thy Ts (p as (t, u)) env = let fun rigrig (a, T) (b, U) uT ts us = if a <> b then cantunify thy p else apfst flat (fold_map (decompose thy Ts) (ts ~~ us) (uT env T U)) in case apply2 (strip_comb o Envir.head_norm env) p of ((Const c, ts), (Const d, us)) => rigrig c d (unifyT thy) ts us | ((Free c, ts), (Free d, us)) => rigrig c d (unifyT thy) ts us | ((Bound i, ts), (Bound j, us)) => rigrig (i, dummyT) (j, dummyT) (K o K) ts us | ((Abs (_, T, t), []), (Abs (_, U, u), [])) => decompose thy (T::Ts) (t, u) (unifyT thy env T U) | ((Abs (_, T, t), []), _) => decompose thy (T::Ts) (t, incr_boundvars 1 u $ Bound 0) env | (_, (Abs (_, T, u), [])) => decompose thy (T::Ts) (incr_boundvars 1 t $ Bound 0, u) env | _ => ([(mk_abs Ts t, mk_abs Ts u)], env) end; fun make_constraints_cprf thy env cprf = let fun add_cnstrt Ts prop prf cs env vTs (t, u) = let val t' = mk_abs Ts t; val u' = mk_abs Ts u in (prop, prf, cs, Pattern.unify (Context.Theory thy) (t', u') env, vTs) handle Pattern.Pattern => let val (cs', env') = decompose thy [] (t', u') env in (prop, prf, cs @ cs', env', vTs) end | Pattern.Unif => cantunify thy (Envir.norm_term env t', Envir.norm_term env u') end; fun mk_cnstrts_atom env vTs prop opTs prf = let val prop_types = type_variables_of prop; val (Ts, env') = (case opTs of NONE => fold_map (mk_tvar o Type.sort_of_atyp) prop_types env | SOME Ts => (Ts, env)); val prop' = subst_atomic_types (prop_types ~~ Ts) (forall_intr_variables_term prop) handle ListPair.UnequalLengths => error ("Wrong number of type arguments for " ^ quote (guess_name prf)) in (prop', change_types (SOME Ts) prf, [], env', vTs) end; fun head_norm (prop, prf, cnstrts, env, vTs) = (Envir.head_norm env prop, prf, cnstrts, env, vTs); fun mk_cnstrts env _ Hs vTs (PBound i) = ((nth Hs i, PBound i, [], env, vTs) handle General.Subscript => error ("mk_cnstrts: bad variable index " ^ string_of_int i)) | mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) = let val (T, env') = (case opT of NONE => mk_tvar [] env | SOME T => (T, env)); val (t, prf, cnstrts, env'', vTs') = mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf; in (Const ("Pure.all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf), cnstrts, env'', vTs') end | mk_cnstrts env Ts Hs vTs (AbsP (s, SOME t, cprf)) = let val (t', _, vTs', env') = infer_type thy env Ts vTs t; val (u, prf, cnstrts, env'', vTs'') = mk_cnstrts env' Ts (t'::Hs) vTs' cprf; in (Logic.mk_implies (t', u), AbsP (s, SOME t', prf), cnstrts, env'', vTs'') end | mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) = let val (t, env') = mk_var env Ts propT; val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf; in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs') end | mk_cnstrts env Ts Hs vTs (cprf1 %% cprf2) = let val (u, prf2, cnstrts, env', vTs') = mk_cnstrts env Ts Hs vTs cprf2 in (case head_norm (mk_cnstrts env' Ts Hs vTs' cprf1) of (Const ("Pure.imp", _) $ u' $ t', prf1, cnstrts', env'', vTs'') => add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts) env'' vTs'' (u, u') | (t, prf1, cnstrts', env'', vTs'') => let val (v, env''') = mk_var env'' Ts propT in add_cnstrt Ts v (prf1 %% prf2) (cnstrts' @ cnstrts) env''' vTs'' (t, Logic.mk_implies (u, v)) end) end | mk_cnstrts env Ts Hs vTs (cprf % SOME t) = let val (t', U, vTs1, env1) = infer_type thy env Ts vTs t in (case head_norm (mk_cnstrts env1 Ts Hs vTs1 cprf) of (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f, prf, cnstrts, env2, vTs2) => let val env3 = unifyT thy env2 T U in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2) end | (u, prf, cnstrts, env2, vTs2) => let val (v, env3) = mk_var env2 Ts (U --> propT); in add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2 (u, Const ("Pure.all", (U --> propT) --> propT) $ v) end) end | mk_cnstrts env Ts Hs vTs (cprf % NONE) = (case head_norm (mk_cnstrts env Ts Hs vTs cprf) of (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f, prf, cnstrts, env', vTs') => let val (t, env'') = mk_var env' Ts T in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs') end | (u, prf, cnstrts, env', vTs') => let val (T, env1) = mk_tvar [] env'; val (v, env2) = mk_var env1 Ts (T --> propT); val (t, env3) = mk_var env2 Ts T in add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs' (u, Const ("Pure.all", (T --> propT) --> propT) $ v) end) | mk_cnstrts env _ _ vTs (prf as PThm ({prop, types = opTs, ...}, _)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PClass (T, c)) = mk_cnstrts_atom env vTs (Logic.mk_of_class (T, c)) NONE prf | mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs) | mk_cnstrts _ _ _ _ MinProof = raise MIN_PROOF () in mk_cnstrts env [] [] Symtab.empty cprf end; (* update list of free variables of constraints *) fun upd_constrs env cs = let val tenv = Envir.term_env env; val tyenv = Envir.type_env env; val dom = [] |> Vartab.fold (cons o #1) tenv |> Vartab.fold (cons o #1) tyenv; val vran = [] |> Vartab.fold (Term.add_var_names o #2 o #2) tenv |> Vartab.fold (Term.add_tvar_namesT o #2 o #2) tyenv; fun check_cs [] = [] | check_cs ((u, p, vs) :: ps) = let val vs' = subtract (op =) dom vs in if vs = vs' then (u, p, vs) :: check_cs ps else (true, p, fold (insert op =) vs' vran) :: check_cs ps end; in check_cs cs end; (* solution of constraints *) fun solve _ [] bigenv = bigenv | solve thy cs bigenv = let fun search _ [] = error ("Unsolvable constraints:\n" ^ Pretty.string_of (Pretty.chunks (map (fn (_, p, _) => Syntax.pretty_flexpair (Syntax.init_pretty_global thy) (apply2 (Envir.norm_term bigenv) p)) cs))) | search env ((u, p as (t1, t2), vs)::ps) = if u then let val tn1 = Envir.norm_term bigenv t1; val tn2 = Envir.norm_term bigenv t2 in if Pattern.pattern tn1 andalso Pattern.pattern tn2 then (Pattern.unify (Context.Theory thy) (tn1, tn2) env, ps) handle Pattern.Unif => cantunify thy (tn1, tn2) else let val (cs', env') = decompose thy [] (tn1, tn2) env in if cs' = [(tn1, tn2)] then apsnd (cons (false, (tn1, tn2), vs)) (search env ps) else search env' (map (fn q => (true, q, vs)) cs' @ ps) end end else apsnd (cons (false, p, vs)) (search env ps); val Envir.Envir {maxidx, ...} = bigenv; val (env, cs') = search (Envir.empty maxidx) cs; in solve thy (upd_constrs env cs') (Envir.merge (bigenv, env)) end; in (* reconstruction of proofs *) fun reconstruct_proof thy prop cprf = let val (cprf' % SOME prop', thawf) = freeze_thaw_prf (cprf % SOME prop); val (t, prf, cs, env, _) = make_constraints_cprf thy (Envir.empty (maxidx_proof cprf ~1)) cprf'; val cs' = map (apply2 (Envir.norm_term env)) ((t, prop') :: cs) |> map (fn p => (true, p, Term.add_var_names (#1 p) (Term.add_var_names (#2 p) []))); val env' = solve thy cs' env; in thawf (norm_proof env' prf) end handle MIN_PROOF () => MinProof; fun prop_of_atom prop Ts = subst_atomic_types (type_variables_of prop ~~ Ts) (forall_intr_variables_term prop); val head_norm = Envir.head_norm Envir.init; fun prop_of0 Hs (PBound i) = nth Hs i | prop_of0 Hs (Abst (s, SOME T, prf)) = Logic.all_const T $ (Abs (s, T, prop_of0 Hs prf)) | prop_of0 Hs (AbsP (_, SOME t, prf)) = Logic.mk_implies (t, prop_of0 (t :: Hs) prf) | prop_of0 Hs (prf % SOME t) = (case head_norm (prop_of0 Hs prf) of Const ("Pure.all", _) $ f => f $ t | _ => error "prop_of: all expected") | prop_of0 Hs (prf1 %% _) = (case head_norm (prop_of0 Hs prf1) of Const ("Pure.imp", _) $ _ $ Q => Q | _ => error "prop_of: ==> expected") | prop_of0 _ (Hyp t) = t | prop_of0 _ (PThm ({prop, types = SOME Ts, ...}, _)) = prop_of_atom prop Ts | prop_of0 _ (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ (PClass (T, c)) = Logic.mk_of_class (T, c) | prop_of0 _ (Oracle (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ _ = error "prop_of: partial proof object"; val prop_of' = Envir.beta_eta_contract oo prop_of0; val prop_of = prop_of' []; (* expand and reconstruct subproofs *) fun expand_name_empty (header: thm_header) = if #name header = "" then SOME "" else NONE; fun expand_proof thy expand_name prf = let fun expand seen maxidx (AbsP (s, t, prf)) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', AbsP (s, t, prf')) end | expand seen maxidx (Abst (s, T, prf)) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', Abst (s, T, prf')) end | expand seen maxidx (prf1 %% prf2) = let val (seen', maxidx', prf1') = expand seen maxidx prf1; val (seen'', maxidx'', prf2') = expand seen' maxidx' prf2; in (seen'', maxidx'', prf1' %% prf2') end | expand seen maxidx (prf % t) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', prf' % t) end | expand seen maxidx (prf as PThm (header, thm_body)) = let val {serial, pos, theory_name, name, prop, types} = header in (case expand_name header of SOME name' => if name' = "" andalso is_some types then let val (seen', maxidx', prf') = (case Inttab.lookup seen serial of NONE => let val prf1 = thm_body_proof_open thm_body |> reconstruct_proof thy prop |> forall_intr_variables prop; val (seen1, maxidx1, prf2) = expand_init seen prf1 val seen2 = seen1 |> Inttab.update (serial, (maxidx1, prf2)); in (seen2, maxidx1, prf2) end | SOME (maxidx1, prf1) => (seen, maxidx1, prf1)); val prf'' = prf' |> incr_indexes (maxidx + 1) |> app_types (maxidx + 1) prop (the types); in (seen', maxidx' + maxidx + 1, prf'') end else if name' <> name then (seen, maxidx, PThm (thm_header serial pos theory_name name' prop types, thm_body)) else (seen, maxidx, prf) | NONE => (seen, maxidx, prf)) end | expand seen maxidx prf = (seen, maxidx, prf) and expand_init seen prf = expand seen (maxidx_proof prf ~1) prf; in #3 (expand_init Inttab.empty prf) end; end; (** promises **) fun fulfill_norm_proof thy ps body0 = let val _ = consolidate_bodies (map #2 ps @ [body0]); val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0; val oracles = unions_oracles (fold (fn (_, PBody {oracles, ...}) => not (null oracles) ? cons oracles) ps [oracles0]); val thms = unions_thms (fold (fn (_, PBody {thms, ...}) => not (null thms) ? cons thms) ps [thms0]); val proof = rew_proof thy proof0; in PBody {oracles = oracles, thms = thms, proof = proof} end; fun fulfill_proof_future thy promises (postproc: proof_body -> proof_body) body = let fun fulfill () = postproc (fulfill_norm_proof thy (map (apsnd Future.join) promises) (Future.join body)); in if null promises then Future.map postproc body else if Future.is_finished body andalso length promises = 1 then Future.map (fn _ => fulfill ()) (snd (hd promises)) else (singleton o Future.forks) {name = "Proofterm.fulfill_proof_future", group = NONE, deps = Future.task_of body :: map (Future.task_of o snd) promises, pri = 1, interrupts = true} fulfill end; (** theorems **) (* standardization of variables for export: only frees and named bounds *) local val declare_names_term = Term.declare_term_frees; val declare_names_term' = fn SOME t => declare_names_term t | NONE => I; val declare_names_proof = fold_proof_terms declare_names_term; fun variant names bs x = #1 (Name.variant x (fold Name.declare bs names)); fun variant_term bs (Abs (x, T, t)) = let val x' = variant (declare_names_term t Name.context) bs x; val t' = variant_term (x' :: bs) t; in Abs (x', T, t') end | variant_term bs (t $ u) = variant_term bs t $ variant_term bs u | variant_term _ t = t; fun variant_proof bs (Abst (x, T, prf)) = let val x' = variant (declare_names_proof prf Name.context) bs x; val prf' = variant_proof (x' :: bs) prf; in Abst (x', T, prf') end | variant_proof bs (AbsP (x, t, prf)) = let val x' = variant (declare_names_term' t (declare_names_proof prf Name.context)) bs x; val t' = Option.map (variant_term bs) t; val prf' = variant_proof (x' :: bs) prf; in AbsP (x', t', prf') end | variant_proof bs (prf % t) = variant_proof bs prf % Option.map (variant_term bs) t | variant_proof bs (prf1 %% prf2) = variant_proof bs prf1 %% variant_proof bs prf2 | variant_proof bs (Hyp t) = Hyp (variant_term bs t) | variant_proof _ prf = prf; val used_frees_type = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I); fun used_frees_term t = fold_types used_frees_type t #> Term.declare_term_frees t; val used_frees_proof = fold_proof_terms_types used_frees_term used_frees_type; val unvarifyT = Term.map_atyps (fn TVar ((a, _), S) => TFree (a, S) | T => T); val unvarify = Term.map_aterms (fn Var ((x, _), T) => Free (x, T) | t => t) #> map_types unvarifyT; val unvarify_proof = map_proof_terms unvarify unvarifyT; fun hidden_types prop proof = let val visible = (fold_types o fold_atyps) (insert (op =)) prop []; val add_hiddenT = fold_atyps (fn T => not (member (op =) visible T) ? insert (op =) T); in rev (fold_proof_terms_types (fold_types add_hiddenT) add_hiddenT proof []) end; fun standard_hidden_types term proof = let val hidden = hidden_types term proof; val idx = Term.maxidx_term term (maxidx_proof proof ~1) + 1; fun smash T = if member (op =) hidden T then (case Type.sort_of_atyp T of [] => dummyT | S => TVar (("'", idx), S)) else T; val smashT = map_atyps smash; in map_proof_terms (map_types smashT) smashT proof end; fun standard_hidden_terms term proof = let fun add_unless excluded x = ((is_Free x orelse is_Var x) andalso not (member (op =) excluded x)) ? insert (op =) x; val visible = fold_aterms (add_unless []) term []; val hidden = fold_proof_terms (fold_aterms (add_unless visible)) proof []; val dummy_term = Term.map_aterms (fn x => if member (op =) hidden x then Term.dummy_pattern (Term.fastype_of x) else x); in proof |> not (null hidden) ? map_proof_terms dummy_term I end; in fun standard_vars used (term, opt_proof) = let val proofs = opt_proof |> Option.map (standard_hidden_types term #> standard_hidden_terms term) |> the_list; val proof_terms = rev (fold (fold_proof_terms_types cons (cons o Logic.mk_type)) proofs []); val used_frees = used |> used_frees_term term |> fold used_frees_proof proofs; val inst = Term_Subst.zero_var_indexes_inst used_frees (term :: proof_terms); val term' = term |> Term_Subst.instantiate inst |> unvarify |> variant_term []; val proofs' = proofs |> map (instantiate inst #> unvarify_proof #> variant_proof []); in (term', try hd proofs') end; fun standard_vars_term used t = #1 (standard_vars used (t, NONE)); val add_standard_vars_term = fold_aterms (fn Free (x, T) => (fn env => (case AList.lookup (op =) env x of NONE => (x, T) :: env | SOME T' => if T = T' then env else raise TYPE ("standard_vars_env: type conflict for variable " ^ quote x, [T, T'], []))) | _ => I); val add_standard_vars = fold_proof_terms add_standard_vars_term; end; (* PThm nodes *) fun prune_body body = if Options.default_bool "prune_proofs" then (Future.map o map_proof_of) (K MinProof) body else body; fun export_enabled () = Options.default_bool "export_proofs"; fun export_standard_enabled () = Options.default_bool "export_standard_proofs"; fun export_proof_boxes_required thy = Context.theory_name thy = Context.PureN orelse (export_enabled () andalso not (export_standard_enabled ())); fun export_proof_boxes bodies = let fun export_thm (i, thm_node) boxes = if Inttab.defined boxes i then boxes else boxes |> Inttab.update (i, thm_node_export thm_node) |> fold export_thm (thm_node_thms thm_node); fun export_body (PBody {thms, ...}) = fold export_thm thms; val exports = Inttab.build (fold export_body bodies) |> Inttab.dest; in List.app (Lazy.force o #2) exports end; local fun unconstrainT_proof algebra classrel_proof arity_proof (ucontext: Logic.unconstrain_context) = let fun hyp_map hyp = (case AList.lookup (op =) (#constraints ucontext) hyp of SOME t => Hyp t | NONE => raise Fail "unconstrainT_proof: missing constraint"); val typ = Term_Subst.map_atypsT_same (Type.strip_sorts o #atyp_map ucontext); fun ofclass (ty, c) = let val ty' = Term.map_atyps (#atyp_map ucontext) ty; in the_single (of_sort_proof algebra classrel_proof arity_proof hyp_map (ty', [c])) end; in Same.commit (map_proof_same (Term_Subst.map_types_same typ) typ ofclass) #> fold_rev (implies_intr_proof o snd) (#constraints ucontext) end; fun export_proof thy i prop prf0 = let val prf = prf0 |> reconstruct_proof thy prop |> apply_preproc thy; val (prop', SOME prf') = (prop, SOME prf) |> standard_vars Name.context; val args = [] |> add_standard_vars_term prop' |> add_standard_vars prf' |> rev; val typargs = [] |> Term.add_tfrees prop' |> fold_proof_terms Term.add_tfrees prf' |> rev; val consts = Sign.consts_of thy; val xml = (typargs, (args, (prop', no_thm_names prf'))) |> let open XML.Encode Term_XML.Encode; val encode_vars = list (pair string typ); val encode_term = encode_standard_term consts; val encode_proof = encode_standard_proof consts; in pair (list (pair string sort)) (pair encode_vars (pair encode_term encode_proof)) end; in Export.export_params {theory = thy, binding = Path.binding0 (Path.make ["proofs", string_of_int i]), executable = false, compress = true, strict = false} xml end; fun prepare_thm_proof unconstrain thy classrel_proof arity_proof (name, pos) shyps hyps concl promises body = let val named = name <> ""; val prop = Logic.list_implies (hyps, concl); val args = prop_args prop; val (ucontext, prop1) = Logic.unconstrainT shyps prop; val PBody {oracles = oracles0, thms = thms0, proof = prf} = body; val body0 = Future.value (PBody {oracles = oracles0, thms = thms0, proof = if proofs_enabled () then fold_rev implies_intr_proof hyps prf else MinProof}); fun new_prf () = let val i = serial (); val unconstrainT = unconstrainT_proof (Sign.classes_of thy) classrel_proof arity_proof ucontext; val postproc = map_proof_of (unconstrainT #> named ? rew_proof thy); in (i, fulfill_proof_future thy promises postproc body0) end; val (i, body') = (*somewhat non-deterministic proof boxes!*) if export_enabled () then new_prf () else (case strip_combt (fst (strip_combP prf)) of (PThm ({serial = ser, name = a, prop = prop', types = NONE, ...}, thm_body'), args') => if (a = "" orelse a = name) andalso prop' = prop1 andalso args' = args then let val Thm_Body {body = body', ...} = thm_body'; val i = if a = "" andalso named then serial () else ser; in (i, body' |> ser <> i ? Future.map (map_proof_of (rew_proof thy))) end else new_prf () | _ => new_prf ()); val open_proof = not named ? rew_proof thy; val export = if export_enabled () then Lazy.lazy (fn () => join_proof body' |> open_proof |> export_proof thy i prop1 handle exn => if Exn.is_interrupt exn then raise Fail ("Interrupt: potential resource problems while exporting proof " ^ string_of_int i) else Exn.reraise exn) else no_export; val thm_body = prune_body body'; val theory_name = Context.theory_long_name thy; val thm = (i, make_thm_node theory_name name prop1 thm_body export); val header = thm_header i ([pos, Position.thread_data ()]) theory_name name prop1 NONE; val head = PThm (header, Thm_Body {open_proof = open_proof, body = thm_body}); val proof = if unconstrain then proof_combt' (head, (map o Option.map o Term.map_types) (#map_atyps ucontext) args) else proof_combP (proof_combt' (head, args), map PClass (#outer_constraints ucontext) @ map Hyp hyps); in (thm, proof) end; in fun thm_proof thy = prepare_thm_proof false thy; fun unconstrain_thm_proof thy classrel_proof arity_proof shyps concl promises body = prepare_thm_proof true thy classrel_proof arity_proof ("", Position.none) shyps [] concl promises body; end; (* PThm identity *) fun get_identity shyps hyps prop prf = let val (_, prop) = Logic.unconstrainT shyps (Logic.list_implies (hyps, prop)) in (case fst (strip_combt (fst (strip_combP prf))) of PThm ({serial, theory_name, name, prop = prop', ...}, _) => if prop = prop' then SOME {serial = serial, theory_name = theory_name, name = name} else NONE | _ => NONE) end; fun get_approximative_name shyps hyps prop prf = Option.map #name (get_identity shyps hyps prop prf) |> the_default ""; (* thm_id *) type thm_id = {serial: serial, theory_name: string}; fun make_thm_id (serial, theory_name) : thm_id = {serial = serial, theory_name = theory_name}; fun thm_header_id ({serial, theory_name, ...}: thm_header) = make_thm_id (serial, theory_name); fun thm_id (serial, thm_node) : thm_id = make_thm_id (serial, thm_node_theory_name thm_node); fun get_id shyps hyps prop prf : thm_id option = (case get_identity shyps hyps prop prf of NONE => NONE | SOME {name = "", ...} => NONE | SOME {serial, theory_name, ...} => SOME (make_thm_id (serial, theory_name))); fun this_id NONE _ = false | this_id (SOME (thm_id: thm_id)) (thm_id': thm_id) = #serial thm_id = #serial thm_id'; (* proof boxes: intermediate PThm nodes *) fun proof_boxes {included, excluded} proofs = let fun boxes_of (Abst (_, _, prf)) = boxes_of prf | boxes_of (AbsP (_, _, prf)) = boxes_of prf | boxes_of (prf % _) = boxes_of prf | boxes_of (prf1 %% prf2) = boxes_of prf1 #> boxes_of prf2 | boxes_of (PThm (header as {serial = i, ...}, thm_body)) = (fn boxes => let val thm_id = thm_header_id header in if Inttab.defined boxes i orelse (excluded thm_id andalso not (included thm_id)) then boxes else let val prf' = thm_body_proof_open thm_body; val boxes' = Inttab.update (i, (header, prf')) boxes; in boxes_of prf' boxes' end end) | boxes_of MinProof = raise MIN_PROOF () | boxes_of _ = I; in Inttab.fold_rev (cons o #2) (Inttab.build (fold boxes_of proofs)) [] end; end; structure Basic_Proofterm = struct datatype proof = datatype Proofterm.proof datatype proof_body = datatype Proofterm.proof_body val op %> = Proofterm.%> end; open Basic_Proofterm; diff --git a/src/Pure/sorts.ML b/src/Pure/sorts.ML --- a/src/Pure/sorts.ML +++ b/src/Pure/sorts.ML @@ -1,492 +1,492 @@ (* Title: Pure/sorts.ML Author: Markus Wenzel and Stefan Berghofer, TU Muenchen The order-sorted algebra of type classes. Classes denote (possibly empty) collections of types that are partially ordered by class inclusion. They are represented symbolically by strings. Sorts are intersections of finitely many classes. They are represented by lists of classes. Normal forms of sorts are sorted lists of minimal classes (wrt. current class inclusion). *) signature SORTS = sig val make: sort list -> sort Ord_List.T val subset: sort Ord_List.T * sort Ord_List.T -> bool val union: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T val subtract: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T val remove_sort: sort -> sort Ord_List.T -> sort Ord_List.T val insert_sort: sort -> sort Ord_List.T -> sort Ord_List.T val insert_typ: typ -> sort Ord_List.T -> sort Ord_List.T val insert_typs: typ list -> sort Ord_List.T -> sort Ord_List.T val insert_term: term -> sort Ord_List.T -> sort Ord_List.T val insert_terms: term list -> sort Ord_List.T -> sort Ord_List.T type algebra val classes_of: algebra -> serial Graph.T val arities_of: algebra -> (class * sort list) list Symtab.table val all_classes: algebra -> class list val super_classes: algebra -> class -> class list val class_less: algebra -> class * class -> bool val class_le: algebra -> class * class -> bool val sort_eq: algebra -> sort * sort -> bool val sort_le: algebra -> sort * sort -> bool val sorts_le: algebra -> sort list * sort list -> bool val inter_sort: algebra -> sort * sort -> sort val minimize_sort: algebra -> sort -> sort val complete_sort: algebra -> sort -> sort val add_class: Context.generic -> class * class list -> algebra -> algebra val add_classrel: Context.generic -> class * class -> algebra -> algebra val add_arities: Context.generic -> string * (class * sort list) list -> algebra -> algebra val empty_algebra: algebra val merge_algebra: Context.generic -> algebra * algebra -> algebra val dest_algebra: algebra list -> algebra -> {classrel: (class * class list) list, arities: (string * sort list * class) list} val subalgebra: Context.generic -> (class -> bool) -> (class * string -> sort list option) -> algebra -> (sort -> sort) * algebra type class_error val class_error: Context.generic -> class_error -> string exception CLASS_ERROR of class_error val has_instance: algebra -> string -> sort -> bool val mg_domain: algebra -> string -> sort -> sort list (*exception CLASS_ERROR*) val meet_sort: algebra -> typ * sort -> sort Vartab.table -> sort Vartab.table (*exception CLASS_ERROR*) val meet_sort_typ: algebra -> typ * sort -> typ -> typ (*exception CLASS_ERROR*) val of_sort: algebra -> typ * sort -> bool val of_sort_derivation: algebra -> {class_relation: typ -> bool -> 'a * class -> class -> 'a, type_constructor: string * typ list -> ('a * class) list list -> class -> 'a, type_variable: typ -> ('a * class) list} -> typ * sort -> 'a list (*exception CLASS_ERROR*) val classrel_derivation: algebra -> ('a * class -> class -> 'a) -> 'a * class -> class -> 'a (*exception CLASS_ERROR*) val witness_sorts: algebra -> string list -> (typ * sort) list -> sort list -> (typ * sort) list end; structure Sorts: SORTS = struct (** ordered lists of sorts **) val make = Ord_List.make Term_Ord.sort_ord; val subset = Ord_List.subset Term_Ord.sort_ord; val union = Ord_List.union Term_Ord.sort_ord; val subtract = Ord_List.subtract Term_Ord.sort_ord; val remove_sort = Ord_List.remove Term_Ord.sort_ord; val insert_sort = Ord_List.insert Term_Ord.sort_ord; fun insert_typ (TFree (_, S)) Ss = insert_sort S Ss | insert_typ (TVar (_, S)) Ss = insert_sort S Ss | insert_typ (Type (_, Ts)) Ss = insert_typs Ts Ss and insert_typs [] Ss = Ss | insert_typs (T :: Ts) Ss = insert_typs Ts (insert_typ T Ss); fun insert_term (Const (_, T)) Ss = insert_typ T Ss | insert_term (Free (_, T)) Ss = insert_typ T Ss | insert_term (Var (_, T)) Ss = insert_typ T Ss | insert_term (Bound _) Ss = Ss | insert_term (Abs (_, T, t)) Ss = insert_term t (insert_typ T Ss) | insert_term (t $ u) Ss = insert_term t (insert_term u Ss); fun insert_terms [] Ss = Ss | insert_terms (t :: ts) Ss = insert_terms ts (insert_term t Ss); (** order-sorted algebra **) (* classes: graph representing class declarations together with proper subclass relation, which needs to be transitive and acyclic. arities: table of association lists of all type arities; (t, ars) means that type constructor t has the arities ars; an element (c, Ss) of ars represents the arity t::(Ss)c. "Coregularity" of the arities structure requires that for any two declarations t::(Ss1)c1 and t::(Ss2)c2 such that c1 <= c2 holds Ss1 <= Ss2. *) datatype algebra = Algebra of {classes: serial Graph.T, arities: (class * sort list) list Symtab.table}; fun classes_of (Algebra {classes, ...}) = classes; fun arities_of (Algebra {arities, ...}) = arities; fun make_algebra (classes, arities) = Algebra {classes = classes, arities = arities}; fun map_classes f (Algebra {classes, arities}) = make_algebra (f classes, arities); fun map_arities f (Algebra {classes, arities}) = make_algebra (classes, f arities); (* classes *) fun all_classes (Algebra {classes, ...}) = Graph.all_preds classes (Graph.maximals classes); val super_classes = Graph.immediate_succs o classes_of; (* class relations *) val class_less : algebra -> class * class -> bool = Graph.is_edge o classes_of; fun class_le algebra (c1, c2) = c1 = c2 orelse class_less algebra (c1, c2); (* sort relations *) fun sort_le algebra (S1: sort, S2: sort) = S1 = S2 orelse forall (fn c2 => exists (fn c1 => class_le algebra (c1, c2)) S1) S2; fun sorts_le algebra (Ss1, Ss2) = ListPair.all (sort_le algebra) (Ss1, Ss2); fun sort_eq algebra (S1, S2) = sort_le algebra (S1, S2) andalso sort_le algebra (S2, S1); (* intersection *) fun inter_class algebra c S = let fun intr [] = [c] | intr (S' as c' :: c's) = if class_le algebra (c', c) then S' else if class_le algebra (c, c') then intr c's else c' :: intr c's in intr S end; fun inter_sort algebra (S1, S2) = sort_strings (fold (inter_class algebra) S1 S2); (* normal forms *) fun minimize_sort _ [] = [] | minimize_sort _ (S as [_]) = S | minimize_sort algebra S = filter (fn c => not (exists (fn c' => class_less algebra (c', c)) S)) S |> sort_distinct string_ord; fun complete_sort algebra = Graph.all_succs (classes_of algebra) o minimize_sort algebra; (** build algebras **) (* classes *) fun err_dup_class c = error ("Duplicate declaration of class: " ^ quote c); fun err_cyclic_classes context css = error (cat_lines (map (fn cs => "Cycle in class relation: " ^ Syntax.string_of_classrel (Syntax.init_pretty context) cs) css)); fun add_class context (c, cs) = map_classes (fn classes => let val classes' = classes |> Graph.new_node (c, serial ()) handle Graph.DUP dup => err_dup_class dup; val classes'' = classes' |> fold Graph.add_edge_trans_acyclic (map (pair c) cs) handle Graph.CYCLES css => err_cyclic_classes context css; in classes'' end); (* arities *) local fun for_classes _ NONE = "" | for_classes ctxt (SOME (c1, c2)) = " for classes " ^ Syntax.string_of_classrel ctxt [c1, c2]; fun err_conflict context t cc (c, Ss) (c', Ss') = let val ctxt = Syntax.init_pretty context in error ("Conflict of type arities" ^ for_classes ctxt cc ^ ":\n " ^ Syntax.string_of_arity ctxt (t, Ss, [c]) ^ " and\n " ^ Syntax.string_of_arity ctxt (t, Ss', [c'])) end; fun coregular context algebra t (c, Ss) ars = let fun conflict (c', Ss') = if class_le algebra (c, c') andalso not (sorts_le algebra (Ss, Ss')) then SOME ((c, c'), (c', Ss')) else if class_le algebra (c', c) andalso not (sorts_le algebra (Ss', Ss)) then SOME ((c', c), (c', Ss')) else NONE; in (case get_first conflict ars of SOME ((c1, c2), (c', Ss')) => err_conflict context t (SOME (c1, c2)) (c, Ss) (c', Ss') | NONE => (c, Ss) :: ars) end; fun complete algebra (c, Ss) = map (rpair Ss) (c :: super_classes algebra c); fun insert context algebra t (c, Ss) ars = (case AList.lookup (op =) ars c of NONE => coregular context algebra t (c, Ss) ars | SOME Ss' => if sorts_le algebra (Ss, Ss') then ars else if sorts_le algebra (Ss', Ss) then coregular context algebra t (c, Ss) (remove (op =) (c, Ss') ars) else err_conflict context t NONE (c, Ss) (c, Ss')); in fun insert_ars context algebra t = fold_rev (insert context algebra t); fun insert_complete_ars context algebra (t, ars) arities = let val ars' = Symtab.lookup_list arities t |> fold_rev (insert_ars context algebra t) (map (complete algebra) ars); in Symtab.update (t, ars') arities end; fun add_arities context arg algebra = algebra |> map_arities (insert_complete_ars context algebra arg); fun add_arities_table context algebra = Symtab.fold (fn (t, ars) => insert_complete_ars context algebra (t, ars)); end; (* classrel *) fun rebuild_arities context algebra = algebra |> map_arities (Symtab.build o add_arities_table context algebra); fun add_classrel context rel = rebuild_arities context o map_classes (fn classes => classes |> Graph.add_edge_trans_acyclic rel handle Graph.CYCLES css => err_cyclic_classes context css); (* empty and merge *) val empty_algebra = make_algebra (Graph.empty, Symtab.empty); fun merge_algebra context (Algebra {classes = classes1, arities = arities1}, Algebra {classes = classes2, arities = arities2}) = let val classes' = Graph.merge_trans_acyclic (op =) (classes1, classes2) handle Graph.DUP c => err_dup_class c | Graph.CYCLES css => err_cyclic_classes context css; val algebra0 = make_algebra (classes', Symtab.empty); val arities' = (case (pointer_eq (classes1, classes2), pointer_eq (arities1, arities2)) of (true, true) => arities1 | (true, false) => (*no completion*) (arities1, arities2) |> Symtab.join (fn t => fn (ars1, ars2) => if pointer_eq (ars1, ars2) then raise Symtab.SAME else insert_ars context algebra0 t ars2 ars1) | (false, true) => (*unary completion*) Symtab.build (add_arities_table context algebra0 arities1) | (false, false) => (*binary completion*) Symtab.build (add_arities_table context algebra0 arities1 #> add_arities_table context algebra0 arities2)); in make_algebra (classes', arities') end; (* destruct *) fun dest_algebra parents (Algebra {classes, arities}) = let fun new_classrel rel = not (exists (fn algebra => class_less algebra rel) parents); val classrel = - (classes, []) |-> Graph.fold (fn (c, (_, (_, ds))) => + build (classes |> Graph.fold (fn (c, (_, (_, ds))) => (case filter (fn d => new_classrel (c, d)) (Graph.Keys.dest ds) of [] => I - | ds' => cons (c, sort_strings ds'))) + | ds' => cons (c, sort_strings ds')))) |> sort_by #1; fun is_arity t ar algebra = member (op =) (Symtab.lookup_list (arities_of algebra) t) ar; fun add_arity t (c, Ss) = not (exists (is_arity t (c, Ss)) parents) ? cons (t, Ss, c); val arities = - (arities, []) |-> Symtab.fold (fn (t, ars) => fold_rev (add_arity t) (sort_by #1 ars)) + build (arities |> Symtab.fold (fn (t, ars) => fold_rev (add_arity t) (sort_by #1 ars))) |> sort_by #1; in {classrel = classrel, arities = arities} end; (* algebra projections *) (* FIXME potentially violates abstract type integrity *) fun subalgebra context P sargs (algebra as Algebra {classes, arities}) = let val restrict_sort = minimize_sort algebra o filter P o Graph.all_succs classes; fun restrict_arity t (c, Ss) = if P c then (case sargs (c, t) of SOME sorts => SOME (c, Ss |> map2 (curry (inter_sort algebra)) sorts |> map restrict_sort) | NONE => NONE) else NONE; val classes' = classes |> Graph.restrict P; val arities' = arities |> Symtab.map (map_filter o restrict_arity); in (restrict_sort, rebuild_arities context (make_algebra (classes', arities'))) end; (** sorts of types **) (* errors -- performance tuning via delayed message composition *) datatype class_error = No_Classrel of class * class | No_Arity of string * class | No_Subsort of sort * sort; fun class_error context = let val ctxt = Syntax.init_pretty context in fn No_Classrel (c1, c2) => "No class relation " ^ Syntax.string_of_classrel ctxt [c1, c2] | No_Arity (a, c) => "No type arity " ^ Syntax.string_of_arity ctxt (a, [], [c]) | No_Subsort (S1, S2) => "Cannot derive subsort relation " ^ Syntax.string_of_sort ctxt S1 ^ " < " ^ Syntax.string_of_sort ctxt S2 end; exception CLASS_ERROR of class_error; (* instances *) fun has_instance algebra a = forall (AList.defined (op =) (Symtab.lookup_list (arities_of algebra) a)); fun mg_domain algebra a S = let val ars = Symtab.lookup_list (arities_of algebra) a; fun dom c = (case AList.lookup (op =) ars c of NONE => raise CLASS_ERROR (No_Arity (a, c)) | SOME Ss => Ss); fun dom_inter c Ss = ListPair.map (inter_sort algebra) (dom c, Ss); in (case S of [] => raise Fail "Unknown domain of empty intersection" | c :: cs => fold dom_inter cs (dom c)) end; (* meet_sort *) fun meet_sort algebra = let fun inters S S' = inter_sort algebra (S, S'); fun meet _ [] = I | meet (TFree (_, S)) S' = if sort_le algebra (S, S') then I else raise CLASS_ERROR (No_Subsort (S, S')) | meet (TVar (v, S)) S' = if sort_le algebra (S, S') then I else Vartab.map_default (v, S) (inters S') | meet (Type (a, Ts)) S = fold2 meet Ts (mg_domain algebra a S); in uncurry meet end; fun meet_sort_typ algebra (T, S) = let val tab = Vartab.build (meet_sort algebra (T, S)); in Term.map_type_tvar (fn (v, _) => TVar (v, (the o Vartab.lookup tab) v)) end; (* of_sort *) fun of_sort algebra = let fun ofS (_, []) = true | ofS (TFree (_, S), S') = sort_le algebra (S, S') | ofS (TVar (_, S), S') = sort_le algebra (S, S') | ofS (Type (a, Ts), S) = let val Ss = mg_domain algebra a S in ListPair.all ofS (Ts, Ss) end handle CLASS_ERROR _ => false; in ofS end; (* animating derivations *) fun of_sort_derivation algebra {class_relation, type_constructor, type_variable} = let val arities = arities_of algebra; fun weaken T D1 S2 = let val S1 = map snd D1 in if S1 = S2 then map fst D1 else S2 |> map (fn c2 => (case D1 |> filter (fn (_, c1) => class_le algebra (c1, c2)) of [d1] => class_relation T true d1 c2 | (d1 :: _ :: _) => class_relation T false d1 c2 | [] => raise CLASS_ERROR (No_Subsort (S1, S2)))) end; fun derive (_, []) = [] | derive (Type (a, Us), S) = let val Ss = mg_domain algebra a S; val dom = map2 (fn U => fn S => derive (U, S) ~~ S) Us Ss; in S |> map (fn c => let val Ss' = the (AList.lookup (op =) (Symtab.lookup_list arities a) c); val dom' = map (fn ((U, d), S') => weaken U d S' ~~ S') ((Us ~~ dom) ~~ Ss'); in type_constructor (a, Us) dom' c end) end | derive (T, S) = weaken T (type_variable T) S; in derive end; fun classrel_derivation algebra class_relation = let fun path (x, c1 :: c2 :: cs) = path (class_relation (x, c1) c2, c2 :: cs) | path (x, _) = x; in fn (x, c1) => fn c2 => (case Graph.irreducible_paths (classes_of algebra) (c1, c2) of [] => raise CLASS_ERROR (No_Classrel (c1, c2)) | cs :: _ => path (x, cs)) end; (* witness_sorts *) fun witness_sorts algebra ground_types hyps sorts = let fun le S1 S2 = sort_le algebra (S1, S2); fun get S2 (T, S1) = if le S1 S2 then SOME (T, S2) else NONE; fun mg_dom t S = SOME (mg_domain algebra t S) handle CLASS_ERROR _ => NONE; fun witn_sort _ [] solved_failed = (SOME (propT, []), solved_failed) | witn_sort path S (solved, failed) = if exists (le S) failed then (NONE, (solved, failed)) else (case get_first (get S) solved of SOME w => (SOME w, (solved, failed)) | NONE => (case get_first (get S) hyps of SOME w => (SOME w, (w :: solved, failed)) | NONE => witn_types path ground_types S (solved, failed))) and witn_sorts path x = fold_map (witn_sort path) x and witn_types _ [] S (solved, failed) = (NONE, (solved, S :: failed)) | witn_types path (t :: ts) S solved_failed = (case mg_dom t S of SOME SS => (*do not descend into stronger args (achieving termination)*) if exists (fn D => le D S orelse exists (le D) path) SS then witn_types path ts S solved_failed else let val (ws, (solved', failed')) = witn_sorts (S :: path) SS solved_failed in if forall is_some ws then let val w = (Type (t, map (#1 o the) ws), S) in (SOME w, (w :: solved', failed')) end else witn_types path ts S (solved', failed') end | NONE => witn_types path ts S solved_failed); in map_filter I (#1 (witn_sorts [] sorts ([], []))) end; end; diff --git a/src/Pure/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2373 +1,2373 @@ (* Title: Pure/thm.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Makarius The very core of Isabelle's Meta Logic: certified types and terms, derivations, theorems, inference rules (including lifting and resolution), oracles. *) infix 0 RS RSN; signature BASIC_THM = sig type ctyp type cterm exception CTERM of string * cterm list type thm type conv = cterm -> thm exception THM of string * int * thm list val RSN: thm * (int * thm) -> thm val RS: thm * thm -> thm end; signature THM = sig include BASIC_THM (*certified types*) val typ_of: ctyp -> typ val global_ctyp_of: theory -> typ -> ctyp val ctyp_of: Proof.context -> typ -> ctyp val dest_ctyp: ctyp -> ctyp list val dest_ctypN: int -> ctyp -> ctyp val dest_ctyp0: ctyp -> ctyp val dest_ctyp1: ctyp -> ctyp val make_ctyp: ctyp -> ctyp list -> ctyp (*certified terms*) val term_of: cterm -> term val typ_of_cterm: cterm -> typ val ctyp_of_cterm: cterm -> ctyp val maxidx_of_cterm: cterm -> int val global_cterm_of: theory -> term -> cterm val cterm_of: Proof.context -> term -> cterm val renamed_term: term -> cterm -> cterm val dest_comb: cterm -> cterm * cterm val dest_fun: cterm -> cterm val dest_arg: cterm -> cterm val dest_fun2: cterm -> cterm val dest_arg1: cterm -> cterm val dest_abs: string option -> cterm -> cterm * cterm val rename_tvar: indexname -> ctyp -> ctyp val var: indexname * ctyp -> cterm val apply: cterm -> cterm -> cterm val lambda_name: string * cterm -> cterm -> cterm val lambda: cterm -> cterm -> cterm val adjust_maxidx_cterm: int -> cterm -> cterm val incr_indexes_cterm: int -> cterm -> cterm val match: cterm * cterm -> ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list val first_order_match: cterm * cterm -> ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list (*theorems*) val fold_terms: (term -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_ctyps: (ctyp -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_cterms: (cterm -> 'a -> 'a) -> thm -> 'a -> 'a val terms_of_tpairs: (term * term) list -> term list val full_prop_of: thm -> term val theory_id: thm -> Context.theory_id val theory_name: thm -> string val maxidx_of: thm -> int val maxidx_thm: thm -> int -> int val shyps_of: thm -> sort Ord_List.T val hyps_of: thm -> term list val prop_of: thm -> term val tpairs_of: thm -> (term * term) list val concl_of: thm -> term val prems_of: thm -> term list val nprems_of: thm -> int val no_prems: thm -> bool val major_prem_of: thm -> term val cprop_of: thm -> cterm val cprem_of: thm -> int -> cterm val cconcl_of: thm -> cterm val cprems_of: thm -> cterm list val chyps_of: thm -> cterm list exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option val theory_of_cterm: cterm -> theory val theory_of_thm: thm -> theory val trim_context_ctyp: ctyp -> ctyp val trim_context_cterm: cterm -> cterm val transfer_ctyp: theory -> ctyp -> ctyp val transfer_cterm: theory -> cterm -> cterm val transfer: theory -> thm -> thm val transfer': Proof.context -> thm -> thm val transfer'': Context.generic -> thm -> thm val join_transfer: theory -> thm -> thm val join_transfer_context: Proof.context * thm -> Proof.context * thm val renamed_prop: term -> thm -> thm val weaken: cterm -> thm -> thm val weaken_sorts: sort list -> cterm -> cterm val proof_bodies_of: thm list -> proof_body list val proof_body_of: thm -> proof_body val proof_of: thm -> proof val reconstruct_proof_of: thm -> Proofterm.proof val consolidate: thm list -> unit val expose_proofs: theory -> thm list -> unit val expose_proof: theory -> thm -> unit val future: thm future -> cterm -> thm val thm_deps: thm -> Proofterm.thm Ord_List.T val extra_shyps: thm -> sort list val strip_shyps: thm -> thm val derivation_closed: thm -> bool val derivation_name: thm -> string val derivation_id: thm -> Proofterm.thm_id option val raw_derivation_name: thm -> string val expand_name: thm -> Proofterm.thm_header -> string option val name_derivation: string * Position.T -> thm -> thm val close_derivation: Position.T -> thm -> thm val trim_context: thm -> thm val axiom: theory -> string -> thm val all_axioms_of: theory -> (string * thm) list val get_tags: thm -> Properties.T val map_tags: (Properties.T -> Properties.T) -> thm -> thm val norm_proof: thm -> thm val adjust_maxidx_thm: int -> thm -> thm (*type classes*) val the_classrel: theory -> class * class -> thm val the_arity: theory -> string * sort list * class -> thm val classrel_proof: theory -> class * class -> proof val arity_proof: theory -> string * sort list * class -> proof (*oracles*) val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory val oracle_space: theory -> Name_Space.T val pretty_oracle: Proof.context -> string -> Pretty.T val extern_oracles: bool -> Proof.context -> (Markup.T * xstring) list val check_oracle: Proof.context -> xstring * Position.T -> string (*inference rules*) val assume: cterm -> thm val implies_intr: cterm -> thm -> thm val implies_elim: thm -> thm -> thm val forall_intr: cterm -> thm -> thm val forall_elim: cterm -> thm -> thm val reflexive: cterm -> thm val symmetric: thm -> thm val transitive: thm -> thm -> thm val beta_conversion: bool -> conv val eta_conversion: conv val eta_long_conversion: conv val abstract_rule: string -> cterm -> thm -> thm val combination: thm -> thm -> thm val equal_intr: thm -> thm -> thm val equal_elim: thm -> thm -> thm val solve_constraints: thm -> thm val flexflex_rule: Proof.context option -> thm -> thm Seq.seq val generalize: Symtab.set * Symtab.set -> int -> thm -> thm val generalize_cterm: Symtab.set * Symtab.set -> int -> cterm -> cterm val generalize_ctyp: Symtab.set -> int -> ctyp -> ctyp val instantiate: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> thm -> thm val instantiate_cterm: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> cterm -> cterm val trivial: cterm -> thm val of_class: ctyp * class -> thm val unconstrainT: thm -> thm val varifyT_global': Term_Subst.TFrees.set -> thm -> ((string * sort) * indexname) list * thm val varifyT_global: thm -> thm val legacy_freezeT: thm -> thm val plain_prop_of: thm -> term val dest_state: thm * int -> (term * term) list * term list * term * term val lift_rule: cterm -> thm -> thm val incr_indexes: int -> thm -> thm val assumption: Proof.context option -> int -> thm -> thm Seq.seq val eq_assumption: int -> thm -> thm val rotate_rule: int -> int -> thm -> thm val permute_prems: int -> int -> thm -> thm val bicompose: Proof.context option -> {flatten: bool, match: bool, incremented: bool} -> bool * thm * int -> int -> thm -> thm Seq.seq val biresolution: Proof.context option -> bool -> (bool * thm) list -> int -> thm -> thm Seq.seq val thynames_of_arity: theory -> string * class -> string list val add_classrel: thm -> theory -> theory val add_arity: thm -> theory -> theory end; structure Thm: THM = struct (*** Certified terms and types ***) (** certified types **) datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: sort Ord_List.T}; fun typ_of (Ctyp {T, ...}) = T; fun global_ctyp_of thy raw_T = let val T = Sign.certify_typ thy raw_T; val maxidx = Term.maxidx_of_typ T; val sorts = Sorts.insert_typ T []; in Ctyp {cert = Context.Certificate thy, T = T, maxidx = maxidx, sorts = sorts} end; val ctyp_of = global_ctyp_of o Proof_Context.theory_of; fun dest_ctyp (Ctyp {cert, T = Type (_, Ts), maxidx, sorts}) = map (fn T => Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}) Ts | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []); fun dest_ctypN n (Ctyp {cert, T, maxidx, sorts}) = let fun err () = raise TYPE ("dest_ctypN", [T], []) in (case T of Type (_, Ts) => Ctyp {cert = cert, T = nth Ts n handle General.Subscript => err (), maxidx = maxidx, sorts = sorts} | _ => err ()) end; val dest_ctyp0 = dest_ctypN 0; val dest_ctyp1 = dest_ctypN 1; fun join_certificate_ctyp (Ctyp {cert, ...}) cert0 = Context.join_certificate (cert0, cert); fun union_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sorts.union sorts0 sorts; fun maxidx_ctyp (Ctyp {maxidx, ...}) maxidx0 = Int.max (maxidx0, maxidx); fun make_ctyp (Ctyp {cert, T, maxidx = _, sorts = _}) cargs = let val As = map typ_of cargs; fun err () = raise TYPE ("make_ctyp", T :: As, []); in (case T of Type (a, args) => Ctyp { cert = fold join_certificate_ctyp cargs cert, maxidx = fold maxidx_ctyp cargs ~1, sorts = fold union_sorts_ctyp cargs [], T = if length args = length cargs then Type (a, As) else err ()} | _ => err ()) end; (** certified terms **) (*certified terms with checked typ, maxidx, and sorts*) datatype cterm = Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: sort Ord_List.T}; exception CTERM of string * cterm list; fun term_of (Cterm {t, ...}) = t; fun typ_of_cterm (Cterm {T, ...}) = T; fun ctyp_of_cterm (Cterm {cert, T, maxidx, sorts, ...}) = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}; fun maxidx_of_cterm (Cterm {maxidx, ...}) = maxidx; fun global_cterm_of thy tm = let val (t, T, maxidx) = Sign.certify_term thy tm; val sorts = Sorts.insert_term t []; in Cterm {cert = Context.Certificate thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end; val cterm_of = global_cterm_of o Proof_Context.theory_of; fun join_certificate0 (Cterm {cert = cert1, ...}, Cterm {cert = cert2, ...}) = Context.join_certificate (cert1, cert2); fun renamed_term t' (Cterm {cert, t, T, maxidx, sorts}) = if t aconv t' then Cterm {cert = cert, t = t', T = T, maxidx = maxidx, sorts = sorts} else raise TERM ("renamed_term: terms disagree", [t, t']); (* destructors *) fun dest_comb (Cterm {t = c $ a, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in (Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_comb ct = raise CTERM ("dest_comb", [ct]); fun dest_fun (Cterm {t = c $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun ct = raise CTERM ("dest_fun", [ct]); fun dest_arg (Cterm {t = c $ a, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg ct = raise CTERM ("dest_arg", [ct]); fun dest_fun2 (Cterm {t = c $ _ $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0; val B = Term.argument_type_of c 1; in Cterm {t = c, T = A --> B --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun2 ct = raise CTERM ("dest_fun2", [ct]); fun dest_arg1 (Cterm {t = c $ a $ _, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]); fun dest_abs a (Cterm {t = Abs (x, T, t), T = Type ("fun", [_, U]), cert, maxidx, sorts}) = let val (y', t') = Term.dest_abs (the_default x a, T, t) in (Cterm {t = Free (y', T), T = T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = t', T = U, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_abs _ ct = raise CTERM ("dest_abs", [ct]); (* constructors *) fun rename_tvar (a, i) (Ctyp {cert, T, maxidx, sorts}) = let val S = (case T of TFree (_, S) => S | TVar (_, S) => S | _ => raise TYPE ("rename_tvar: no variable", [T], [])); val _ = if i < 0 then raise TYPE ("rename_tvar: bad index", [TVar ((a, i), S)], []) else (); in Ctyp {cert = cert, T = TVar ((a, i), S), maxidx = Int.max (i, maxidx), sorts = sorts} end; fun var ((x, i), Ctyp {cert, T, maxidx, sorts}) = if i < 0 then raise TERM ("var: bad index", [Var ((x, i), T)]) else Cterm {cert = cert, t = Var ((x, i), T), T = T, maxidx = Int.max (i, maxidx), sorts = sorts}; fun apply (cf as Cterm {t = f, T = Type ("fun", [dty, rty]), maxidx = maxidx1, sorts = sorts1, ...}) (cx as Cterm {t = x, T, maxidx = maxidx2, sorts = sorts2, ...}) = if T = dty then Cterm {cert = join_certificate0 (cf, cx), t = f $ x, T = rty, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} else raise CTERM ("apply: types don't agree", [cf, cx]) | apply cf cx = raise CTERM ("apply: first arg is not a function", [cf, cx]); fun lambda_name (x, ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...}) (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) = let val t = Term.lambda_name (x, t1) t2 in Cterm {cert = join_certificate0 (ct1, ct2), t = t, T = T1 --> T2, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} end; fun lambda t u = lambda_name ("", t) u; (* indexes *) fun adjust_maxidx_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if maxidx = i then ct else if maxidx < i then Cterm {maxidx = i, cert = cert, t = t, T = T, sorts = sorts} else Cterm {maxidx = Int.max (maxidx_of_term t, i), cert = cert, t = t, T = T, sorts = sorts}; fun incr_indexes_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if i < 0 then raise CTERM ("negative increment", [ct]) else if i = 0 then ct else Cterm {cert = cert, t = Logic.incr_indexes ([], [], i) t, T = Logic.incr_tvar i T, maxidx = maxidx + i, sorts = sorts}; (*** Derivations and Theorems ***) (* sort constraints *) type constraint = {theory: theory, typ: typ, sort: sort}; local val constraint_ord : constraint ord = Context.theory_id_ord o apply2 (Context.theory_id o #theory) ||| Term_Ord.typ_ord o apply2 #typ ||| Term_Ord.sort_ord o apply2 #sort; val smash_atyps = map_atyps (fn TVar (_, S) => Term.aT S | TFree (_, S) => Term.aT S | T => T); in val union_constraints = Ord_List.union constraint_ord; fun insert_constraints thy (T, S) = let val ignored = S = [] orelse (case T of TFree (_, S') => S = S' | TVar (_, S') => S = S' | _ => false); in if ignored then I else Ord_List.insert constraint_ord {theory = thy, typ = smash_atyps T, sort = S} end; fun insert_constraints_env thy env = let val tyenv = Envir.type_env env; fun insert ([], _) = I | insert (S, T) = insert_constraints thy (Envir.norm_type tyenv T, S); in tyenv |> Vartab.fold (insert o #2) end; end; (* datatype thm *) datatype thm = Thm of deriv * (*derivation*) {cert: Context.certificate, (*background theory certificate*) tags: Properties.T, (*additional annotations/comments*) maxidx: int, (*maximum index of any Var or TVar*) constraints: constraint Ord_List.T, (*implicit proof obligations for sort constraints*) shyps: sort Ord_List.T, (*sort hypotheses*) hyps: term Ord_List.T, (*hypotheses*) tpairs: (term * term) list, (*flex-flex pairs*) prop: term} (*conclusion*) and deriv = Deriv of {promises: (serial * thm future) Ord_List.T, body: Proofterm.proof_body}; type conv = cterm -> thm; (*errors involving theorems*) exception THM of string * int * thm list; fun rep_thm (Thm (_, args)) = args; fun fold_terms f (Thm (_, {tpairs, prop, hyps, ...})) = fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps; fun fold_atomic_ctyps f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun ctyp T = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms o fold_types o fold_atyps) (f o ctyp) th end; fun fold_atomic_cterms f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun cterm t T = Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms o fold_aterms) (fn t as Const (_, T) => f (cterm t T) | t as Free (_, T) => f (cterm t T) | t as Var (_, T) => f (cterm t T) | _ => I) th end; fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs []; fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u'; fun union_tpairs ts us = Library.merge eq_tpairs (ts, us); val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u); fun attach_tpairs tpairs prop = Logic.list_implies (map Logic.mk_equals tpairs, prop); fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop; val union_hyps = Ord_List.union Term_Ord.fast_term_ord; val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord; val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord; fun join_certificate1 (Cterm {cert = cert1, ...}, Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); fun join_certificate2 (Thm (_, {cert = cert1, ...}), Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); (* basic components *) val cert_of = #cert o rep_thm; val theory_id = Context.certificate_theory_id o cert_of; val theory_name = Context.theory_id_name o theory_id; val maxidx_of = #maxidx o rep_thm; fun maxidx_thm th i = Int.max (maxidx_of th, i); val shyps_of = #shyps o rep_thm; val hyps_of = #hyps o rep_thm; val prop_of = #prop o rep_thm; val tpairs_of = #tpairs o rep_thm; val concl_of = Logic.strip_imp_concl o prop_of; val prems_of = Logic.strip_imp_prems o prop_of; val nprems_of = Logic.count_prems o prop_of; fun no_prems th = nprems_of th = 0; fun major_prem_of th = (case prems_of th of prem :: _ => Logic.strip_assums_concl prem | [] => raise THM ("major_prem_of: rule with no premises", 0, [th])); fun cprop_of (Thm (_, {cert, maxidx, shyps, prop, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, t = prop, sorts = shyps}; fun cprem_of (th as Thm (_, {cert, maxidx, shyps, prop, ...})) i = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])}; fun cconcl_of (th as Thm (_, {cert, maxidx, shyps, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = concl_of th}; fun cprems_of (th as Thm (_, {cert, maxidx, shyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = t}) (prems_of th); fun chyps_of (Thm (_, {cert, shyps, hyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t}) hyps; (* implicit theory context *) exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option; fun theory_of_cterm (ct as Cterm {cert, ...}) = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct], [], NONE); fun theory_of_thm th = Context.certificate_theory (cert_of th) handle ERROR msg => raise CONTEXT (msg, [], [], [th], NONE); fun trim_context_ctyp cT = (case cT of Ctyp {cert = Context.Certificate_Id _, ...} => cT | Ctyp {cert = Context.Certificate thy, T, maxidx, sorts} => Ctyp {cert = Context.Certificate_Id (Context.theory_id thy), T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_cterm ct = (case ct of Cterm {cert = Context.Certificate_Id _, ...} => ct | Cterm {cert = Context.Certificate thy, t, T, maxidx, sorts} => Cterm {cert = Context.Certificate_Id (Context.theory_id thy), t = t, T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_thm th = (case th of Thm (_, {constraints = _ :: _, ...}) => raise THM ("trim_context: pending sort constraints", 0, [th]) | Thm (_, {cert = Context.Certificate_Id _, ...}) => th | Thm (der, {cert = Context.Certificate thy, tags, maxidx, constraints = [], shyps, hyps, tpairs, prop}) => Thm (der, {cert = Context.Certificate_Id (Context.theory_id thy), tags = tags, maxidx = maxidx, constraints = [], shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})); fun transfer_ctyp thy' cT = let val Ctyp {cert, T, maxidx, sorts} = cT; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [cT], [], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then cT else Ctyp {cert = cert', T = T, maxidx = maxidx, sorts = sorts} end; fun transfer_cterm thy' ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [ct], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then ct else Cterm {cert = cert', t = t, T = T, maxidx = maxidx, sorts = sorts} end; fun transfer thy' th = let val Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) = th; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [], [th], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then th else Thm (der, {cert = cert', tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; val transfer' = transfer o Proof_Context.theory_of; val transfer'' = transfer o Context.theory_of; fun join_transfer thy th = (Context.subthy_id (theory_id th, Context.theory_id thy) ? transfer thy) th; fun join_transfer_context (ctxt, th) = if Context.subthy_id (theory_id th, Context.theory_id (Proof_Context.theory_of ctxt)) then (ctxt, transfer' ctxt th) else (Context.raw_transfer (theory_of_thm th) ctxt, th); (* matching *) local fun gen_match match (ct1 as Cterm {t = t1, sorts = sorts1, ...}, ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) = let val cert = join_certificate0 (ct1, ct2); val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct1, ct2], [], NONE); val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty); val sorts = Sorts.union sorts1 sorts2; fun mk_cTinst ((a, i), (S, T)) = (((a, i), S), Ctyp {T = T, cert = cert, maxidx = maxidx2, sorts = sorts}); fun mk_ctinst ((x, i), (U, t)) = let val T = Envir.subst_type Tinsts U in (((x, i), T), Cterm {t = t, T = T, cert = cert, maxidx = maxidx2, sorts = sorts}) end; in (Vartab.fold (cons o mk_cTinst) Tinsts [], Vartab.fold (cons o mk_ctinst) tinsts []) end; in val match = gen_match Pattern.match; val first_order_match = gen_match Pattern.first_order_match; end; (*implicit alpha-conversion*) fun renamed_prop prop' (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if prop aconv prop' then Thm (der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) else raise TERM ("renamed_prop: props disagree", [prop, prop']); fun make_context ths NONE cert = (Context.Theory (Context.certificate_theory cert) handle ERROR msg => raise CONTEXT (msg, [], [], ths, NONE)) | make_context ths (SOME ctxt) cert = let val thy_id = Context.certificate_theory_id cert; val thy_id' = Context.theory_id (Proof_Context.theory_of ctxt); in if Context.subthy_id (thy_id, thy_id') then Context.Proof ctxt else raise CONTEXT ("Bad context", [], [], ths, SOME (Context.Proof ctxt)) end; fun make_context_certificate ths opt_ctxt cert = let val context = make_context ths opt_ctxt cert; val cert' = Context.Certificate (Context.theory_of context); in (context, cert') end; (*explicit weakening: maps |- B to A |- B*) fun weaken raw_ct th = let val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct; val Thm (der, {tags, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; in if T <> propT then raise THM ("weaken: assumptions must have type prop", 0, []) else if maxidxA <> ~1 then raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, []) else Thm (der, {cert = join_certificate1 (ct, th), tags = tags, maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = insert_hyps A hyps, tpairs = tpairs, prop = prop}) end; fun weaken_sorts raw_sorts ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val thy = theory_of_cterm ct; val more_sorts = Sorts.make (map (Sign.certify_sort thy) raw_sorts); val sorts' = Sorts.union sorts more_sorts; in Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = sorts'} end; (** derivations and promised proofs **) fun make_deriv promises oracles thms proof = Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}}; val empty_deriv = make_deriv [] [] [] MinProof; (* inference rules *) val promise_ord: (serial * thm future) ord = fn ((i, _), (j, _)) => int_ord (j, i); fun bad_proofs i = error ("Illegal level of detail for proof objects: " ^ string_of_int i); fun deriv_rule2 f (Deriv {promises = ps1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}}) (Deriv {promises = ps2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) = let val ps = Ord_List.union promise_ord ps1 ps2; val oracles = Proofterm.unions_oracles [oracles1, oracles2]; val thms = Proofterm.unions_thms [thms1, thms2]; val prf = (case ! Proofterm.proofs of 2 => f prf1 prf2 | 1 => MinProof | 0 => MinProof | i => bad_proofs i); in make_deriv ps oracles thms prf end; fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv; fun deriv_rule0 make_prf = if ! Proofterm.proofs <= 1 then empty_deriv else deriv_rule1 I (make_deriv [] [] [] (make_prf ())); fun deriv_rule_unconditional f (Deriv {promises, body = PBody {oracles, thms, proof}}) = make_deriv promises oracles thms (f proof); (* fulfilled proofs *) fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises; fun join_promises [] = () | join_promises promises = join_promises_of (Future.joins (map snd promises)) and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms)); fun fulfill_body (th as Thm (Deriv {promises, body}, _)) = let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises)) in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end; fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms); val proof_body_of = singleton proof_bodies_of; val proof_of = Proofterm.proof_of o proof_body_of; fun reconstruct_proof_of thm = Proofterm.reconstruct_proof (theory_of_thm thm) (prop_of thm) (proof_of thm); val consolidate = ignore o proof_bodies_of; fun expose_proofs thy thms = if Proofterm.export_proof_boxes_required thy then Proofterm.export_proof_boxes (proof_bodies_of (map (transfer thy) thms)) else (); fun expose_proof thy = expose_proofs thy o single; (* future rule *) fun future_result i orig_cert orig_shyps orig_prop thm = let fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]); val Thm (Deriv {promises, ...}, {cert, constraints, shyps, hyps, tpairs, prop, ...}) = thm; val _ = Context.eq_certificate (cert, orig_cert) orelse err "bad theory"; val _ = prop aconv orig_prop orelse err "bad prop"; val _ = null constraints orelse err "bad sort constraints"; val _ = null tpairs orelse err "bad flex-flex constraints"; val _ = null hyps orelse err "bad hyps"; val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps"; val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies"; val _ = join_promises promises; in thm end; fun future future_thm ct = let val Cterm {cert = cert, t = prop, T, maxidx, sorts} = ct; val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]); val _ = if Proofterm.proofs_enabled () then raise CTERM ("future: proof terms enabled", [ct]) else (); val i = serial (); val future = future_thm |> Future.map (future_result i cert sorts prop); in Thm (make_deriv [(i, future)] [] [] MinProof, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end; (** Axioms **) fun axiom thy name = (case Name_Space.lookup (Theory.axiom_table thy) name of SOME prop => let val der = deriv_rule0 (fn () => Proofterm.axm_proof name prop); val cert = Context.Certificate thy; val maxidx = maxidx_of_term prop; val shyps = Sorts.insert_term prop []; in Thm (der, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = shyps, hyps = [], tpairs = [], prop = prop}) end | NONE => raise THEORY ("No axiom " ^ quote name, [thy])); fun all_axioms_of thy = map (fn (name, _) => (name, axiom thy name)) (Theory.all_axioms_of thy); (* tags *) val get_tags = #tags o rep_thm; fun map_tags f (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = Thm (der, {cert = cert, tags = f tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (* technical adjustments *) fun norm_proof (th as Thm (der, args)) = Thm (deriv_rule1 (Proofterm.rew_proof (theory_of_thm th)) der, args); fun adjust_maxidx_thm i (th as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if maxidx = i then th else if maxidx < i then Thm (der, {maxidx = i, cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) else Thm (der, {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i), cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (*** Theory data ***) (* type classes *) structure Aritytab = Table( type key = string * sort list * class; val ord = fast_string_ord o apply2 #1 ||| fast_string_ord o apply2 #3 ||| list_ord Term_Ord.sort_ord o apply2 #2; ); datatype classes = Classes of {classrels: thm Symreltab.table, arities: (thm * string * serial) Aritytab.table}; fun make_classes (classrels, arities) = Classes {classrels = classrels, arities = arities}; val empty_classes = make_classes (Symreltab.empty, Aritytab.empty); (*see Theory.at_begin hook for transitive closure of classrels and arity completion*) fun merge_classes (Classes {classrels = classrels1, arities = arities1}, Classes {classrels = classrels2, arities = arities2}) = let val classrels' = Symreltab.merge (K true) (classrels1, classrels2); val arities' = Aritytab.merge (K true) (arities1, arities2); in make_classes (classrels', arities') end; (* data *) structure Data = Theory_Data ( type T = unit Name_Space.table * (*oracles: authentic derivation names*) classes; (*type classes within the logic*) val empty : T = (Name_Space.empty_table Markup.oracleN, empty_classes); val extend = I; fun merge ((oracles1, sorts1), (oracles2, sorts2)) : T = (Name_Space.merge_tables (oracles1, oracles2), merge_classes (sorts1, sorts2)); ); val get_oracles = #1 o Data.get; val map_oracles = Data.map o apfst; val get_classes = (fn (_, Classes args) => args) o Data.get; val get_classrels = #classrels o get_classes; val get_arities = #arities o get_classes; fun map_classes f = (Data.map o apsnd) (fn Classes {classrels, arities} => make_classes (f (classrels, arities))); fun map_classrels f = map_classes (fn (classrels, arities) => (f classrels, arities)); fun map_arities f = map_classes (fn (classrels, arities) => (classrels, f arities)); (* type classes *) fun the_classrel thy (c1, c2) = (case Symreltab.lookup (get_classrels thy) (c1, c2) of SOME thm => transfer thy thm | NONE => error ("Unproven class relation " ^ Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2])); fun the_arity thy (a, Ss, c) = (case Aritytab.lookup (get_arities thy) (a, Ss, c) of SOME (thm, _, _) => transfer thy thm | NONE => error ("Unproven type arity " ^ Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c]))); val classrel_proof = proof_of oo the_classrel; val arity_proof = proof_of oo the_arity; (* solve sort constraints by pro-forma proof *) local fun union_digest (oracles1, thms1) (oracles2, thms2) = (Proofterm.unions_oracles [oracles1, oracles2], Proofterm.unions_thms [thms1, thms2]); fun thm_digest (Thm (Deriv {body = PBody {oracles, thms, ...}, ...}, _)) = (oracles, thms); fun constraint_digest ({theory = thy, typ, sort, ...}: constraint) = Sorts.of_sort_derivation (Sign.classes_of thy) {class_relation = fn _ => fn _ => fn (digest, c1) => fn c2 => if c1 = c2 then ([], []) else union_digest digest (thm_digest (the_classrel thy (c1, c2))), type_constructor = fn (a, _) => fn dom => fn c => let val arity_digest = thm_digest (the_arity thy (a, (map o map) #2 dom, c)) in (fold o fold) (union_digest o #1) dom arity_digest end, type_variable = fn T => map (pair ([], [])) (Type.sort_of_atyp T)} (typ, sort); in fun solve_constraints (thm as Thm (_, {constraints = [], ...})) = thm | solve_constraints (thm as Thm (der, args)) = let val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args; val thy = Context.certificate_theory cert; val bad_thys = constraints |> map_filter (fn {theory = thy', ...} => if Context.eq_thy (thy, thy') then NONE else SOME thy'); val () = if null bad_thys then () else raise THEORY ("solve_constraints: bad theories for theorem\n" ^ Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys); val Deriv {promises, body = PBody {oracles, thms, proof}} = der; val (oracles', thms') = (oracles, thms) |> fold (fold union_digest o constraint_digest) constraints; val body' = PBody {oracles = oracles', thms = thms', proof = proof}; in Thm (Deriv {promises = promises, body = body'}, {constraints = [], cert = cert, tags = tags, maxidx = maxidx, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; end; (*Dangling sort constraints of a thm*) fun extra_shyps (th as Thm (_, {shyps, ...})) = Sorts.subtract (fold_terms Sorts.insert_term th []) shyps; (*Remove extra sorts that are witnessed by type signature information*) fun strip_shyps thm = (case thm of Thm (_, {shyps = [], ...}) => thm | Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) => let val thy = theory_of_thm thm; val algebra = Sign.classes_of thy; val minimize = Sorts.minimize_sort algebra; val le = Sorts.sort_le algebra; fun lt (S1, S2) = le (S1, S2) andalso not (le (S2, S1)); fun rel (S1, S2) = if S1 = S2 then [] else [(Term.aT S1, S2)]; val present = (fold_terms o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm []; val extra = fold (Sorts.remove_sort o #2) present shyps; val witnessed = Sign.witness_sorts thy present extra; val non_witnessed = fold (Sorts.remove_sort o #2) witnessed extra |> map (`minimize); val extra' = non_witnessed |> map_filter (fn (S, _) => if non_witnessed |> exists (fn (S', _) => lt (S', S)) then NONE else SOME S) |> Sorts.make; val constrs' = non_witnessed |> maps (fn (S1, S2) => let val S0 = the (find_first (fn S => le (S, S1)) extra') in rel (S0, S1) @ rel (S1, S2) end); val constraints' = fold (insert_constraints thy) (witnessed @ constrs') constraints; val shyps' = fold (Sorts.insert_sort o #2) present extra'; in Thm (deriv_rule_unconditional (Proofterm.strip_shyps_proof algebra present witnessed extra') der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop}) end) |> solve_constraints; (*** Closed theorems with official name ***) (*non-deterministic, depends on unknown promises*) fun derivation_closed (Thm (Deriv {body, ...}, _)) = Proofterm.compact_proof (Proofterm.proof_of body); (*non-deterministic, depends on unknown promises*) fun raw_derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (Proofterm.proof_of body); fun expand_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = let val self_id = (case Proofterm.get_identity shyps hyps prop (Proofterm.proof_of body) of NONE => K false | SOME {serial, ...} => fn (header: Proofterm.thm_header) => serial = #serial header); fun expand header = if self_id header orelse #name header = "" then SOME "" else NONE; in expand end; (*deterministic name of finished proof*) fun derivation_name (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (proof_of thm); (*identified PThm node*) fun derivation_id (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_id shyps hyps prop (proof_of thm); (*dependencies of PThm node*) fun thm_deps (thm as Thm (Deriv {promises = [], body = PBody {thms, ...}, ...}, _)) = (case (derivation_id thm, thms) of (SOME {serial = i, ...}, [(j, thm_node)]) => if i = j then Proofterm.thm_node_thms thm_node else thms | _ => thms) | thm_deps thm = raise THM ("thm_deps: bad promises", 0, [thm]); fun name_derivation name_pos = strip_shyps #> (fn thm as Thm (der, args) => let val thy = theory_of_thm thm; val Deriv {promises, body} = der; val {shyps, hyps, prop, tpairs, ...} = args; val _ = null tpairs orelse raise THM ("name_derivation: bad flex-flex constraints", 0, [thm]); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.thm_proof thy (classrel_proof thy) (arity_proof thy) name_pos shyps hyps prop ps body; val der' = make_deriv [] [] [pthm] proof; in Thm (der', args) end); fun close_derivation pos = solve_constraints #> (fn thm => if not (null (tpairs_of thm)) orelse derivation_closed thm then thm else name_derivation ("", pos) thm); val trim_context = solve_constraints #> trim_context_thm; (*** Oracles ***) fun add_oracle (b, oracle_fn) thy = let val (name, oracles') = Name_Space.define (Context.Theory thy) true (b, ()) (get_oracles thy); val thy' = map_oracles (K oracles') thy; fun invoke_oracle arg = let val Cterm {cert = cert2, t = prop, T, maxidx, sorts} = oracle_fn arg in if T <> propT then raise THM ("Oracle's result must have type prop: " ^ name, 0, []) else let val (oracle, prf) = (case ! Proofterm.proofs of 2 => (((name, Position.thread_data ()), SOME prop), Proofterm.oracle_proof name prop) | 1 => (((name, Position.thread_data ()), SOME prop), MinProof) | 0 => (((name, Position.none), NONE), MinProof) | i => bad_proofs i); in Thm (make_deriv [] [oracle] [] prf, {cert = Context.join_certificate (Context.Certificate thy', cert2), tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end end; in ((name, invoke_oracle), thy') end; val oracle_space = Name_Space.space_of_table o get_oracles; fun pretty_oracle ctxt = Name_Space.pretty ctxt (oracle_space (Proof_Context.theory_of ctxt)); fun extern_oracles verbose ctxt = map #1 (Name_Space.markup_table verbose ctxt (get_oracles (Proof_Context.theory_of ctxt))); fun check_oracle ctxt = Name_Space.check (Context.Proof ctxt) (get_oracles (Proof_Context.theory_of ctxt)) #> #1; (*** Meta rules ***) (** primitive rules **) (*The assumption rule A |- A*) fun assume raw_ct = let val Cterm {cert, t = prop, T, maxidx, sorts} = adjust_maxidx_cterm ~1 raw_ct in if T <> propT then raise THM ("assume: prop", 0, []) else if maxidx <> ~1 then raise THM ("assume: variables", maxidx, []) else Thm (deriv_rule0 (fn () => Proofterm.Hyp prop), {cert = cert, tags = [], maxidx = ~1, constraints = [], shyps = sorts, hyps = [prop], tpairs = [], prop = prop}) end; (*Implication introduction [A] : B ------- A \ B *) fun implies_intr (ct as Cterm {t = A, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...})) = if T <> propT then raise THM ("implies_intr: assumptions must have type prop", 0, [th]) else Thm (deriv_rule1 (Proofterm.implies_intr_proof A) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = remove_hyps A hyps, tpairs = tpairs, prop = Logic.mk_implies (A, prop)}); (*Implication elimination A \ B A ------------ B *) fun implies_elim thAB thA = let val Thm (derA, {maxidx = maxidx1, hyps = hypsA, constraints = constraintsA, shyps = shypsA, tpairs = tpairsA, prop = propA, ...}) = thA and Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...}) = thAB; fun err () = raise THM ("implies_elim: major premise", 0, [thAB, thA]); in (case prop of Const ("Pure.imp", _) $ A $ B => if A aconv propA then Thm (deriv_rule2 (curry Proofterm.%%) der derA, {cert = join_certificate2 (thAB, thA), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraintsA constraints, shyps = Sorts.union shypsA shyps, hyps = union_hyps hypsA hyps, tpairs = union_tpairs tpairsA tpairs, prop = B}) else err () | _ => err ()) end; (*Forall introduction. The Free or Var x must not be free in the hypotheses. [x] : A ------ \x. A *) fun forall_intr (ct as Cterm {maxidx = maxidx1, t = x, T, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = let fun result a = Thm (deriv_rule1 (Proofterm.forall_intr_proof (a, x) NONE) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.all_const T $ Abs (a, T, abstract_over (x, prop))}); fun check_occs a x ts = if exists (fn t => Logic.occs (x, t)) ts then raise THM ("forall_intr: variable " ^ quote a ^ " free in assumptions", 0, [th]) else (); in (case x of Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result a) | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result a) | _ => raise THM ("forall_intr: not a variable", 0, [th])) end; (*Forall elimination \x. A ------ A[t/x] *) fun forall_elim (ct as Cterm {t, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of Const ("Pure.all", Type ("fun", [Type ("fun", [qary, _]), _])) $ A => if T <> qary then raise THM ("forall_elim: type mismatch", 0, [th]) else Thm (deriv_rule1 (Proofterm.% o rpair (SOME t)) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Term.betapply (A, t)}) | _ => raise THM ("forall_elim: not quantified", 0, [th])); (* Equality *) (*Reflexivity t \ t *) fun reflexive (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t)}); (*Symmetry t \ u ------ u \ t *) fun symmetric (th as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of (eq as Const ("Pure.eq", _)) $ t $ u => Thm (deriv_rule1 Proofterm.symmetric_proof der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = eq $ u $ t}) | _ => raise THM ("symmetric", 0, [th])); (*Transitivity t1 \ u u \ t2 ------------------ t1 \ t2 *) fun transitive th1 th2 = let val Thm (der1, {maxidx = maxidx1, hyps = hyps1, constraints = constraints1, shyps = shyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, hyps = hyps2, constraints = constraints2, shyps = shyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("transitive: " ^ msg, 0, [th1, th2]); in case (prop1, prop2) of ((eq as Const ("Pure.eq", Type (_, [U, _]))) $ t1 $ u, Const ("Pure.eq", _) $ u' $ t2) => if not (u aconv u') then err "middle term" else Thm (deriv_rule2 (Proofterm.transitive_proof U u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = eq $ t1 $ t2}) | _ => err "premises" end; (*Beta-conversion (\x. t) u \ t[u/x] fully beta-reduces the term if full = true *) fun beta_conversion full (Cterm {cert, t, T = _, maxidx, sorts}) = let val t' = if full then Envir.beta_norm t else (case t of Abs (_, _, bodt) $ u => subst_bound (u, bodt) | _ => raise THM ("beta_conversion: not a redex", 0, [])); in Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t')}) end; fun eta_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_contract t)}); fun eta_long_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_long [] t)}); (*The abstraction rule. The Free or Var x must not be free in the hypotheses. The bound variable will be named "a" (since x will be something like x320) t \ u -------------- \x. t \ \x. u *) fun abstract_rule a (Cterm {t = x, T, sorts, ...}) (th as Thm (der, {cert, maxidx, hyps, constraints, shyps, tpairs, prop, ...})) = let val (t, u) = Logic.dest_equals prop handle TERM _ => raise THM ("abstract_rule: premise not an equality", 0, [th]); val result = Thm (deriv_rule1 (Proofterm.abstract_rule_proof (a, x)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.mk_equals (Abs (a, T, abstract_over (x, t)), Abs (a, T, abstract_over (x, u)))}); fun check_occs a x ts = if exists (fn t => Logic.occs (x, t)) ts then raise THM ("abstract_rule: variable " ^ quote a ^ " free in assumptions", 0, [th]) else (); in (case x of Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result) | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result) | _ => raise THM ("abstract_rule: not a variable", 0, [th])) end; (*The combination rule f \ g t \ u ------------- f t \ g u *) fun combination th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun chktypes fT tT = (case fT of Type ("fun", [T1, _]) => if T1 <> tT then raise THM ("combination: types", 0, [th1, th2]) else () | _ => raise THM ("combination: not function type", 0, [th1, th2])); in (case (prop1, prop2) of (Const ("Pure.eq", Type ("fun", [fT, _])) $ f $ g, Const ("Pure.eq", Type ("fun", [tT, _])) $ t $ u) => (chktypes fT tT; Thm (deriv_rule2 (Proofterm.combination_proof f g t u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (f $ t, g $ u)})) | _ => raise THM ("combination: premises", 0, [th1, th2])) end; (*Equality introduction A \ B B \ A ---------------- A \ B *) fun equal_intr th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_intr: " ^ msg, 0, [th1, th2]); in (case (prop1, prop2) of (Const("Pure.imp", _) $ A $ B, Const("Pure.imp", _) $ B' $ A') => if A aconv A' andalso B aconv B' then Thm (deriv_rule2 (Proofterm.equal_intr_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (A, B)}) else err "not equal" | _ => err "premises") end; (*The equal propositions rule A \ B A --------- B *) fun equal_elim th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_elim: " ^ msg, 0, [th1, th2]); in (case prop1 of Const ("Pure.eq", _) $ A $ B => if prop2 aconv A then Thm (deriv_rule2 (Proofterm.equal_elim_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = B}) else err "not equal" | _ => err "major premise") end; (**** Derived rules ****) (*Smash unifies the list of term pairs leaving no flex-flex pairs. Instantiates the theorem and deletes trivial tpairs. Resulting sequence may contain multiple elements if the tpairs are not all flex-flex.*) fun flexflex_rule opt_ctxt = solve_constraints #> (fn th => let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val (context, cert') = make_context_certificate [th] opt_ctxt cert; in Unify.smash_unifiers context tpairs (Envir.empty maxidx) |> Seq.map (fn env => if Envir.is_empty env then th else let val tpairs' = tpairs |> map (apply2 (Envir.norm_term env)) (*remove trivial tpairs, of the form t \ t*) |> filter_out (op aconv); val der' = deriv_rule1 (Proofterm.norm_proof' env) der; val constraints' = insert_constraints_env (Context.certificate_theory cert') env constraints; val prop' = Envir.norm_term env prop; val maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop'); val shyps = Envir.insert_sorts env shyps; in Thm (der', {cert = cert', tags = [], maxidx = maxidx, constraints = constraints', shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end) end); (*Generalization of fixed variables A -------------------- A[?'a/'a, ?x/x, ...] *) fun generalize (tfrees, frees) idx th = if Symtab.is_empty tfrees andalso Symtab.is_empty frees then th else let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]); val bad_type = if Symtab.is_empty tfrees then K false else Term.exists_subtype (fn TFree (a, _) => Symtab.defined tfrees a | _ => false); fun bad_term (Free (x, T)) = bad_type T orelse Symtab.defined frees x | bad_term (Var (_, T)) = bad_type T | bad_term (Const (_, T)) = bad_type T | bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t | bad_term (t $ u) = bad_term t orelse bad_term u | bad_term (Bound _) = false; val _ = exists bad_term hyps andalso raise THM ("generalize: variable free in assumptions", 0, [th]); val generalize = Term_Subst.generalize (tfrees, frees) idx; val prop' = generalize prop; val tpairs' = map (apply2 generalize) tpairs; val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop'); in Thm (deriv_rule1 (Proofterm.generalize_proof (tfrees, frees) idx prop) der, {cert = cert, tags = [], maxidx = maxidx', constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end; fun generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) = if Symtab.is_empty tfrees andalso Symtab.is_empty frees then ct else if idx <= maxidx then raise CTERM ("generalize_cterm: bad index", [ct]) else Cterm {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, t = Term_Subst.generalize (tfrees, frees) idx t, maxidx = Int.max (maxidx, idx)}; fun generalize_ctyp tfrees idx (cT as Ctyp {cert, T, maxidx, sorts}) = if Symtab.is_empty tfrees then cT else if idx <= maxidx then raise CTERM ("generalize_ctyp: bad index", []) else Ctyp {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, maxidx = Int.max (maxidx, idx)}; (*Instantiation of schematic variables A -------------------- A[t1/v1, ..., tn/vn] *) local fun add_cert cert_of (_, c) cert = Context.join_certificate (cert, cert_of c); val add_instT_cert = add_cert (fn Ctyp {cert, ...} => cert); val add_inst_cert = add_cert (fn Cterm {cert, ...} => cert); fun add_sorts sorts_of (_, c) sorts = Sorts.union (sorts_of c) sorts; val add_instT_sorts = add_sorts (fn Ctyp {sorts, ...} => sorts); val add_inst_sorts = add_sorts (fn Cterm {sorts, ...} => sorts); fun add_instT thy (v as (_, S), Ctyp {T = U, maxidx, ...}) = if Sign.of_sort thy (U, S) then Term_Subst.TVars.add (v, (U, maxidx)) else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy S, [U], []); fun add_inst thy (v as (_, T), Cterm {t = u, T = U, maxidx, ...}) = if T = U then Term_Subst.Vars.add (v, (u, maxidx)) else let fun pretty_typing t ty = Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy ty]; val msg = Pretty.string_of (Pretty.block [Pretty.str "instantiate: type conflict", Pretty.fbrk, pretty_typing (Var v) T, Pretty.fbrk, pretty_typing u U]) in raise TYPE (msg, [T, U], [Var v, u]) end; fun prep_insts (instT, inst) (cert, sorts) = let val cert' = cert |> fold add_instT_cert instT |> fold add_inst_cert inst; val thy' = Context.certificate_theory cert' handle ERROR msg => raise CONTEXT (msg, map #2 instT, map #2 inst, [], NONE); val sorts' = sorts |> fold add_instT_sorts instT |> fold add_inst_sorts inst; val instT' = Term_Subst.TVars.build (fold (add_instT thy') instT); val inst' = Term_Subst.Vars.build (fold (add_inst thy') inst); in ((instT', inst'), (cert', sorts')) end; in (*Left-to-right replacements: ctpairs = [..., (vi, ti), ...]. Instantiates distinct Vars by terms of same type. Does NOT normalize the resulting theorem!*) fun instantiate ([], []) th = th | instantiate (instT, inst) th = let val Thm (der, {cert, hyps, constraints, shyps, tpairs, prop, ...}) = th; val ((instT', inst'), (cert', shyps')) = prep_insts (instT, inst) (cert, shyps) handle CONTEXT (msg, cTs, cts, ths, context) => raise CONTEXT (msg, cTs, cts, th :: ths, context); val subst = Term_Subst.instantiate_maxidx (instT', inst'); val (prop', maxidx1) = subst prop ~1; val (tpairs', maxidx') = fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1; val thy' = Context.certificate_theory cert'; val constraints' = Term_Subst.TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) instT' constraints; in Thm (deriv_rule1 (fn d => Proofterm.instantiate (Term_Subst.TVars.map (K #1) instT', Term_Subst.Vars.map (K #1) inst') d) der, {cert = cert', tags = [], maxidx = maxidx', constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs', prop = prop'}) |> solve_constraints end handle TYPE (msg, _, _) => raise THM (msg, 0, [th]); fun instantiate_cterm ([], []) ct = ct | instantiate_cterm (instT, inst) ct = let val Cterm {cert, t, T, sorts, ...} = ct; val ((instT', inst'), (cert', sorts')) = prep_insts (instT, inst) (cert, sorts); val subst = Term_Subst.instantiate_maxidx (instT', inst'); val substT = Term_Subst.instantiateT_maxidx instT'; val (t', maxidx1) = subst t ~1; val (T', maxidx') = substT T maxidx1; in Cterm {cert = cert', t = t', T = T', sorts = sorts', maxidx = maxidx'} end handle TYPE (msg, _, _) => raise CTERM (msg, [ct]); end; (*The trivial implication A \ A, justified by assume and forall rules. A can contain Vars, not so for assume!*) fun trivial (Cterm {cert, t = A, T, maxidx, sorts}) = if T <> propT then raise THM ("trivial: the term must have type prop", 0, []) else Thm (deriv_rule0 (fn () => Proofterm.trivial_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_implies (A, A)}); (*Axiom-scheme reflecting signature contents T :: c ------------------- OFCLASS(T, c_class) *) fun of_class (cT, raw_c) = let val Ctyp {cert, T, ...} = cT; val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [cT], [], [], NONE); val c = Sign.certify_class thy raw_c; val Cterm {t = prop, maxidx, sorts, ...} = global_cterm_of thy (Logic.mk_of_class (T, c)); in if Sign.of_sort thy (T, [c]) then Thm (deriv_rule0 (fn () => Proofterm.PClass (T, c)), {cert = cert, tags = [], maxidx = maxidx, constraints = insert_constraints thy (T, [c]) [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, []) end |> solve_constraints; (*Internalize sort constraints of type variables*) val unconstrainT = strip_shyps #> (fn thm as Thm (der, args) => let val Deriv {promises, body} = der; val {cert, shyps, hyps, tpairs, prop, ...} = args; val thy = theory_of_thm thm; fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]); val _ = null hyps orelse err "bad hyps"; val _ = null tpairs orelse err "bad flex-flex constraints"; val tfrees = rev (Term.add_tfree_names prop []); val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.unconstrain_thm_proof thy (classrel_proof thy) (arity_proof thy) shyps prop ps body; val der' = make_deriv [] [] [pthm] proof; val prop' = Proofterm.thm_node_prop (#2 pthm); in Thm (der', {cert = cert, tags = [], maxidx = maxidx_of_term prop', constraints = [], shyps = [[]], (*potentially redundant*) hyps = [], tpairs = [], prop = prop'}) end); (*Replace all TFrees not fixed or in the hyps by new TVars*) fun varifyT_global' fixed (Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = let val tfrees = fold Term_Subst.add_tfrees hyps fixed; val prop1 = attach_tpairs tpairs prop; val (al, prop2) = Type.varify_global tfrees prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in (al, Thm (deriv_rule1 (Proofterm.varify_proof prop tfrees) der, {cert = cert, tags = [], maxidx = Int.max (0, maxidx), constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3})) end; val varifyT_global = #2 o varifyT_global' Term_Subst.TFrees.empty; (*Replace all TVars by TFrees that are often new*) fun legacy_freezeT (Thm (der, {cert, constraints, shyps, hyps, tpairs, prop, ...})) = let val prop1 = attach_tpairs tpairs prop; val prop2 = Type.legacy_freeze prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in Thm (deriv_rule1 (Proofterm.legacy_freezeT prop1) der, {cert = cert, tags = [], maxidx = maxidx_of_term prop2, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3}) end; fun plain_prop_of raw_thm = let val thm = strip_shyps raw_thm; fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]); in if not (null (hyps_of thm)) then err "theorem may not contain hypotheses" else if not (null (extra_shyps thm)) then err "theorem may not contain sort hypotheses" else if not (null (tpairs_of thm)) then err "theorem may not contain flex-flex pairs" else prop_of thm end; (*** Inference rules for tactics ***) (*Destruct proof state into constraints, other goals, goal(i), rest *) fun dest_state (state as Thm (_, {prop,tpairs,...}), i) = (case Logic.strip_prems(i, [], prop) of (B::rBs, C) => (tpairs, rev rBs, B, C) | _ => raise THM("dest_state", i, [state])) handle TERM _ => raise THM("dest_state", i, [state]); (*Prepare orule for resolution by lifting it over the parameters and assumptions of goal.*) fun lift_rule goal orule = let val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal; val inc = gmax + 1; val lift_abs = Logic.lift_abs inc gprop; val lift_all = Logic.lift_all inc gprop; val Thm (der, {maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = orule; val (As, B) = Logic.strip_horn prop; in if T <> propT then raise THM ("lift_rule: the term must have type prop", 0, []) else Thm (deriv_rule1 (Proofterm.lift_proof gprop inc prop) der, {cert = join_certificate1 (goal, orule), tags = [], maxidx = maxidx + inc, constraints = constraints, shyps = Sorts.union shyps sorts, (*sic!*) hyps = hyps, tpairs = map (apply2 lift_abs) tpairs, prop = Logic.list_implies (map lift_all As, lift_all B)}) end; fun incr_indexes i (thm as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = if i < 0 then raise THM ("negative increment", 0, [thm]) else if i = 0 then thm else Thm (deriv_rule1 (Proofterm.incr_indexes i) der, {cert = cert, tags = [], maxidx = maxidx + i, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = map (apply2 (Logic.incr_indexes ([], [], i))) tpairs, prop = Logic.incr_indexes ([], [], i) prop}); (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) fun assumption opt_ctxt i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (context, cert') = make_context_certificate [state] opt_ctxt cert; val (tpairs, Bs, Bi, C) = dest_state (state, i); fun newth n (env, tpairs) = let val normt = Envir.norm_term env; fun assumption_proof prf = Proofterm.assumption_proof (map normt Bs) (normt Bi) n prf; in Thm (deriv_rule1 (assumption_proof #> not (Envir.is_empty env) ? Proofterm.norm_proof' env) der, {tags = [], maxidx = Envir.maxidx_of env, constraints = insert_constraints_env (Context.certificate_theory cert') env constraints, shyps = Envir.insert_sorts env shyps, hyps = hyps, tpairs = if Envir.is_empty env then tpairs else map (apply2 normt) tpairs, prop = if Envir.is_empty env then Logic.list_implies (Bs, C) (*avoid wasted normalizations*) else normt (Logic.list_implies (Bs, C)) (*normalize the new rule fully*), cert = cert'}) end; val (close, asms, concl) = Logic.assum_problems (~1, Bi); val concl' = close concl; fun addprfs [] _ = Seq.empty | addprfs (asm :: rest) n = Seq.make (fn () => Seq.pull (Seq.mapp (newth n) (if Term.could_unify (asm, concl) then (Unify.unifiers (context, Envir.empty maxidx, (close asm, concl') :: tpairs)) else Seq.empty) (addprfs rest (n + 1)))) in addprfs asms 1 end; (*Solve subgoal Bi of proof state B1...Bn/C by assumption. Checks if Bi's conclusion is alpha/eta-convertible to one of its assumptions*) fun eq_assumption i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val (_, asms, concl) = Logic.assum_problems (~1, Bi); in (case find_index (fn asm => Envir.aeconv (asm, concl)) asms of ~1 => raise THM ("eq_assumption", 0, [state]) | n => Thm (deriv_rule1 (Proofterm.assumption_proof Bs Bi (n + 1)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs, C)})) end; (*For rotate_tac: fast rotation of assumptions of subgoal i*) fun rotate_rule k i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val params = Term.strip_all_vars Bi; val rest = Term.strip_all_body Bi; val asms = Logic.strip_imp_prems rest val concl = Logic.strip_imp_concl rest; val n = length asms; val m = if k < 0 then n + k else k; val Bi' = if 0 = m orelse m = n then Bi else if 0 < m andalso m < n then let val (ps, qs) = chop m asms in Logic.list_all (params, Logic.list_implies (qs @ ps, concl)) end else raise THM ("rotate_rule", k, [state]); in Thm (deriv_rule1 (Proofterm.rotate_proof Bs Bi' params asms m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs @ [Bi'], C)}) end; (*Rotates a rule's premises to the left by k, leaving the first j premises unchanged. Does nothing if k=0 or if k equals n-j, where n is the number of premises. Useful with eresolve_tac and underlies defer_tac*) fun permute_prems j k rl = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = rl; val prems = Logic.strip_imp_prems prop and concl = Logic.strip_imp_concl prop; val moved_prems = List.drop (prems, j) and fixed_prems = List.take (prems, j) handle General.Subscript => raise THM ("permute_prems: j", j, [rl]); val n_j = length moved_prems; val m = if k < 0 then n_j + k else k; val (prems', prop') = if 0 = m orelse m = n_j then (prems, prop) else if 0 < m andalso m < n_j then let val (ps, qs) = chop m moved_prems; val prems' = fixed_prems @ qs @ ps; in (prems', Logic.list_implies (prems', concl)) end else raise THM ("permute_prems: k", k, [rl]); in Thm (deriv_rule1 (Proofterm.permute_prems_proof prems' j m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) end; (* strip_apply f B A strips off all assumptions/parameters from A introduced by lifting over B, and applies f to remaining part of A*) fun strip_apply f = let fun strip (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ A2 $ B2) = Logic.mk_implies (A2, strip B1 B2) | strip ((c as Const ("Pure.all", _)) $ Abs (_, _, t1)) ( Const ("Pure.all", _) $ Abs (a, T, t2)) = c $ Abs (a, T, strip t1 t2) | strip _ A = f A in strip end; fun strip_lifted (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ _ $ B2) = strip_lifted B1 B2 | strip_lifted (Const ("Pure.all", _) $ Abs (_, _, t1)) (Const ("Pure.all", _) $ Abs (_, _, t2)) = strip_lifted t1 t2 | strip_lifted _ A = A; (*Use the alist to rename all bound variables and some unknowns in a term dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); Preserves unknowns in tpairs and on lhs of dpairs. *) fun rename_bvs [] _ _ _ _ = K I | rename_bvs al dpairs tpairs B As = let val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I); val vids = [] |> fold (add_var o fst) dpairs |> fold (add_var o fst) tpairs |> fold (add_var o snd) tpairs; val vids' = fold (add_var o strip_lifted B) As []; (*unknowns appearing elsewhere be preserved!*) val al' = distinct ((op =) o apply2 fst) (filter_out (fn (x, y) => not (member (op =) vids' x) orelse member (op =) vids x orelse member (op =) vids y) al); val unchanged = filter_out (AList.defined (op =) al') vids'; fun del_clashing clash xs _ [] qs = if clash then del_clashing false xs xs qs [] else qs | del_clashing clash xs ys ((p as (x, y)) :: ps) qs = if member (op =) ys y then del_clashing true (x :: xs) (x :: ys) ps qs else del_clashing clash xs (y :: ys) ps (p :: qs); val al'' = del_clashing false unchanged unchanged al' []; fun rename (t as Var ((x, i), T)) = (case AList.lookup (op =) al'' x of SOME y => Var ((y, i), T) | NONE => t) | rename (Abs (x, T, t)) = Abs (the_default x (AList.lookup (op =) al x), T, rename t) | rename (f $ t) = rename f $ rename t | rename t = t; fun strip_ren f Ai = f rename B Ai in strip_ren end; (*Function to rename bounds/unknowns in the argument, lifted over B*) fun rename_bvars dpairs = rename_bvs (fold_rev Term.match_bvars dpairs []) dpairs; (*** RESOLUTION ***) (** Lifting optimizations **) (*strip off pairs of assumptions/parameters in parallel -- they are identical because of lifting*) fun strip_assums2 (Const("Pure.imp", _) $ _ $ B1, Const("Pure.imp", _) $ _ $ B2) = strip_assums2 (B1,B2) | strip_assums2 (Const("Pure.all",_)$Abs(a,T,t1), Const("Pure.all",_)$Abs(_,_,t2)) = let val (B1,B2) = strip_assums2 (t1,t2) in (Abs(a,T,B1), Abs(a,T,B2)) end | strip_assums2 BB = BB; (*Faster normalization: skip assumptions that were lifted over*) fun norm_term_skip env 0 t = Envir.norm_term env t | norm_term_skip env n (Const ("Pure.all", _) $ Abs (a, T, t)) = let val T' = Envir.norm_type (Envir.type_env env) T (*Must instantiate types of parameters because they are flattened; this could be a NEW parameter*) in Logic.all_const T' $ Abs (a, T', norm_term_skip env n t) end | norm_term_skip env n (Const ("Pure.imp", _) $ A $ B) = Logic.mk_implies (A, norm_term_skip env (n - 1) B) | norm_term_skip _ _ _ = error "norm_term_skip: too few assumptions??"; (*unify types of schematic variables (non-lifted case)*) fun unify_var_types context (th1, th2) env = let fun unify_vars (T :: Us) = fold (fn U => Pattern.unify_types context (T, U)) Us | unify_vars _ = I; val add_vars = full_prop_of #> fold_aterms (fn Var v => Vartab.insert_list (op =) v | _ => I); val vars = Vartab.build (add_vars th1 #> add_vars th2); in SOME (Vartab.fold (unify_vars o #2) vars env) end handle Pattern.Unif => NONE; (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) Unifies B with Bi, replacing subgoal i (1 <= i <= n) If match then forbid instantiations in proof state If lifted then shorten the dpair using strip_assums2. If eres_flg then simultaneously proves A1 by assumption. nsubgoal is the number of new subgoals (written m above). Curried so that resolution calls dest_state only once. *) local exception COMPOSE in fun bicompose_aux opt_ctxt {flatten, match, incremented} (state, (stpairs, Bs, Bi, C), lifted) (eres_flg, orule, nsubgoal) = let val Thm (sder, {maxidx=smax, constraints = constraints2, shyps = shyps2, hyps = hyps2, ...}) = state and Thm (rder, {maxidx=rmax, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs=rtpairs, prop=rprop,...}) = orule (*How many hyps to skip over during normalization*) and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0) val (context, cert) = make_context_certificate [state, orule] opt_ctxt (join_certificate2 (state, orule)); (*Add new theorem with prop = "\Bs; As\ \ C" to thq*) fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) = let val normt = Envir.norm_term env; (*perform minimal copying here by examining env*) val (ntpairs, normp) = if Envir.is_empty env then (tpairs, (Bs @ As, C)) else let val ntps = map (apply2 normt) tpairs in if Envir.above env smax then (*no assignments in state; normalize the rule only*) if lifted then (ntps, (Bs @ map (norm_term_skip env nlift) As, C)) else (ntps, (Bs @ map normt As, C)) else if match then raise COMPOSE else (*normalize the new rule fully*) (ntps, (map normt (Bs @ As), normt C)) end val constraints' = union_constraints constraints1 constraints2 |> insert_constraints_env (Context.certificate_theory cert) env; fun bicompose_proof prf1 prf2 = Proofterm.bicompose_proof flatten (map normt Bs) (map normt As) A oldAs n (nlift+1) prf1 prf2 val th = Thm (deriv_rule2 (if Envir.is_empty env then bicompose_proof else if Envir.above env smax then bicompose_proof o Proofterm.norm_proof' env else Proofterm.norm_proof' env oo bicompose_proof) rder' sder, {tags = [], maxidx = Envir.maxidx_of env, constraints = constraints', shyps = Envir.insert_sorts env (Sorts.union shyps1 shyps2), hyps = union_hyps hyps1 hyps2, tpairs = ntpairs, prop = Logic.list_implies normp, cert = cert}) in Seq.cons th thq end handle COMPOSE => thq; val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop) handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]); (*Modify assumptions, deleting n-th if n>0 for e-resolution*) fun newAs(As0, n, dpairs, tpairs) = let val (As1, rder') = if not lifted then (As0, rder) else let val rename = rename_bvars dpairs tpairs B As0 in (map (rename strip_apply) As0, deriv_rule1 (Proofterm.map_proof_terms (rename K) I) rder) end; in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n) handle TERM _ => raise THM("bicompose: 1st premise", 0, [orule]) end; val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); val dpairs = BBi :: (rtpairs@stpairs); (*elim-resolution: try each assumption in turn*) fun eres _ [] = raise THM ("bicompose: no premises", 0, [orule, state]) | eres env (A1 :: As) = let val A = SOME A1; val (close, asms, concl) = Logic.assum_problems (nlift + 1, A1); val concl' = close concl; fun tryasms [] _ = Seq.empty | tryasms (asm :: rest) n = if Term.could_unify (asm, concl) then let val asm' = close asm in (case Seq.pull (Unify.unifiers (context, env, (asm', concl') :: dpairs)) of NONE => tryasms rest (n + 1) | cell as SOME ((_, tpairs), _) => Seq.it_right (addth A (newAs (As, n, [BBi, (concl', asm')], tpairs))) (Seq.make (fn () => cell), Seq.make (fn () => Seq.pull (tryasms rest (n + 1))))) end else tryasms rest (n + 1); in tryasms asms 1 end; (*ordinary resolution*) fun res env = (case Seq.pull (Unify.unifiers (context, env, dpairs)) of NONE => Seq.empty | cell as SOME ((_, tpairs), _) => Seq.it_right (addth NONE (newAs (rev rAs, 0, [BBi], tpairs))) (Seq.make (fn () => cell), Seq.empty)); val env0 = Envir.empty (Int.max (rmax, smax)); in (case if incremented then SOME env0 else unify_var_types context (state, orule) env0 of NONE => Seq.empty | SOME env => if eres_flg then eres env (rev rAs) else res env) end; end; fun bicompose opt_ctxt flags arg i state = bicompose_aux opt_ctxt flags (state, dest_state (state,i), false) arg; (*Quick test whether rule is resolvable with the subgoal with hyps Hs and conclusion B. If eres_flg then checks 1st premise of rule also*) fun could_bires (Hs, B, eres_flg, rule) = let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs | could_reshyp [] = false; (*no premise -- illegal*) in Term.could_unify(concl_of rule, B) andalso (not eres_flg orelse could_reshyp (prems_of rule)) end; (*Bi-resolution of a state with a list of (flag,rule) pairs. Puts the rule above: rule/state. Renames vars in the rules. *) fun biresolution opt_ctxt match brules i state = let val (stpairs, Bs, Bi, C) = dest_state(state,i); val lift = lift_rule (cprem_of state i); val B = Logic.strip_assums_concl Bi; val Hs = Logic.strip_assums_hyp Bi; val compose = bicompose_aux opt_ctxt {flatten = true, match = match, incremented = true} (state, (stpairs, Bs, Bi, C), true); fun res [] = Seq.empty | res ((eres_flg, rule)::brules) = if Config.get_generic (make_context [state] opt_ctxt (cert_of state)) Pattern.unify_trace_failure orelse could_bires (Hs, B, eres_flg, rule) then Seq.make (*delay processing remainder till needed*) (fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule), res brules)) else res brules in Seq.flat (res brules) end; (*Resolution: exactly one resolvent must be produced*) fun tha RSN (i, thb) = (case Seq.chop 2 (biresolution NONE false [(false, tha)] i thb) of ([th], _) => solve_constraints th | ([], _) => raise THM ("RSN: no unifiers", i, [tha, thb]) | _ => raise THM ("RSN: multiple unifiers", i, [tha, thb])); (*Resolution: P \ Q, Q \ R gives P \ R*) fun tha RS thb = tha RSN (1,thb); (**** Type classes ****) fun standard_tvars thm = let val thy = theory_of_thm thm; val tvars = rev (Term.add_tvars (prop_of thm) []); val names = Name.invent Name.context Name.aT (length tvars); val tinst = map2 (fn (ai, S) => fn b => ((ai, S), global_ctyp_of thy (TVar ((b, 0), S)))) tvars names; in instantiate (tinst, []) thm end (* class relations *) val is_classrel = Symreltab.defined o get_classrels; fun complete_classrels thy = let fun complete (c, (_, (all_preds, all_succs))) (finished1, thy1) = let fun compl c1 c2 (finished2, thy2) = if is_classrel thy2 (c1, c2) then (finished2, thy2) else (false, thy2 |> (map_classrels o Symreltab.update) ((c1, c2), (the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy2) |> trim_context)); val proven = is_classrel thy1; val preds = Graph.Keys.fold (fn c1 => proven (c1, c) ? cons c1) all_preds []; val succs = Graph.Keys.fold (fn c2 => proven (c, c2) ? cons c2) all_succs []; in fold_product compl preds succs (finished1, thy1) end; in (case Graph.fold complete (Sorts.classes_of (Sign.classes_of thy)) (true, thy) of (true, _) => NONE | (_, thy') => SOME thy') end; (* type arities *) fun thynames_of_arity thy (a, c) = - (get_arities thy, []) |-> Aritytab.fold - (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser)) + build (get_arities thy |> Aritytab.fold + (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser))) |> sort (int_ord o apply2 #2) |> map #1; fun insert_arity_completions thy ((t, Ss, c), (th, thy_name, ser)) (finished, arities) = let val completions = Sign.super_classes thy c |> map_filter (fn c1 => if Aritytab.defined arities (t, Ss, c1) then NONE else let val th1 = (th RS the_classrel thy (c, c1)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy) |> trim_context; in SOME ((t, Ss, c1), (th1, thy_name, ser)) end); val finished' = finished andalso null completions; val arities' = fold Aritytab.update completions arities; in (finished', arities') end; fun complete_arities thy = let val arities = get_arities thy; val (finished, arities') = Aritytab.fold (insert_arity_completions thy) arities (true, get_arities thy); in if finished then NONE else SOME (map_arities (K arities') thy) end; val _ = Theory.setup (Theory.at_begin complete_classrels #> Theory.at_begin complete_arities); (* primitive rules *) fun add_classrel raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (c1, c2) = Logic.dest_classrel prop; in thy |> Sign.primitive_classrel (c1, c2) |> map_classrels (Symreltab.update ((c1, c2), th')) |> perhaps complete_classrels |> perhaps complete_arities end; fun add_arity raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (t, Ss, c) = Logic.dest_arity prop; val ar = ((t, Ss, c), (th', Context.theory_name thy, serial ())); in thy |> Sign.primitive_arity (t, Ss, [c]) |> map_arities (Aritytab.update ar #> curry (insert_arity_completions thy ar) true #> #2) end; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Pure/type.ML b/src/Pure/type.ML --- a/src/Pure/type.ML +++ b/src/Pure/type.ML @@ -1,714 +1,714 @@ (* Title: Pure/type.ML Author: Tobias Nipkow, Lawrence C Paulson, and Markus Wenzel Type signatures and certified types, special treatment of type vars, matching and unification of types, extend and merge type signatures. *) signature TYPE = sig (*constraints*) val mark_polymorphic: typ -> typ val constraint: typ -> term -> term val constraint_type: Proof.context -> typ -> typ val strip_constraints: term -> term val appl_error: Proof.context -> term -> typ -> term -> typ -> string (*type signatures and certified types*) datatype decl = LogicalType of int | Abbreviation of string list * typ * bool | Nonterminal type tsig val eq_tsig: tsig * tsig -> bool val rep_tsig: tsig -> {classes: Name_Space.T * Sorts.algebra, default: sort, types: decl Name_Space.table, log_types: string list} val change_base: bool -> tsig -> tsig val change_ignore: tsig -> tsig val empty_tsig: tsig val class_space: tsig -> Name_Space.T val defaultS: tsig -> sort val logical_types: tsig -> string list val eq_sort: tsig -> sort * sort -> bool val subsort: tsig -> sort * sort -> bool val of_sort: tsig -> typ * sort -> bool val inter_sort: tsig -> sort * sort -> sort val cert_class: tsig -> class -> class val cert_sort: tsig -> sort -> sort val minimize_sort: tsig -> sort -> sort val witness_sorts: tsig -> (typ * sort) list -> sort list -> (typ * sort) list type mode val mode_default: mode val mode_syntax: mode val mode_abbrev: mode val get_mode: Proof.context -> mode val set_mode: mode -> Proof.context -> Proof.context val restore_mode: Proof.context -> Proof.context -> Proof.context val type_space: tsig -> Name_Space.T val type_alias: Name_Space.naming -> binding -> string -> tsig -> tsig val check_decl: Context.generic -> tsig -> xstring * Position.T -> (string * Position.report list) * decl val the_decl: tsig -> string * Position.T -> decl val cert_typ_mode: mode -> tsig -> typ -> typ val cert_typ: tsig -> typ -> typ val arity_number: tsig -> string -> int val arity_sorts: Context.generic -> tsig -> string -> sort -> sort list (*special treatment of type vars*) val sort_of_atyp: typ -> sort val strip_sorts: typ -> typ val strip_sorts_dummy: typ -> typ val no_tvars: typ -> typ val varify_global: Term_Subst.TFrees.set -> term -> ((string * sort) * indexname) list * term val legacy_freeze_thaw_type: typ -> typ * (typ -> typ) val legacy_freeze_type: typ -> typ val legacy_freeze_thaw: term -> term * (term -> term) val legacy_freeze: term -> term (*matching and unification*) exception TYPE_MATCH type tyenv = (sort * typ) Vartab.table val lookup: tyenv -> indexname * sort -> typ option val devar: tyenv -> typ -> typ val typ_match: tsig -> typ * typ -> tyenv -> tyenv val typ_instance: tsig -> typ * typ -> bool val raw_match: typ * typ -> tyenv -> tyenv val raw_matches: typ list * typ list -> tyenv -> tyenv val could_match: typ * typ -> bool val could_matches: typ list * typ list -> bool val raw_instance: typ * typ -> bool exception TUNIFY val unify: tsig -> typ * typ -> tyenv * int -> tyenv * int val raw_unify: typ * typ -> tyenv -> tyenv val raw_unifys: typ list * typ list -> tyenv -> tyenv val could_unify: typ * typ -> bool val could_unifys: typ list * typ list -> bool val unified: tyenv -> typ * typ -> bool (*extend and merge type signatures*) val add_class: Context.generic -> binding * class list -> tsig -> tsig val hide_class: bool -> string -> tsig -> tsig val set_defsort: sort -> tsig -> tsig val add_type: Context.generic -> binding * int -> tsig -> tsig val add_abbrev: Context.generic -> binding * string list * typ -> tsig -> tsig val add_nonterminal: Context.generic -> binding -> tsig -> tsig val hide_type: bool -> string -> tsig -> tsig val add_arity: Context.generic -> arity -> tsig -> tsig val add_classrel: Context.generic -> class * class -> tsig -> tsig val merge_tsig: Context.generic -> tsig * tsig -> tsig end; structure Type: TYPE = struct (** constraints **) (*indicate polymorphic Vars*) fun mark_polymorphic T = Type ("_polymorphic_", [T]); fun constraint T t = if T = dummyT then t else Const ("_type_constraint_", T --> T) $ t; fun constraint_type ctxt T = let fun err () = error ("Malformed internal type constraint: " ^ Syntax.string_of_typ ctxt T); in (case T of Type ("fun", [A, B]) => if A = B then A else err () | _ => err ()) end; fun strip_constraints (Const ("_type_constraint_", _) $ t) = strip_constraints t | strip_constraints (t $ u) = strip_constraints t $ strip_constraints u | strip_constraints (Abs (x, T, t)) = Abs (x, T, strip_constraints t) | strip_constraints a = a; fun appl_error ctxt (Const ("_type_constraint_", Type ("fun", [T, _]))) _ u U = cat_lines ["Failed to meet type constraint:", "", Pretty.string_of (Pretty.block [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt u, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt U]), Pretty.string_of (Pretty.block [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt T])] | appl_error ctxt t T u U = cat_lines ["Type error in application: " ^ (case T of Type ("fun", _) => "incompatible operand type" | _ => "operator not of function type"), "", Pretty.string_of (Pretty.block [Pretty.str "Operator:", Pretty.brk 2, Syntax.pretty_term ctxt t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt T]), Pretty.string_of (Pretty.block [Pretty.str "Operand:", Pretty.brk 3, Syntax.pretty_term ctxt u, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt U])]; (** type signatures and certified types **) (* type declarations *) datatype decl = LogicalType of int | Abbreviation of string list * typ * bool | Nonterminal; (* type tsig *) datatype tsig = TSig of { classes: Name_Space.T * Sorts.algebra, (*order-sorted algebra of type classes*) default: sort, (*default sort on input*) types: decl Name_Space.table, (*declared types*) log_types: string list}; (*logical types sorted by number of arguments*) fun eq_tsig (TSig {classes = classes1, default = default1, types = types1, log_types = _}, TSig {classes = classes2, default = default2, types = types2, log_types = _}) = pointer_eq (classes1, classes2) andalso default1 = default2 andalso pointer_eq (types1, types2); fun rep_tsig (TSig comps) = comps; fun make_tsig (classes, default, types, log_types) = TSig {classes = classes, default = default, types = types, log_types = log_types}; fun change_base begin (TSig {classes, default, types, log_types}) = make_tsig (classes, default, Name_Space.change_base begin types, log_types); fun change_ignore (TSig {classes, default, types, log_types}) = make_tsig (classes, default, Name_Space.change_ignore types, log_types); fun build_tsig (classes, default, types) = let val log_types = Name_Space.fold_table (fn (c, LogicalType n) => cons (c, n) | _ => I) types [] |> Library.sort (int_ord o apply2 snd) |> map fst; in make_tsig (classes, default, types, log_types) end; fun map_tsig f (TSig {classes, default, types, log_types = _}) = build_tsig (f (classes, default, types)); val empty_tsig = build_tsig ((Name_Space.empty Markup.classN, Sorts.empty_algebra), [], Name_Space.empty_table Markup.type_nameN); (* classes and sorts *) val class_space = #1 o #classes o rep_tsig; fun defaultS (TSig {default, ...}) = default; fun logical_types (TSig {log_types, ...}) = log_types; fun eq_sort (TSig {classes, ...}) = Sorts.sort_eq (#2 classes); fun subsort (TSig {classes, ...}) = Sorts.sort_le (#2 classes); fun of_sort (TSig {classes, ...}) = Sorts.of_sort (#2 classes); fun inter_sort (TSig {classes, ...}) = Sorts.inter_sort (#2 classes); fun cert_class (TSig {classes = (_, algebra), ...}) c = if can (Graph.get_entry (Sorts.classes_of algebra)) c then c else raise TYPE ("Undeclared class: " ^ quote c, [], []); val cert_sort = map o cert_class; fun minimize_sort (TSig {classes, ...}) = Sorts.minimize_sort (#2 classes); fun witness_sorts (TSig {classes, log_types, ...}) = Sorts.witness_sorts (#2 classes) log_types; (* certification mode *) datatype mode = Mode of {normalize: bool, logical: bool}; val mode_default = Mode {normalize = true, logical = true}; val mode_syntax = Mode {normalize = true, logical = false}; val mode_abbrev = Mode {normalize = false, logical = false}; structure Mode = Proof_Data ( type T = mode; fun init _ = mode_default; ); val get_mode = Mode.get; fun set_mode mode = Mode.map (K mode); fun restore_mode ctxt = set_mode (get_mode ctxt); (* types *) val type_space = Name_Space.space_of_table o #types o rep_tsig; fun type_alias naming binding name = map_tsig (fn (classes, default, types) => (classes, default, (Name_Space.alias_table naming binding name types))); fun undecl_type c = "Undeclared type constructor: " ^ quote c; fun lookup_type (TSig {types, ...}) = Name_Space.lookup types; fun check_decl context (TSig {types, ...}) (c, pos) = Name_Space.check_reports context types (c, [pos]); fun the_decl tsig (c, pos) = (case lookup_type tsig c of NONE => error (undecl_type c ^ Position.here pos) | SOME decl => decl); (* certified types *) fun bad_nargs t = "Bad number of arguments for type constructor: " ^ quote t; local fun inst_typ env (Type (c, Ts)) = Type (c, map (inst_typ env) Ts) | inst_typ env (T as TFree (x, _)) = the_default T (AList.lookup (op =) env x) | inst_typ _ T = T; in fun cert_typ_mode (Mode {normalize, logical}) tsig ty = let fun err msg = raise TYPE (msg, [ty], []); val check_logical = if logical then fn c => err ("Illegal occurrence of syntactic type: " ^ quote c) else fn _ => (); fun cert (T as Type (c, Ts)) = let val Ts' = map cert Ts; fun nargs n = if length Ts <> n then err (bad_nargs c) else (); in (case the_decl tsig (c, Position.none) of LogicalType n => (nargs n; Type (c, Ts')) | Abbreviation (vs, U, syn) => (nargs (length vs); if syn then check_logical c else (); if normalize then inst_typ (vs ~~ Ts') U else Type (c, Ts')) | Nonterminal => (nargs 0; check_logical c; T)) end | cert (TFree (x, S)) = TFree (x, cert_sort tsig S) | cert (TVar (xi as (_, i), S)) = if i < 0 then err ("Malformed type variable: " ^ quote (Term.string_of_vname xi)) else TVar (xi, cert_sort tsig S); val ty' = cert ty; in if ty = ty' then ty else ty' end; (*avoid copying of already normal type*) val cert_typ = cert_typ_mode mode_default; end; (* type arities *) fun arity_number tsig a = (case lookup_type tsig a of SOME (LogicalType n) => n | _ => error (undecl_type a)); fun arity_sorts _ tsig a [] = replicate (arity_number tsig a) [] | arity_sorts context (TSig {classes, ...}) a S = Sorts.mg_domain (#2 classes) a S handle Sorts.CLASS_ERROR err => error (Sorts.class_error context err); (** special treatment of type vars **) (* sort_of_atyp *) fun sort_of_atyp (TFree (_, S)) = S | sort_of_atyp (TVar (_, S)) = S | sort_of_atyp T = raise TYPE ("sort_of_atyp", [T], []); (* strip_sorts *) val strip_sorts = map_atyps (fn TFree (x, _) => TFree (x, []) | TVar (xi, _) => TVar (xi, [])); val strip_sorts_dummy = map_atyps (fn TFree (x, _) => TFree (x, dummyS) | TVar (xi, _) => TVar (xi, dummyS)); (* no_tvars *) fun no_tvars T = (case Term.add_tvarsT T [] of [] => T | vs => raise TYPE ("Illegal schematic type variable(s): " ^ commas_quote (map (Term.string_of_vname o #1) (rev vs)), [T], [])); (* varify_global *) fun varify_global fixed t = let val fs = - (t, []) |-> (Term.fold_types o Term.fold_atyps) - (fn TFree v => if Term_Subst.TFrees.defined fixed v then I else insert (op =) v | _ => I); + build (t |> (Term.fold_types o Term.fold_atyps) + (fn TFree v => if Term_Subst.TFrees.defined fixed v then I else insert (op =) v | _ => I)); val used = Name.context |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t; val fmap = fs ~~ map (rpair 0) (#1 (fold_map Name.variant (map fst fs) used)); fun thaw (f as (_, S)) = (case AList.lookup (op =) fmap f of NONE => TFree f | SOME xi => TVar (xi, S)); in (fmap, map_types (map_type_tfree thaw) t) end; (* freeze_thaw: freeze TVars in a term; return the "thaw" inverse *) local fun new_name ix (pairs, used) = let val v = singleton (Name.variant_list used) (string_of_indexname ix) in ((ix, v) :: pairs, v :: used) end; fun freeze_one alist (ix, sort) = TFree (the (AList.lookup (op =) alist ix), sort) handle Option.Option => raise TYPE ("Failure during freezing of ?" ^ string_of_indexname ix, [], []); fun thaw_one alist (a, sort) = TVar (the (AList.lookup (op =) alist a), sort) handle Option.Option => TFree (a, sort); in fun legacy_freeze_thaw_type T = let val used = Term.add_tfree_namesT T []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvarsT T [])) ([], used); in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end; val legacy_freeze_type = #1 o legacy_freeze_thaw_type; fun legacy_freeze_thaw t = let val used = Term.add_tfree_names t []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvars t [])) ([], used); in (case alist of [] => (t, fn x => x) (*nothing to do!*) | _ => (map_types (map_type_tvar (freeze_one alist)) t, map_types (map_type_tfree (thaw_one (map swap alist))))) end; val legacy_freeze = #1 o legacy_freeze_thaw; end; (** matching and unification of types **) type tyenv = (sort * typ) Vartab.table; fun tvar_clash ixn S S' = raise TYPE ("Type variable has two distinct sorts", [TVar (ixn, S), TVar (ixn, S')], []); fun lookup tye (ixn, S) = (case Vartab.lookup tye ixn of NONE => NONE | SOME (S', T) => if S = S' then SOME T else tvar_clash ixn S S'); (* matching *) exception TYPE_MATCH; fun typ_match tsig = let fun match (TVar (v, S), T) subs = (case lookup subs (v, S) of NONE => if of_sort tsig (T, S) then Vartab.update_new (v, (S, T)) subs else raise TYPE_MATCH | SOME U => if U = T then subs else raise TYPE_MATCH) | match (Type (a, Ts), Type (b, Us)) subs = if a <> b then raise TYPE_MATCH else matches (Ts, Us) subs | match (TFree x, TFree y) subs = if x = y then subs else raise TYPE_MATCH | match _ _ = raise TYPE_MATCH and matches (T :: Ts, U :: Us) subs = matches (Ts, Us) (match (T, U) subs) | matches _ subs = subs; in match end; fun typ_instance tsig (T, U) = (Vartab.build (typ_match tsig (U, T)); true) handle TYPE_MATCH => false; (*purely structural matching*) fun raw_match (TVar (v, S), T) subs = (case lookup subs (v, S) of NONE => Vartab.update_new (v, (S, T)) subs | SOME U => if U = T then subs else raise TYPE_MATCH) | raw_match (Type (a, Ts), Type (b, Us)) subs = if a <> b then raise TYPE_MATCH else raw_matches (Ts, Us) subs | raw_match (TFree x, TFree y) subs = if x = y then subs else raise TYPE_MATCH | raw_match _ _ = raise TYPE_MATCH and raw_matches (T :: Ts, U :: Us) subs = raw_matches (Ts, Us) (raw_match (T, U) subs) | raw_matches ([], []) subs = subs | raw_matches _ _ = raise TYPE_MATCH; (*fast matching filter*) fun could_match (Type (a, Ts), Type (b, Us)) = a = b andalso could_matches (Ts, Us) | could_match (TFree (a, _), TFree (b, _)) = a = b | could_match (TVar _, _) = true | could_match _ = false and could_matches (T :: Ts, U :: Us) = could_match (T, U) andalso could_matches (Ts, Us) | could_matches ([], []) = true | could_matches _ = false; fun raw_instance (T, U) = if could_match (U, T) then (Vartab.build (raw_match (U, T)); true) handle TYPE_MATCH => false else false; (* unification *) exception TUNIFY; (*occurs check*) fun occurs v tye = let fun occ (Type (_, Ts)) = exists occ Ts | occ (TFree _) = false | occ (TVar (w, S)) = Term.eq_ix (v, w) orelse (case lookup tye (w, S) of NONE => false | SOME U => occ U); in occ end; (*chase variable assignments; if devar returns a type var then it must be unassigned*) fun devar tye (T as TVar v) = (case lookup tye v of SOME U => devar tye U | NONE => T) | devar _ T = T; (*order-sorted unification*) fun unify (TSig {classes = (_, classes), ...}) TU (tyenv, maxidx) = let val tyvar_count = Unsynchronized.ref maxidx; fun gen_tyvar S = TVar ((Name.aT, Unsynchronized.inc tyvar_count), S); fun mg_domain a S = Sorts.mg_domain classes a S handle Sorts.CLASS_ERROR _ => raise TUNIFY; fun meet (_, []) tye = tye | meet (TVar (xi, S'), S) tye = if Sorts.sort_le classes (S', S) then tye else Vartab.update_new (xi, (S', gen_tyvar (Sorts.inter_sort classes (S', S)))) tye | meet (TFree (_, S'), S) tye = if Sorts.sort_le classes (S', S) then tye else raise TUNIFY | meet (Type (a, Ts), S) tye = meets (Ts, mg_domain a S) tye and meets (T :: Ts, S :: Ss) tye = meets (Ts, Ss) (meet (devar tye T, S) tye) | meets _ tye = tye; fun unif (ty1, ty2) tye = (case (devar tye ty1, devar tye ty2) of (T as TVar (v, S1), U as TVar (w, S2)) => if Term.eq_ix (v, w) then if S1 = S2 then tye else tvar_clash v S1 S2 else if Sorts.sort_le classes (S1, S2) then Vartab.update_new (w, (S2, T)) tye else if Sorts.sort_le classes (S2, S1) then Vartab.update_new (v, (S1, U)) tye else let val S = gen_tyvar (Sorts.inter_sort classes (S1, S2)) in Vartab.update_new (v, (S1, S)) (Vartab.update_new (w, (S2, S)) tye) end | (TVar (v, S), T) => if occurs v tye T then raise TUNIFY else meet (T, S) (Vartab.update_new (v, (S, T)) tye) | (T, TVar (v, S)) => if occurs v tye T then raise TUNIFY else meet (T, S) (Vartab.update_new (v, (S, T)) tye) | (Type (a, Ts), Type (b, Us)) => if a <> b then raise TUNIFY else unifs (Ts, Us) tye | (T, U) => if T = U then tye else raise TUNIFY) and unifs (T :: Ts, U :: Us) tye = unifs (Ts, Us) (unif (T, U) tye) | unifs _ tye = tye; in (unif TU tyenv, ! tyvar_count) end; (*purely structural unification*) fun raw_unify (ty1, ty2) tye = (case (devar tye ty1, devar tye ty2) of (T as TVar (v, S1), TVar (w, S2)) => if Term.eq_ix (v, w) then if S1 = S2 then tye else tvar_clash v S1 S2 else Vartab.update_new (w, (S2, T)) tye | (TVar (v, S), T) => if occurs v tye T then raise TUNIFY else Vartab.update_new (v, (S, T)) tye | (T, TVar (v, S)) => if occurs v tye T then raise TUNIFY else Vartab.update_new (v, (S, T)) tye | (Type (a, Ts), Type (b, Us)) => if a <> b then raise TUNIFY else raw_unifys (Ts, Us) tye | (T, U) => if T = U then tye else raise TUNIFY) and raw_unifys (T :: Ts, U :: Us) tye = raw_unifys (Ts, Us) (raw_unify (T, U) tye) | raw_unifys ([], []) tye = tye | raw_unifys _ _ = raise TUNIFY; (*fast unification filter*) fun could_unify (Type (a, Ts), Type (b, Us)) = a = b andalso could_unifys (Ts, Us) | could_unify (TFree (a, _), TFree (b, _)) = a = b | could_unify (TVar _, _) = true | could_unify (_, TVar _) = true | could_unify _ = false and could_unifys (T :: Ts, U :: Us) = could_unify (T, U) andalso could_unifys (Ts, Us) | could_unifys ([], []) = true | could_unifys _ = false; (*equality with respect to a type environment*) fun unified tye = let fun unif (T, T') = (case (devar tye T, devar tye T') of (Type (s, Ts), Type (s', Ts')) => s = s' andalso unifs (Ts, Ts') | (U, U') => U = U') and unifs ([], []) = true | unifs (T :: Ts, T' :: Ts') = unif (T', T') andalso unifs (Ts, Ts') | unifs _ = false; in if Vartab.is_empty tye then op = else unif end; (** extend and merge type signatures **) (* classes *) fun add_class context (c, cs) tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val cs' = map (cert_class tsig) cs handle TYPE (msg, _, _) => error msg; val _ = Binding.check c; val (c', space') = space |> Name_Space.declare context true c; val classes' = classes |> Sorts.add_class context (c', cs'); in ((space', classes'), default, types) end); fun hide_class fully c = map_tsig (fn ((space, classes), default, types) => ((Name_Space.hide fully c space, classes), default, types)); (* arities *) fun add_arity context (t, Ss, S) tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val _ = (case lookup_type tsig t of SOME (LogicalType n) => if length Ss <> n then error (bad_nargs t) else () | SOME _ => error ("Logical type constructor expected: " ^ quote t) | NONE => error (undecl_type t)); val (Ss', S') = (map (cert_sort tsig) Ss, cert_sort tsig S) handle TYPE (msg, _, _) => error msg; val classes' = classes |> Sorts.add_arities context ((t, map (fn c' => (c', Ss')) S')); in ((space, classes'), default, types) end); (* classrel *) fun add_classrel context rel tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val rel' = apply2 (cert_class tsig) rel handle TYPE (msg, _, _) => error msg; val classes' = classes |> Sorts.add_classrel context rel'; in ((space, classes'), default, types) end); (* default sort *) fun set_defsort S tsig = tsig |> map_tsig (fn (classes, _, types) => (classes, cert_sort tsig S handle TYPE (msg, _, _) => error msg, types)); (* types *) local fun new_decl context (c, decl) types = (Binding.check c; #2 (Name_Space.define context true (c, decl) types)); fun map_types f = map_tsig (fn (classes, default, types) => let val types' = f types; val _ = not (Name_Space.defined types' "dummy") orelse Name_Space.intern (Name_Space.space_of_table types') "dummy" = "dummy" orelse error "Illegal declaration of dummy type"; in (classes, default, types') end); fun syntactic tsig (Type (c, Ts)) = (case lookup_type tsig c of SOME Nonterminal => true | _ => false) orelse exists (syntactic tsig) Ts | syntactic _ _ = false; in fun add_type context (c, n) = if n < 0 then error ("Bad type constructor declaration " ^ Binding.print c) else map_types (new_decl context (c, LogicalType n)); fun add_abbrev context (a, vs, rhs) tsig = tsig |> map_types (fn types => let fun err msg = cat_error msg ("The error(s) above occurred in type abbreviation " ^ Binding.print a); val rhs' = strip_sorts (no_tvars (cert_typ_mode mode_syntax tsig rhs)) handle TYPE (msg, _, _) => err msg; val _ = (case duplicates (op =) vs of [] => [] | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups)); val _ = (case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of [] => [] | extras => err ("Extra variables on rhs: " ^ commas_quote extras)); in types |> new_decl context (a, Abbreviation (vs, rhs', syntactic tsig rhs')) end); fun add_nonterminal context = map_types o new_decl context o rpair Nonterminal; end; fun hide_type fully c = map_tsig (fn (classes, default, types) => (classes, default, Name_Space.hide_table fully c types)); (* merge type signatures *) fun merge_tsig context (tsig1, tsig2) = let val (TSig {classes = (space1, classes1), default = default1, types = types1, log_types = _}) = tsig1; val (TSig {classes = (space2, classes2), default = default2, types = types2, log_types = _}) = tsig2; val space' = Name_Space.merge (space1, space2); val classes' = Sorts.merge_algebra context (classes1, classes2); val default' = Sorts.inter_sort classes' (default1, default2); val types' = Name_Space.merge_tables (types1, types2); in build_tsig ((space', classes'), default', types') end; end;