diff --git a/src/Pure/Isar/class_declaration.ML b/src/Pure/Isar/class_declaration.ML --- a/src/Pure/Isar/class_declaration.ML +++ b/src/Pure/Isar/class_declaration.ML @@ -1,394 +1,391 @@ (* Title: Pure/Isar/class_declaration.ML Author: Florian Haftmann, TU Muenchen Declaring classes and subclass relations. *) signature CLASS_DECLARATION = sig val class: binding -> class list -> Element.context_i list -> theory -> string * local_theory val class_cmd: binding -> xstring list -> Element.context list -> theory -> string * local_theory val prove_subclass: tactic -> class -> local_theory -> local_theory val subclass: class -> local_theory -> Proof.state val subclass_cmd: xstring -> local_theory -> Proof.state end; structure Class_Declaration: CLASS_DECLARATION = struct (** class definitions **) local (* calculating class-related rules including canonical interpretation *) fun calculate thy class sups base_sort param_map assm_axiom = let val thy_ctxt = Proof_Context.init_global thy; val certT = Thm.trim_context_ctyp o Thm.global_ctyp_of thy; val cert = Thm.trim_context_cterm o Thm.global_cterm_of thy; (* instantiation of canonical interpretation *) val a_tfree = (Name.aT, base_sort); val a_type = TFree a_tfree; val param_map_const = (map o apsnd) Const param_map; val param_map_inst = param_map |> map (fn (x, (c, T)) => let val T' = map_atyps (K a_type) T in ((x, T'), cert (Const (c, T'))) end); val const_morph = Element.instantiate_normalize_morphism ([], param_map_inst); val typ_morph = Element.instantiate_normalize_morphism ([(a_tfree, certT (Term.aT [class]))], []) val (([raw_props], _, [(_, raw_inst_morph)], _, export_morph), _) = thy_ctxt |> Expression.cert_goal_expression ([(class, (("", false), (Expression.Named param_map_const, [])))], []); val (props, inst_morph) = if null param_map then (raw_props |> map (Morphism.term typ_morph), raw_inst_morph $> typ_morph) else (raw_props, raw_inst_morph); (*FIXME proper handling in locale.ML / expression.ML would be desirable*) (* witness for canonical interpretation *) val some_prop = try the_single props; val some_witn = Option.map (fn prop => let val sup_axioms = map_filter (fst o Class.rules thy) sups; val loc_intro_tac = (case Locale.intros_of thy class of (_, NONE) => all_tac | (_, SOME intro) => ALLGOALS (resolve_tac thy_ctxt [intro])); val tac = loc_intro_tac THEN ALLGOALS (Proof_Context.fact_tac thy_ctxt (sup_axioms @ the_list assm_axiom)); in Element.prove_witness thy_ctxt prop tac end) some_prop; val some_axiom = Option.map (Element.conclude_witness thy_ctxt) some_witn; (* canonical interpretation *) val base_morph = inst_morph $> Morphism.binding_morphism "class_binding" (Binding.prefix false (Class.class_prefix class)) $> Element.satisfy_morphism (the_list some_witn); val eq_morph = Element.eq_morphism thy (Class.these_defs thy sups); (* assm_intro *) fun prove_assm_intro thm = let val ((_, [thm']), _) = Variable.import true [thm] thy_ctxt; val const_eq_morph = (case eq_morph of SOME eq_morph => const_morph $> eq_morph | NONE => const_morph); val thm'' = Morphism.thm const_eq_morph thm'; in Goal.prove_sorry_global thy [] [] (Thm.prop_of thm'') (fn {context = ctxt, ...} => ALLGOALS (Proof_Context.fact_tac ctxt [thm''])) - |> tap (Thm.expose_proof thy) end; val some_assm_intro = Option.map prove_assm_intro (fst (Locale.intros_of thy class)); (* of_class *) val of_class_prop_concl = Logic.mk_of_class (a_type, class); val of_class_prop = (case some_prop of NONE => of_class_prop_concl | SOME prop => Logic.mk_implies (Morphism.term const_morph ((map_types o map_atyps) (K a_type) prop), of_class_prop_concl)); val sup_of_classes = map (snd o Class.rules thy) sups; val loc_axiom_intros = map Drule.export_without_context_open (Locale.axioms_of thy class); val axclass_intro = #intro (Axclass.get_info thy class); val base_sort_trivs = Thm.of_sort (Thm.global_ctyp_of thy a_type, base_sort); fun tac ctxt = REPEAT (SOMEGOAL (match_tac ctxt (axclass_intro :: sup_of_classes @ loc_axiom_intros @ base_sort_trivs) ORELSE' assume_tac ctxt)); - val of_class = - Goal.prove_sorry_global thy [] [] of_class_prop (tac o #context) - |> tap (Thm.expose_proof thy); + val of_class = Goal.prove_sorry_global thy [] [] of_class_prop (tac o #context); in (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) end; (* reading and processing class specifications *) fun prep_class_elems prep_decl thy sups raw_elems = let (* user space type system: only permits 'a type variable, improves towards 'a *) val algebra = Sign.classes_of thy; val inter_sort = curry (Sorts.inter_sort algebra); val proto_base_sort = if null sups then Sign.defaultS thy else fold inter_sort (map (Class.base_sort thy) sups) []; val is_param = member (op =) (map fst (Class.these_params thy sups)); val base_constraints = (map o apsnd) (map_type_tfree (K (TVar ((Name.aT, 0), proto_base_sort))) o fst o fst o snd) (Class.these_operations thy sups); val singleton_fixate = burrow_types (fn Ts => let val tfrees = fold Term.add_tfreesT Ts []; val inferred_sort = (fold o fold_atyps) (fn TVar (_, S) => inter_sort S | _ => I) Ts []; val fixate_sort = (case tfrees of [] => inferred_sort | [(a, S)] => if a <> Name.aT then error ("No type variable other than " ^ Name.aT ^ " allowed in class specification") else if Sorts.sort_le algebra (S, inferred_sort) then S else error ("Type inference imposes additional sort constraint " ^ Syntax.string_of_sort_global thy inferred_sort ^ " of type parameter " ^ Name.aT ^ " of sort " ^ Syntax.string_of_sort_global thy S) | _ => error "Multiple type variables in class specification"); val fixateT = Term.aT fixate_sort; in (map o map_atyps) (fn T as TVar (xi, _) => if Type_Infer.is_param xi then fixateT else T | T => T) Ts end); fun unify_params ts = let val param_Ts = (fold o fold_aterms) (fn Free (v, T) => if is_param v then fold_atyps (insert (op =)) T else I | _ => I) ts []; val param_namesT = map_filter (try (fst o dest_TVar)) param_Ts; val param_T = if null param_namesT then NONE else SOME (case get_first (try dest_TFree) param_Ts of SOME v_sort => TFree v_sort | NONE => TVar (hd param_namesT, proto_base_sort)); in case param_T of NONE => ts | SOME T => map (subst_TVars (map (rpair T) param_namesT)) ts end; (* preprocessing elements, retrieving base sort from type-checked elements *) val raw_supexpr = (map (fn sup => (sup, (("", false), (Expression.Positional [], [])))) sups, []); val init_class_body = fold (Proof_Context.add_const_constraint o apsnd SOME) base_constraints #> Class.redeclare_operations thy sups #> Context.proof_map (Syntax_Phases.term_check 0 "singleton_fixate" (K singleton_fixate)); val ((raw_supparams, _, raw_inferred_elems, _), _) = Proof_Context.init_global thy |> Context.proof_map (Syntax_Phases.term_check 0 "unify_params" (K unify_params)) |> prep_decl raw_supexpr init_class_body raw_elems; fun filter_element (Element.Fixes []) = NONE | filter_element (e as Element.Fixes _) = SOME e | filter_element (Element.Constrains []) = NONE | filter_element (e as Element.Constrains _) = SOME e | filter_element (Element.Assumes []) = NONE | filter_element (e as Element.Assumes _) = SOME e | filter_element (Element.Defines _) = error ("\"defines\" element not allowed in class specification.") | filter_element (Element.Notes _) = error ("\"notes\" element not allowed in class specification.") | filter_element (Element.Lazy_Notes _) = error ("\"notes\" element not allowed in class specification."); val inferred_elems = map_filter filter_element raw_inferred_elems; fun fold_element_types f (Element.Fixes fxs) = fold (fn (_, SOME T, _) => f T) fxs | fold_element_types f (Element.Constrains cnstrs) = fold (f o snd) cnstrs | fold_element_types f (Element.Assumes assms) = fold (fold (fn (t, ts) => fold_types f t #> (fold o fold_types) f ts) o snd) assms; val base_sort = if null inferred_elems then proto_base_sort else (case (fold o fold_element_types) Term.add_tfreesT inferred_elems [] of [] => error "No type variable in class specification" | [(_, sort)] => sort | _ => error "Multiple type variables in class specification"); val supparams = map (fn ((c, T), _) => (c, map_atyps (K (Term.aT base_sort)) T)) raw_supparams; val supparam_names = map fst supparams; fun mk_param ((c, _), _) = Free (c, (the o AList.lookup (op =) supparams) c); val supexpr = (map (fn sup => (sup, (("", false), (Expression.Positional (map (SOME o mk_param) (Locale.params_of thy sup)), [])))) sups, map (fn (c, T) => (Binding.name c, SOME T, NoSyn)) supparams); in (base_sort, supparam_names, supexpr, inferred_elems) end; val cert_class_elems = prep_class_elems Expression.cert_declaration; val read_class_elems = prep_class_elems Expression.cert_read_declaration; fun prep_class_spec prep_class prep_class_elems thy raw_supclasses raw_elems = let val thy_ctxt = Proof_Context.init_global thy; (* prepare import *) val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy)); val sups = Sign.minimize_sort thy (map (prep_class thy_ctxt) raw_supclasses); val _ = (case filter_out (Class.is_class thy) sups of [] => () | no_classes => error ("No (proper) classes: " ^ commas_quote no_classes)); val raw_supparams = (map o apsnd) (snd o snd) (Class.these_params thy sups); val raw_supparam_names = map fst raw_supparams; val _ = if has_duplicates (op =) raw_supparam_names then error ("Duplicate parameter(s) in superclasses: " ^ (commas_quote (duplicates (op =) raw_supparam_names))) else (); (* infer types and base sort *) val (base_sort, supparam_names, supexpr, inferred_elems) = prep_class_elems thy sups raw_elems; val sup_sort = inter_sort base_sort sups; (* process elements as class specification *) val class_ctxt = Class.begin sups base_sort thy_ctxt; val ((_, _, syntax_elems, _), _) = class_ctxt |> Expression.cert_declaration supexpr I inferred_elems; fun check_vars e vs = if null vs then error ("No type variable in part of specification element " ^ Pretty.string_of (Pretty.chunks (Element.pretty_ctxt class_ctxt e))) else (); fun check_element (e as Element.Fixes fxs) = List.app (fn (_, SOME T, _) => check_vars e (Term.add_tfreesT T [])) fxs | check_element (e as Element.Assumes assms) = List.app (fn (_, ts_pss) => List.app (fn (t, _) => check_vars e (Term.add_tfrees t [])) ts_pss) assms | check_element _ = (); val _ = List.app check_element syntax_elems; fun fork_syn (Element.Fixes xs) = fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs #>> Element.Fixes | fork_syn x = pair x; val (elems, global_syntax) = fold_map fork_syn syntax_elems []; in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (elems, global_syntax)) end; val cert_class_spec = prep_class_spec (K I) cert_class_elems; val read_class_spec = prep_class_spec Proof_Context.read_class read_class_elems; (* class establishment *) fun add_consts class base_sort sups supparam_names global_syntax thy = let (*FIXME simplify*) val supconsts = supparam_names |> AList.make (snd o the o AList.lookup (op =) (Class.these_params thy sups)) |> (map o apsnd o apsnd o map_atyps) (K (Term.aT [class])); val all_params = Locale.params_of thy class; val raw_params = (snd o chop (length supparam_names)) all_params; fun add_const ((raw_c, raw_ty), _) thy = let val b = Binding.name raw_c; val c = Sign.full_name thy b; val ty = map_atyps (K (Term.aT base_sort)) raw_ty; val ty0 = Type.strip_sorts ty; val ty' = map_atyps (K (Term.aT [class])) ty0; val syn = (the_default NoSyn o AList.lookup Binding.eq_name global_syntax) b; in thy |> Sign.declare_const_global ((b, ty0), syn) |> snd |> pair ((Variable.check_name b, ty), (c, ty')) end; in thy |> Sign.add_path (Class.class_prefix class) |> fold_map add_const raw_params ||> Sign.restore_naming thy |-> (fn params => pair (supconsts @ (map o apfst) fst params, params)) end; fun adjungate_axclass bname class base_sort sups supsort supparam_names global_syntax thy = let (*FIXME simplify*) fun globalize param_map = map_aterms (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty) | t => t); val raw_pred = Locale.intros_of thy class |> fst |> Option.map (Logic.unvarify_global o Logic.strip_imp_concl o Thm.prop_of); fun get_axiom thy = (case #axioms (Axclass.get_info thy class) of [] => NONE | [thm] => SOME thm); in thy |> add_consts class base_sort sups supparam_names global_syntax |-> (fn (param_map, params) => Axclass.define_class (bname, supsort) (map (fst o snd) params) [(Binding.empty_atts, Option.map (globalize param_map) raw_pred |> the_list)] #> snd #> `get_axiom #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params #> pair (param_map, params, assm_axiom))) end; fun gen_class prep_class_spec b raw_supclasses raw_elems thy = let val class = Sign.full_name thy b; val prefix = Binding.qualify true "class"; val (((sups, supparam_names), (supsort, base_sort, supexpr)), (elems, global_syntax)) = prep_class_spec thy raw_supclasses raw_elems; in thy |> Expression.add_locale b (prefix b) supexpr elems |> snd |> Local_Theory.exit_global |> adjungate_axclass b class base_sort sups supsort supparam_names global_syntax |-> (fn (param_map, params, assm_axiom) => `(fn thy => calculate thy class sups base_sort param_map assm_axiom) #-> (fn (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) => Locale.add_registration_theory {inst = (class, base_morph), mixin = Option.map (rpair true) eq_morph, export = export_morph} #> Class.register class sups params base_sort base_morph export_morph some_axiom some_assm_intro of_class #> Global_Theory.store_thm (prefix (Binding.qualified_name (class ^ ".of_class.intro")), of_class))) |> snd |> Named_Target.init class |> pair class end; in val class = gen_class cert_class_spec; val class_cmd = gen_class read_class_spec; end; (*local*) (** subclass relations **) local fun gen_subclass prep_class do_proof raw_sup lthy = let val thy = Proof_Context.theory_of lthy; val proto_sup = prep_class thy raw_sup; val proto_sub = (case Named_Target.class_of lthy of SOME class => class | NONE => error "Not in a class target"); val (sub, sup) = Axclass.cert_classrel thy (proto_sub, proto_sup); val expr = ([(sup, (("", false), (Expression.Positional [], [])))], []); val (([props], _, deps, _, export), goal_ctxt) = Expression.cert_goal_expression expr lthy; val some_prop = try the_single props; val some_dep_morph = try the_single (map snd deps); fun after_qed some_wit = Class.register_subclass (sub, sup) some_dep_morph some_wit export; in do_proof after_qed some_prop goal_ctxt end; fun user_proof after_qed some_prop = Element.witness_proof (after_qed o try the_single o the_single) [the_list some_prop]; fun tactic_proof tac after_qed some_prop ctxt = after_qed (Option.map (fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt; in fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac); fun subclass x = gen_subclass (K I) user_proof x; fun subclass_cmd x = gen_subclass (Proof_Context.read_class o Proof_Context.init_global) user_proof x; end; (*local*) end; diff --git a/src/Pure/Isar/element.ML b/src/Pure/Isar/element.ML --- a/src/Pure/Isar/element.ML +++ b/src/Pure/Isar/element.ML @@ -1,473 +1,473 @@ (* Title: Pure/Isar/element.ML Author: Makarius Explicit data structures for some Isar language elements, with derived logical operations. *) signature ELEMENT = sig type ('typ, 'term) obtain = binding * ((binding * 'typ option * mixfix) list * 'term list) type obtains = (string, string) obtain list type obtains_i = (typ, term) obtain list datatype ('typ, 'term) stmt = Shows of (Attrib.binding * ('term * 'term list) list) list | Obtains of ('typ, 'term) obtain list type statement = (string, string) stmt type statement_i = (typ, term) stmt datatype ('typ, 'term, 'fact) ctxt = Fixes of (binding * 'typ option * mixfix) list | Constrains of (string * 'typ) list | Assumes of (Attrib.binding * ('term * 'term list) list) list | Defines of (Attrib.binding * ('term * 'term list)) list | Notes of string * (Attrib.binding * ('fact * Token.src list) list) list | Lazy_Notes of string * (binding * 'fact lazy) type context = (string, string, Facts.ref) ctxt type context_i = (typ, term, thm list) ctxt val map_ctxt: {binding: binding -> binding, typ: 'typ -> 'a, term: 'term -> 'b, pattern: 'term -> 'b, fact: 'fact -> 'c, attrib: Token.src -> Token.src} -> ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt val map_ctxt_attrib: (Token.src -> Token.src) -> ('typ, 'term, 'fact) ctxt -> ('typ, 'term, 'fact) ctxt val transform_ctxt: morphism -> context_i -> context_i val pretty_stmt: Proof.context -> statement_i -> Pretty.T list val pretty_ctxt: Proof.context -> context_i -> Pretty.T list val pretty_ctxt_no_attribs: Proof.context -> context_i -> Pretty.T list val pretty_statement: Proof.context -> string -> thm -> Pretty.T type witness val prove_witness: Proof.context -> term -> tactic -> witness val witness_proof: (witness list list -> Proof.context -> Proof.context) -> term list list -> Proof.context -> Proof.state val witness_proof_eqs: (witness list list -> thm list -> Proof.context -> Proof.context) -> term list list -> term list -> Proof.context -> Proof.state val witness_local_proof: (witness list list -> Proof.state -> Proof.state) -> string -> term list list -> Proof.context -> Proof.state -> Proof.state val witness_local_proof_eqs: (witness list list -> thm list -> Proof.state -> Proof.state) -> string -> term list list -> term list -> Proof.context -> Proof.state -> Proof.state val transform_witness: morphism -> witness -> witness val conclude_witness: Proof.context -> witness -> thm val pretty_witness: Proof.context -> witness -> Pretty.T val instantiate_normalize_morphism: ((string * sort) * ctyp) list * ((string * typ) * cterm) list -> morphism val satisfy_morphism: witness list -> morphism val eq_term_morphism: theory -> term list -> morphism option val eq_morphism: theory -> thm list -> morphism option val init: context_i -> Context.generic -> Context.generic val activate_i: context_i -> Proof.context -> context_i * Proof.context val activate: (typ, term, Facts.ref) ctxt -> Proof.context -> context_i * Proof.context end; structure Element: ELEMENT = struct (** language elements **) (* statement *) type ('typ, 'term) obtain = binding * ((binding * 'typ option * mixfix) list * 'term list); type obtains = (string, string) obtain list; type obtains_i = (typ, term) obtain list; datatype ('typ, 'term) stmt = Shows of (Attrib.binding * ('term * 'term list) list) list | Obtains of ('typ, 'term) obtain list; type statement = (string, string) stmt; type statement_i = (typ, term) stmt; (* context *) datatype ('typ, 'term, 'fact) ctxt = Fixes of (binding * 'typ option * mixfix) list | Constrains of (string * 'typ) list | Assumes of (Attrib.binding * ('term * 'term list) list) list | Defines of (Attrib.binding * ('term * 'term list)) list | Notes of string * (Attrib.binding * ('fact * Token.src list) list) list | Lazy_Notes of string * (binding * 'fact lazy); type context = (string, string, Facts.ref) ctxt; type context_i = (typ, term, thm list) ctxt; fun map_ctxt {binding, typ, term, pattern, fact, attrib} = fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) => (binding x, Option.map typ T, mx))) | Constrains xs => Constrains (xs |> map (fn (x, T) => (Variable.check_name (binding (Binding.name x)), typ T))) | Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) => ((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map pattern ps))))) | Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) => ((binding a, map attrib atts), (term t, map pattern ps)))) | Notes (kind, facts) => Notes (kind, facts |> map (fn ((a, atts), bs) => ((binding a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts))))) | Lazy_Notes (kind, (a, ths)) => Lazy_Notes (kind, (binding a, Lazy.map fact ths)); fun map_ctxt_attrib attrib = map_ctxt {binding = I, typ = I, term = I, pattern = I, fact = I, attrib = attrib}; fun transform_ctxt phi = map_ctxt {binding = Morphism.binding phi, typ = Morphism.typ phi, term = Morphism.term phi, pattern = Morphism.term phi, fact = Morphism.fact phi, attrib = map (Token.transform phi)}; (** pretty printing **) fun pretty_items _ _ [] = [] | pretty_items keyword sep (x :: ys) = Pretty.block [Pretty.keyword2 keyword, Pretty.brk 1, x] :: map (fn y => Pretty.block [Pretty.str " ", Pretty.keyword2 sep, Pretty.brk 1, y]) ys; (* pretty_stmt *) fun pretty_stmt ctxt = let val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; val prt_term = Pretty.quote o Syntax.pretty_term ctxt; val prt_terms = separate (Pretty.keyword2 "and") o map prt_term; val prt_binding = Attrib.pretty_binding ctxt; val prt_name = Proof_Context.pretty_name ctxt; fun prt_show (a, ts) = Pretty.block (Pretty.breaks (prt_binding a ":" @ prt_terms (map fst ts))); fun prt_var (x, SOME T, _) = Pretty.block [prt_name (Binding.name_of x), Pretty.str " ::", Pretty.brk 1, prt_typ T] | prt_var (x, NONE, _) = prt_name (Binding.name_of x); val prt_vars = separate (Pretty.keyword2 "and") o map prt_var; fun prt_obtain (_, ([], props)) = Pretty.block (Pretty.breaks (prt_terms props)) | prt_obtain (_, (vars, props)) = Pretty.block (Pretty.breaks (prt_vars vars @ [Pretty.keyword2 "where"] @ prt_terms props)); in fn Shows shows => pretty_items "shows" "and" (map prt_show shows) | Obtains obtains => pretty_items "obtains" "|" (map prt_obtain obtains) end; (* pretty_ctxt *) fun gen_pretty_ctxt show_attribs ctxt = let val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; val prt_term = Pretty.quote o Syntax.pretty_term ctxt; val prt_thm = Pretty.cartouche o Thm.pretty_thm ctxt; val prt_name = Proof_Context.pretty_name ctxt; fun prt_binding (b, atts) = Attrib.pretty_binding ctxt (b, if show_attribs then atts else []); fun prt_fact (ths, atts) = if not show_attribs orelse null atts then map prt_thm ths else Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths)) :: Attrib.pretty_attribs ctxt atts; fun prt_mixfix NoSyn = [] | prt_mixfix mx = [Pretty.brk 2, Mixfix.pretty_mixfix mx]; fun prt_fix (x, SOME T, mx) = Pretty.block (prt_name (Binding.name_of x) :: Pretty.str " ::" :: Pretty.brk 1 :: prt_typ T :: prt_mixfix mx) | prt_fix (x, NONE, mx) = Pretty.block (prt_name (Binding.name_of x) :: prt_mixfix mx); fun prt_constrain (x, T) = prt_fix (Binding.name x, SOME T, NoSyn); fun prt_asm (a, ts) = Pretty.block (Pretty.breaks (prt_binding a ":" @ map (prt_term o fst) ts)); fun prt_def (a, (t, _)) = Pretty.block (Pretty.breaks (prt_binding a ":" @ [prt_term t])); fun prt_note (a, ths) = Pretty.block (Pretty.breaks (flat (prt_binding a " =" :: map prt_fact ths))); fun notes_kind "" = "notes" | notes_kind kind = "notes " ^ kind; in fn Fixes fixes => pretty_items "fixes" "and" (map prt_fix fixes) | Constrains xs => pretty_items "constrains" "and" (map prt_constrain xs) | Assumes asms => pretty_items "assumes" "and" (map prt_asm asms) | Defines defs => pretty_items "defines" "and" (map prt_def defs) | Notes (kind, facts) => pretty_items (notes_kind kind) "and" (map prt_note facts) | Lazy_Notes (kind, (a, ths)) => pretty_items (notes_kind kind) "and" [prt_note ((a, []), [(Lazy.force ths, [])])] end; val pretty_ctxt = gen_pretty_ctxt true; val pretty_ctxt_no_attribs = gen_pretty_ctxt false; (* pretty_statement *) local fun standard_elim ctxt th = (case Object_Logic.elim_concl ctxt th of SOME C => let val thesis = Var ((Auto_Bind.thesisN, Thm.maxidx_of th + 1), fastype_of C); val th' = Thm.instantiate ([], [(Term.dest_Var C, Thm.cterm_of ctxt thesis)]) th; in (th', true) end | NONE => (th, false)); fun thm_name ctxt kind th prts = let val head = if Thm.has_name_hint th then Pretty.block [Pretty.keyword1 kind, Pretty.brk 1, Proof_Context.pretty_name ctxt (Long_Name.base_name (Thm.get_name_hint th)), Pretty.str ":"] else Pretty.keyword1 kind in Pretty.block (Pretty.fbreaks (head :: prts)) end; fun obtain prop ctxt = let val ((ps, prop'), ctxt') = Variable.focus NONE prop ctxt; fun fix (x, T) = (Binding.name (Variable.revert_fixed ctxt' x), SOME T, NoSyn); val xs = map (fix o #2) ps; val As = Logic.strip_imp_prems prop'; in ((Binding.empty, (xs, As)), ctxt') end; in fun pretty_statement ctxt kind raw_th = let val (th, is_elim) = standard_elim ctxt (Raw_Simplifier.norm_hhf ctxt raw_th); val ((_, [th']), ctxt') = Variable.import true [th] (Variable.set_body true ctxt); val prop = Thm.prop_of th'; val (prems, concl) = Logic.strip_horn prop; val concl_term = Object_Logic.drop_judgment ctxt concl; val (assumes, cases) = chop_suffix (fn prem => is_elim andalso concl aconv Logic.strip_assums_concl prem) prems; val is_thesis = if null cases then K false else fn v => v aconv concl_term; val fixes = rev (fold_aterms (fn v as Free (x, T) => if Variable.is_newly_fixed ctxt' ctxt x andalso not (is_thesis v) then insert (op =) (Variable.revert_fixed ctxt' x, T) else I | _ => I) prop []); in pretty_ctxt ctxt' (Fixes (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) fixes)) @ pretty_ctxt ctxt' (Assumes (map (fn t => (Binding.empty_atts, [(t, [])])) assumes)) @ (if null cases then pretty_stmt ctxt' (Shows [(Binding.empty_atts, [(concl, [])])]) else let val (clauses, ctxt'') = fold_map obtain cases ctxt' in pretty_stmt ctxt'' (Obtains clauses) end) end |> thm_name ctxt kind raw_th; end; (** logical operations **) (* witnesses -- hypotheses as protected facts *) datatype witness = Witness of term * thm; val mark_witness = Logic.protect; fun witness_prop (Witness (t, _)) = t; fun witness_hyps (Witness (_, th)) = Thm.hyps_of th; fun map_witness f (Witness witn) = Witness (f witn); fun transform_witness phi = map_witness (fn (t, th) => (Morphism.term phi t, Morphism.thm phi th)); fun prove_witness ctxt t tac = Witness (t, Goal.prove ctxt [] [] (mark_witness t) (fn _ => resolve_tac ctxt [Drule.protectI] 1 THEN tac) - |> Thm.expose_derivation \<^here>); + |> Thm.close_derivation \<^here>); local val refine_witness = Proof.refine_singleton (Method.Basic (fn ctxt => CONTEXT_TACTIC o K (ALLGOALS (CONJUNCTS (ALLGOALS (CONJUNCTS (TRYALL (resolve_tac ctxt [Drule.protectI])))))))); fun gen_witness_proof proof after_qed wit_propss eq_props = let val propss = (map o map) (fn prop => (mark_witness prop, [])) wit_propss @ [map (rpair []) eq_props]; fun after_qed' thmss = - let val (wits, eqs) = split_last ((map o map) (Thm.expose_derivation \<^here>) thmss); + let val (wits, eqs) = split_last ((map o map) (Thm.close_derivation \<^here>) thmss); in after_qed ((map2 o map2) (curry Witness) wit_propss wits) eqs end; in proof after_qed' propss #> refine_witness end; fun proof_local cmd goal_ctxt after_qed propp = let fun after_qed' (result_ctxt, results) state' = after_qed (burrow (Proof_Context.export result_ctxt (Proof.context_of state')) results) state'; in Proof.map_context (K goal_ctxt) #> Proof.internal_goal (K (K ())) (Proof_Context.get_mode goal_ctxt) true cmd NONE after_qed' [] [] (map (pair Binding.empty_atts) propp) #> snd end; in fun witness_proof after_qed wit_propss = gen_witness_proof (Proof.theorem NONE) (fn wits => fn _ => after_qed wits) wit_propss []; val witness_proof_eqs = gen_witness_proof (Proof.theorem NONE); fun witness_local_proof after_qed cmd wit_propss goal_ctxt = gen_witness_proof (proof_local cmd goal_ctxt) (fn wits => fn _ => after_qed wits) wit_propss []; fun witness_local_proof_eqs after_qed cmd wit_propss eq_props goal_ctxt = gen_witness_proof (proof_local cmd goal_ctxt) after_qed wit_propss eq_props; end; fun conclude_witness ctxt (Witness (_, th)) = Goal.conclude th |> Raw_Simplifier.norm_hhf_protect ctxt - |> Thm.expose_derivation \<^here>; + |> Thm.close_derivation \<^here>; fun pretty_witness ctxt witn = let val prt_term = Pretty.quote o Syntax.pretty_term ctxt in Pretty.block (prt_term (witness_prop witn) :: (if Config.get ctxt show_hyps then [Pretty.brk 2, Pretty.list "[" "]" (map prt_term (witness_hyps witn))] else [])) end; (* instantiate frees, with beta normalization *) fun instantiate_normalize_morphism insts = Morphism.instantiate_frees_morphism insts $> Morphism.term_morphism "beta_norm" Envir.beta_norm $> Morphism.thm_morphism "beta_conversion" (Conv.fconv_rule (Thm.beta_conversion true)); (* satisfy hypotheses *) local val norm_term = Envir.beta_eta_contract; val norm_conv = Drule.beta_eta_conversion; val norm_cterm = Thm.rhs_of o norm_conv; fun find_witness witns hyp = (case find_first (fn Witness (t, _) => hyp aconv t) witns of NONE => let val hyp' = norm_term hyp in find_first (fn Witness (t, _) => hyp' aconv norm_term t) witns end | some => some); fun compose_witness (Witness (_, th)) r = let val th' = Goal.conclude th; val A = Thm.cprem_of r 1; in Thm.implies_elim (Conv.gconv_rule norm_conv 1 r) (Conv.fconv_rule norm_conv (Thm.instantiate (Thm.match (apply2 norm_cterm (Thm.cprop_of th', A))) th')) end; in fun satisfy_thm witns thm = (Thm.chyps_of thm, thm) |-> fold (fn hyp => (case find_witness witns (Thm.term_of hyp) of NONE => I | SOME w => Thm.implies_intr hyp #> compose_witness w)); val satisfy_morphism = Morphism.thm_morphism "Element.satisfy" o satisfy_thm; end; (* rewriting with equalities *) (* for activating declarations only *) fun eq_term_morphism _ [] = NONE | eq_term_morphism thy props = let fun decomp_simp prop = let val ctxt = Proof_Context.init_global thy; val _ = Logic.count_prems prop = 0 orelse error ("Bad conditional rewrite rule " ^ Syntax.string_of_term ctxt prop); val lhsrhs = Logic.dest_equals prop handle TERM _ => error ("Rewrite rule not a meta-equality " ^ Syntax.string_of_term ctxt prop); in lhsrhs end; val phi = Morphism.morphism "Element.eq_term_morphism" {binding = [], typ = [], term = [Pattern.rewrite_term thy (map decomp_simp props) []], fact = [fn _ => error "Illegal application of Element.eq_term_morphism"]}; in SOME phi end; fun eq_morphism _ [] = NONE | eq_morphism thy thms = let (* FIXME proper context!? *) fun rewrite th = rewrite_rule (Proof_Context.init_global (Thm.theory_of_thm th)) thms th; val phi = Morphism.morphism "Element.eq_morphism" {binding = [], typ = [], term = [Raw_Simplifier.rewrite_term thy thms []], fact = [map rewrite]}; in SOME phi end; (** activate in context **) (* init *) fun init (Fixes fixes) = Context.map_proof (Proof_Context.add_fixes fixes #> #2) | init (Constrains _) = I | init (Assumes asms) = Context.map_proof (fn ctxt => let val asms' = Attrib.map_specs (map (Attrib.attribute ctxt)) asms; val (_, ctxt') = ctxt |> fold Proof_Context.augment (maps (map #1 o #2) asms') |> Proof_Context.add_assms Assumption.assume_export asms'; in ctxt' end) | init (Defines defs) = Context.map_proof (fn ctxt => let val defs' = Attrib.map_specs (map (Attrib.attribute ctxt)) defs; val asms = defs' |> map (fn (b, (t, ps)) => let val (_, t') = Local_Defs.cert_def ctxt (K []) t (* FIXME adapt ps? *) in (t', (b, [(t', ps)])) end); val (_, ctxt') = ctxt |> fold Proof_Context.augment (map #1 asms) |> Proof_Context.add_assms Local_Defs.def_export (map #2 asms); in ctxt' end) | init (Notes (kind, facts)) = Attrib.generic_notes kind facts #> #2 | init (Lazy_Notes (kind, ths)) = Attrib.lazy_notes kind ths; (* activate *) fun activate_i elem ctxt = let val elem' = (case (map_ctxt_attrib o map) Token.init_assignable elem of Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) => ((Thm.def_binding_optional (Binding.name (#1 (#1 (Local_Defs.cert_def ctxt (K []) t)))) a, atts), (t, ps)))) | e => e); val ctxt' = Context.proof_map (init elem') ctxt; in ((map_ctxt_attrib o map) Token.closure elem', ctxt') end; fun activate raw_elem ctxt = let val elem = raw_elem |> map_ctxt {binding = I, typ = I, term = I, pattern = I, fact = Proof_Context.get_fact ctxt, attrib = Attrib.check_src ctxt} in activate_i elem ctxt end; end; diff --git a/src/Pure/Isar/expression.ML b/src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML +++ b/src/Pure/Isar/expression.ML @@ -1,887 +1,886 @@ (* Title: Pure/Isar/expression.ML Author: Clemens Ballarin, TU Muenchen Locale expressions and user interface layer of locales. *) signature EXPRESSION = sig (* Locale expressions *) datatype 'term map = Positional of 'term option list | Named of (string * 'term) list type 'term rewrites = (Attrib.binding * 'term) list type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list type expression_i = (string, term) expr * (binding * typ option * mixfix) list type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list (* Processing of context statements *) val cert_statement: Element.context_i list -> Element.statement_i -> Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context val read_statement: Element.context list -> Element.statement -> Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context (* Declaring locales *) val cert_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context_i list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) val cert_read_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) (*FIXME*) val read_declaration: expression -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) val add_locale: binding -> binding -> expression_i -> Element.context_i list -> theory -> string * local_theory val add_locale_cmd: binding -> binding -> expression -> Element.context list -> theory -> string * local_theory (* Processing of locale expressions *) val cert_goal_expression: expression_i -> Proof.context -> (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context val read_goal_expression: expression -> Proof.context -> (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context (* Diagnostic *) val pretty_dependencies: Proof.context -> bool -> expression -> Pretty.T end; structure Expression : EXPRESSION = struct datatype ctxt = datatype Element.ctxt; (*** Expressions ***) datatype 'term map = Positional of 'term option list | Named of (string * 'term) list; type 'term rewrites = (Attrib.binding * 'term) list; type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list; type expression_i = (string, term) expr * (binding * typ option * mixfix) list; type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list; (** Internalise locale names in expr **) fun check_expr thy instances = map (apfst (Locale.check thy)) instances; (** Parameters of expression **) (*Sanity check of instantiations and extraction of implicit parameters. The latter only occurs iff strict = false. Positional instantiations are extended to match full length of parameter list of instantiated locale.*) fun parameters_of thy strict (expr, fixed) = let val ctxt = Proof_Context.init_global thy; fun reject_dups message xs = (case duplicates (op =) xs of [] => () | dups => error (message ^ commas dups)); fun parm_eq ((p1, mx1), (p2, mx2)) = p1 = p2 andalso (Mixfix.equal (mx1, mx2) orelse error ("Conflicting syntax for parameter " ^ quote p1 ^ " in expression" ^ Position.here_list [Mixfix.pos_of mx1, Mixfix.pos_of mx2])); fun params_loc loc = Locale.params_of thy loc |> map (apfst #1); fun params_inst (loc, (prfx, (Positional insts, eqns))) = let val ps = params_loc loc; val d = length ps - length insts; val insts' = if d < 0 then error ("More arguments than parameters in instantiation of locale " ^ quote (Locale.markup_name ctxt loc)) else insts @ replicate d NONE; val ps' = (ps ~~ insts') |> map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE); in (ps', (loc, (prfx, (Positional insts', eqns)))) end | params_inst (loc, (prfx, (Named insts, eqns))) = let val _ = reject_dups "Duplicate instantiation of the following parameter(s): " (map fst insts); val ps' = (insts, params_loc loc) |-> fold (fn (p, _) => fn ps => if AList.defined (op =) ps p then AList.delete (op =) p ps else error (quote p ^ " not a parameter of instantiated expression")); in (ps', (loc, (prfx, (Named insts, eqns)))) end; fun params_expr is = let val (is', ps') = fold_map (fn i => fn ps => let val (ps', i') = params_inst i; val ps'' = distinct parm_eq (ps @ ps'); in (i', ps'') end) is [] in (ps', is') end; val (implicit, expr') = params_expr expr; val implicit' = map #1 implicit; val fixed' = map (Variable.check_name o #1) fixed; val _ = reject_dups "Duplicate fixed parameter(s): " fixed'; val implicit'' = if strict then [] else let val _ = reject_dups "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed'); in map (fn (x, mx) => (Binding.name x, NONE, mx)) implicit end; in (expr', implicit'' @ fixed) end; (** Read instantiation **) (* Parse positional or named instantiation *) local fun prep_inst prep_term ctxt parms (Positional insts) = (insts ~~ parms) |> map (fn (NONE, p) => Free (p, dummyT) | (SOME t, _) => prep_term ctxt t) | prep_inst prep_term ctxt parms (Named insts) = parms |> map (fn p => (case AList.lookup (op =) insts p of SOME t => prep_term ctxt t | NONE => Free (p, dummyT))); in fun parse_inst x = prep_inst Syntax.parse_term x; fun make_inst x = prep_inst (K I) x; end; (* Instantiation morphism *) fun inst_morphism params ((prfx, mandatory), insts') ctxt = let (* parameters *) val parm_types = map #2 params; val type_parms = fold Term.add_tfreesT parm_types []; (* type inference *) val parm_types' = map (Type_Infer.paramify_vars o Logic.varifyT_global) parm_types; val type_parms' = fold Term.add_tvarsT parm_types' []; val checked = (map (Logic.mk_type o TVar) type_parms' @ map2 Type.constraint parm_types' insts') |> Syntax.check_terms (Config.put Type_Infer.object_logic false ctxt) val (type_parms'', insts'') = chop (length type_parms') checked; (* context *) val ctxt' = fold Proof_Context.augment checked ctxt; val certT = Thm.trim_context_ctyp o Thm.ctyp_of ctxt'; val cert = Thm.trim_context_cterm o Thm.cterm_of ctxt'; (* instantiation *) val instT = (type_parms ~~ map Logic.dest_type type_parms'') |> map_filter (fn (v, T) => if TFree v = T then NONE else SOME (v, T)); val cert_inst = ((map #1 params ~~ map (Term_Subst.instantiateT_frees instT) parm_types) ~~ insts'') |> map_filter (fn (v, t) => if Free v = t then NONE else SOME (v, cert t)); in (Element.instantiate_normalize_morphism (map (apsnd certT) instT, cert_inst) $> Morphism.binding_morphism "Expression.inst" (Binding.prefix mandatory prfx), ctxt') end; (*** Locale processing ***) (** Parsing **) fun parse_elem prep_typ prep_term ctxt = Element.map_ctxt {binding = I, typ = prep_typ ctxt, term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt), pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt), fact = I, attrib = I}; fun prepare_stmt prep_prop prep_obtains ctxt stmt = (case stmt of Element.Shows raw_shows => raw_shows |> (map o apsnd o map) (fn (t, ps) => (prep_prop (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t, map (prep_prop (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) | Element.Obtains raw_obtains => let val ((_, thesis), thesis_ctxt) = Obtain.obtain_thesis ctxt; val obtains = prep_obtains thesis_ctxt thesis raw_obtains; in map (fn (b, t) => ((b, []), [(t, [])])) obtains end); (** Simultaneous type inference: instantiations + elements + statement **) local fun mk_type T = (Logic.mk_type T, []); fun mk_term t = (t, []); fun mk_propp (p, pats) = (Type.constraint propT p, pats); fun dest_type (T, []) = Logic.dest_type T; fun dest_term (t, []) = t; fun dest_propp (p, pats) = (p, pats); fun extract_inst (_, (_, ts)) = map mk_term ts; fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs)); fun extract_eqns es = map (mk_term o snd) es; fun restore_eqns (es, cs) = map2 (fn (b, _) => fn c => (b, dest_term c)) es cs; fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs | extract_elem (Notes _) = [] | extract_elem (Lazy_Notes _) = []; fun restore_elem (Fixes fixes, css) = (fixes ~~ css) |> map (fn ((x, _, mx), cs) => (x, cs |> map dest_type |> try hd, mx)) |> Fixes | restore_elem (Constrains csts, css) = (csts ~~ css) |> map (fn ((x, _), cs) => (x, cs |> map dest_type |> hd)) |> Constrains | restore_elem (Assumes asms, css) = (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes | restore_elem (Defines defs, css) = (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines | restore_elem (elem as Notes _, _) = elem | restore_elem (elem as Lazy_Notes _, _) = elem; fun prep (_, pats) (ctxt, t :: ts) = let val ctxt' = Proof_Context.augment t ctxt in ((t, Syntax.check_props (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats), (ctxt', ts)) end; fun check cs ctxt = let val (cs', (ctxt', _)) = fold_map prep cs (ctxt, Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) (map fst cs)); in (cs', ctxt') end; in fun check_autofix insts eqnss elems concl ctxt = let val inst_cs = map extract_inst insts; val eqns_cs = map extract_eqns eqnss; val elem_css = map extract_elem elems; val concl_cs = (map o map) mk_propp (map snd concl); (* Type inference *) val (inst_cs' :: eqns_cs' :: css', ctxt') = (fold_burrow o fold_burrow) check (inst_cs :: eqns_cs :: elem_css @ [concl_cs]) ctxt; val (elem_css', [concl_cs']) = chop (length elem_css) css'; in ((map restore_inst (insts ~~ inst_cs'), map restore_eqns (eqnss ~~ eqns_cs'), map restore_elem (elems ~~ elem_css'), map fst concl ~~ concl_cs'), ctxt') end; end; (** Prepare locale elements **) fun declare_elem prep_var (Fixes fixes) ctxt = let val (vars, _) = fold_map prep_var fixes ctxt in ctxt |> Proof_Context.add_fixes vars |> snd end | declare_elem prep_var (Constrains csts) ctxt = ctxt |> fold_map (fn (x, T) => prep_var (Binding.name x, SOME T, NoSyn)) csts |> snd | declare_elem _ (Assumes _) ctxt = ctxt | declare_elem _ (Defines _) ctxt = ctxt | declare_elem _ (Notes _) ctxt = ctxt | declare_elem _ (Lazy_Notes _) ctxt = ctxt; (** Finish locale elements **) fun finish_inst ctxt (loc, (prfx, inst)) = let val thy = Proof_Context.theory_of ctxt; val (morph, _) = inst_morphism (map #1 (Locale.params_of thy loc)) (prfx, inst) ctxt; in (loc, morph) end; fun finish_fixes (parms: (string * typ) list) = map (fn (binding, _, mx) => let val x = Binding.name_of binding in (binding, AList.lookup (op =) parms x, mx) end); local fun closeup _ _ false elem = elem | closeup (outer_ctxt, ctxt) parms true elem = let (* FIXME consider closing in syntactic phase -- before type checking *) fun close_frees t = let val rev_frees = Term.fold_aterms (fn Free (x, T) => if Variable.is_fixed outer_ctxt x orelse AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t []; in fold (Logic.all o Free) rev_frees t end; fun no_binds [] = [] | no_binds _ = error "Illegal term bindings in context element"; in (case elem of Assumes asms => Assumes (asms |> map (fn (a, propps) => (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) => let val ((c, _), t') = Local_Defs.cert_def ctxt (K []) (close_frees t) in ((Thm.def_binding_optional (Binding.name c) name, atts), (t', no_binds ps)) end)) | e => e) end; in fun finish_elem _ parms _ (Fixes fixes) = Fixes (finish_fixes parms fixes) | finish_elem _ _ _ (Constrains _) = Constrains [] | finish_elem ctxts parms do_close (Assumes asms) = closeup ctxts parms do_close (Assumes asms) | finish_elem ctxts parms do_close (Defines defs) = closeup ctxts parms do_close (Defines defs) | finish_elem _ _ _ (elem as Notes _) = elem | finish_elem _ _ _ (elem as Lazy_Notes _) = elem; end; (** Process full context statement: instantiations + elements + statement **) (* Interleave incremental parsing and type inference over entire parsed stretch. *) local fun abs_def ctxt = Thm.cterm_of ctxt #> Assumption.assume ctxt #> Local_Defs.abs_def_rule ctxt #> Thm.prop_of; fun prep_full_context_statement parse_typ parse_prop prep_obtains prep_var_elem prep_inst prep_eqns prep_attr prep_var_inst prep_expr {strict, do_close, fixed_frees} raw_import init_body raw_elems raw_stmt ctxt1 = let val thy = Proof_Context.theory_of ctxt1; val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import); fun prep_insts_cumulative (loc, (prfx, (inst, eqns))) (i, insts, eqnss, ctxt) = let val params = map #1 (Locale.params_of thy loc); val inst' = prep_inst ctxt (map #1 params) inst; val parm_types' = params |> map (#2 #> Logic.varifyT_global #> Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) #> Type_Infer.paramify_vars); val inst'' = map2 Type.constraint parm_types' inst'; val insts' = insts @ [(loc, (prfx, inst''))]; val ((insts'', _, _, _), ctxt2) = check_autofix insts' [] [] [] ctxt; val inst''' = insts'' |> List.last |> snd |> snd; val (inst_morph, _) = inst_morphism params (prfx, inst''') ctxt; val ctxt' = Locale.activate_declarations (loc, inst_morph) ctxt2 handle ERROR msg => if null eqns then error msg else (Locale.tracing ctxt1 (msg ^ "\nFalling back to reading rewrites clause before activation."); ctxt2); val attrss = map (apsnd (map (prep_attr ctxt)) o fst) eqns; val eqns' = (prep_eqns ctxt' o map snd) eqns; val eqnss' = [attrss ~~ eqns']; val ((_, [eqns''], _, _), _) = check_autofix insts'' eqnss' [] [] ctxt'; val rewrite_morph = eqns' |> map (abs_def ctxt') |> Variable.export_terms ctxt' ctxt |> Element.eq_term_morphism (Proof_Context.theory_of ctxt) |> the_default Morphism.identity; val ctxt'' = Locale.activate_declarations (loc, inst_morph $> rewrite_morph) ctxt; val eqnss' = eqnss @ [attrss ~~ Variable.export_terms ctxt' ctxt eqns']; in (i + 1, insts', eqnss', ctxt'') end; fun prep_elem raw_elem ctxt = let val ctxt' = ctxt |> Context_Position.set_visible false |> declare_elem prep_var_elem raw_elem |> Context_Position.restore_visible ctxt; val elems' = parse_elem parse_typ parse_prop ctxt' raw_elem; in (elems', ctxt') end; val fors = fold_map prep_var_inst fixed ctxt1 |> fst; val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd; val (_, insts', eqnss', ctxt3) = fold prep_insts_cumulative raw_insts (0, [], [], ctxt2); fun prep_stmt elems ctxt = check_autofix insts' [] elems (prepare_stmt parse_prop prep_obtains ctxt raw_stmt) ctxt; val _ = if fixed_frees then () else (case fold (fold (Variable.add_frees ctxt3) o snd o snd) insts' [] of [] => () | frees => error ("Illegal free variables in expression: " ^ commas_quote (map (Syntax.string_of_term ctxt3 o Free) (rev frees)))); val ((insts, _, elems', concl), ctxt4) = ctxt3 |> init_body |> fold_map prep_elem raw_elems |-> prep_stmt; (* parameters from expression and elements *) val xs = maps (fn Fixes fixes => map (Variable.check_name o #1) fixes | _ => []) (Fixes fors :: elems'); val (parms, ctxt5) = fold_map Proof_Context.inferred_param xs ctxt4; val fors' = finish_fixes parms fors; val fixed = map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fors'; val deps = map (finish_inst ctxt5) insts; val elems'' = map (finish_elem (ctxt1, ctxt5) parms do_close) elems'; in ((fixed, deps, eqnss', elems'', concl), (parms, ctxt5)) end; in fun cert_full_context_statement x = prep_full_context_statement (K I) (K I) Obtain.cert_obtains Proof_Context.cert_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; fun cert_read_full_context_statement x = prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains Proof_Context.read_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; fun read_full_context_statement x = prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains Proof_Context.read_var parse_inst Syntax.read_props Attrib.check_src Proof_Context.read_var check_expr x; end; (* Context statement: elements + statement *) local fun prep_statement prep activate raw_elems raw_stmt ctxt = let val ((_, _, _, elems, concl), _) = prep {strict = true, do_close = false, fixed_frees = true} ([], []) I raw_elems raw_stmt ctxt; val ctxt' = ctxt |> Proof_Context.set_stmt true |> fold_map activate elems |> #2 |> Proof_Context.restore_stmt ctxt; in (concl, ctxt') end; in fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x; fun read_statement x = prep_statement read_full_context_statement Element.activate x; end; (* Locale declaration: import + elements *) fun fix_params params = Proof_Context.add_fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params) #> snd; local fun prep_declaration prep activate raw_import init_body raw_elems ctxt = let val ((fixed, deps, eqnss, elems, _), (parms, ctxt0)) = prep {strict = false, do_close = true, fixed_frees = false} raw_import init_body raw_elems (Element.Shows []) ctxt; val _ = null (flat eqnss) orelse error "Illegal rewrites clause(s) in declaration of locale"; (* Declare parameters and imported facts *) val ctxt' = ctxt |> fix_params fixed |> fold (Context.proof_map o Locale.activate_facts NONE) deps; val (elems', ctxt'') = ctxt' |> Proof_Context.set_stmt true |> fold_map activate elems ||> Proof_Context.restore_stmt ctxt'; in ((fixed, deps, elems', ctxt''), (parms, ctxt0)) end; in fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x; fun cert_read_declaration x = prep_declaration cert_read_full_context_statement Element.activate x; fun read_declaration x = prep_declaration read_full_context_statement Element.activate x; end; (* Locale expression to set up a goal *) local fun props_of thy (name, morph) = let val (asm, defs) = Locale.specification_of thy name in map (Morphism.term morph) (the_list asm @ defs) end; fun prep_goal_expression prep expression ctxt = let val thy = Proof_Context.theory_of ctxt; val ((fixed, deps, eqnss, _, _), _) = prep {strict = true, do_close = true, fixed_frees = true} expression I [] (Element.Shows []) ctxt; (* proof obligations *) val propss = map (props_of thy) deps; val eq_propss = (map o map) snd eqnss; val goal_ctxt = ctxt |> fix_params fixed |> (fold o fold) Proof_Context.augment (propss @ eq_propss); val export = Proof_Context.export_morphism goal_ctxt ctxt; val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export; val exp_term = Term_Subst.zero_var_indexes o Morphism.term export; val exp_typ = Logic.type_map exp_term; val export' = Morphism.morphism "Expression.prep_goal" {binding = [], typ = [exp_typ], term = [exp_term], fact = [exp_fact]}; in ((propss, eq_propss, deps, eqnss, export'), goal_ctxt) end; in fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x; fun read_goal_expression x = prep_goal_expression read_full_context_statement x; end; (*** Locale declarations ***) (* extract specification text *) val norm_term = Envir.beta_norm oo Term.subst_atomic; fun bind_def ctxt eq (xs, env, eqs) = let val _ = Local_Defs.cert_def ctxt (K []) eq; val ((y, T), b) = Local_Defs.abs_def eq; val b' = norm_term env b; fun err msg = error (msg ^ ": " ^ quote y); in (case filter (fn (Free (y', _), _) => y = y' | _ => false) env of [] => (Term.add_frees b' xs, (Free (y, T), b') :: env, eq :: eqs) | dups => if forall (fn (_, b'') => b' aconv b'') dups then (xs, env, eqs) else err "Attempt to redefine variable") end; (* text has the following structure: (((exts, exts'), (ints, ints')), (xs, env, defs)) where exts: external assumptions (terms in assumes elements) exts': dito, normalised wrt. env ints: internal assumptions (terms in assumptions from insts) ints': dito, normalised wrt. env xs: the free variables in exts' and ints' and rhss of definitions, this includes parameters except defined parameters env: list of term pairs encoding substitutions, where the first term is a free variable; substitutions represent defines elements and the rhs is normalised wrt. the previous env defs: the equations from the defines elements *) fun eval_text _ _ (Fixes _) text = text | eval_text _ _ (Constrains _) text = text | eval_text _ is_ext (Assumes asms) (((exts, exts'), (ints, ints')), (xs, env, defs)) = let val ts = maps (map #1 o #2) asms; val ts' = map (norm_term env) ts; val spec' = if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) else ((exts, exts'), (ints @ ts, ints' @ ts')); in (spec', (fold Term.add_frees ts' xs, env, defs)) end | eval_text ctxt _ (Defines defs) (spec, binds) = (spec, fold (bind_def ctxt o #1 o #2) defs binds) | eval_text _ _ (Notes _) text = text | eval_text _ _ (Lazy_Notes _) text = text; fun eval_inst ctxt (loc, morph) text = let val thy = Proof_Context.theory_of ctxt; val (asm, defs) = Locale.specification_of thy loc; val asm' = Option.map (Morphism.term morph) asm; val defs' = map (Morphism.term morph) defs; val text' = text |> (if is_some asm then eval_text ctxt false (Assumes [(Binding.empty_atts, [(the asm', [])])]) else I) |> (if not (null defs) then eval_text ctxt false (Defines (map (fn def => (Binding.empty_atts, (def, []))) defs')) else I) (* FIXME clone from locale.ML *) in text' end; fun eval_elem ctxt elem text = eval_text ctxt true elem text; fun eval ctxt deps elems = let val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [], [])); val ((spec, (_, _, defs))) = fold (eval_elem ctxt) elems text'; in (spec, defs) end; (* axiomsN: name of theorem set with destruct rules for locale predicates, also name suffix of delta predicates and assumptions. *) val axiomsN = "axioms"; local (* introN: name of theorems for introduction rules of locale and delta predicates *) val introN = "intro"; fun atomize_spec ctxt ts = let val t = Logic.mk_conjunction_balanced ts; val body = Object_Logic.atomize_term ctxt t; val bodyT = Term.fastype_of body; in if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of ctxt t)) else (body, bodyT, Object_Logic.atomize ctxt (Thm.cterm_of ctxt t)) end; (* achieve plain syntax for locale predicates (without "PROP") *) fun aprop_tr' n c = let val c' = Lexicon.mark_const c; fun tr' (_: Proof.context) T args = if T <> dummyT andalso length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.const c', args) else raise Match; in (c', tr') end; (* define one predicate including its intro rule and axioms - binding: predicate name - parms: locale parameters - defs: thms representing substitutions from defines elements - ts: terms representing locale assumptions (not normalised wrt. defs) - norm_ts: terms representing locale assumptions (normalised wrt. defs) - thy: the theory *) fun def_pred binding parms defs ts norm_ts thy = let val name = Sign.full_name thy binding; val thy_ctxt = Proof_Context.init_global thy; val (body, bodyT, body_eq) = atomize_spec thy_ctxt norm_ts; val env = Term.add_free_names body []; val xs = filter (member (op =) env o #1) parms; val Ts = map #2 xs; val extraTs = (subtract (op =) (fold Term.add_tfreesT Ts []) (Term.add_tfrees body [])) |> sort_by #1 |> map TFree; val predT = map Term.itselfT extraTs ---> Ts ---> bodyT; val args = map Logic.mk_type extraTs @ map Free xs; val head = Term.list_comb (Const (name, predT), args); val statement = Object_Logic.ensure_propT thy_ctxt head; val ([pred_def], defs_thy) = thy |> bodyT = propT ? Sign.typed_print_translation [aprop_tr' (length args) name] |> Sign.declare_const_global ((binding, predT), NoSyn) |> snd |> Global_Theory.add_defs false [((Thm.def_binding binding, Logic.mk_equals (head, body)), [])]; val defs_ctxt = Proof_Context.init_global defs_thy |> Variable.declare_term head; val intro = Goal.prove_global defs_thy [] norm_ts statement (fn {context = ctxt, ...} => rewrite_goals_tac ctxt [pred_def] THEN compose_tac defs_ctxt (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN compose_tac defs_ctxt (false, - Conjunction.intr_balanced (map (Thm.assume o Thm.cterm_of defs_ctxt) norm_ts), 0) 1) - |> tap (Thm.expose_proof defs_thy); + Conjunction.intr_balanced (map (Thm.assume o Thm.cterm_of defs_ctxt) norm_ts), 0) 1); val conjuncts = (Drule.equal_elim_rule2 OF [body_eq, rewrite_rule defs_ctxt [pred_def] (Thm.assume (Thm.cterm_of defs_ctxt statement))]) |> Conjunction.elim_balanced (length ts); val (_, axioms_ctxt) = defs_ctxt |> Assumption.add_assumes (maps Thm.chyps_of (defs @ conjuncts)); val axioms = ts ~~ conjuncts |> map (fn (t, ax) => Element.prove_witness axioms_ctxt t (rewrite_goals_tac axioms_ctxt defs THEN compose_tac axioms_ctxt (false, ax, 0) 1)); in ((statement, intro, axioms), defs_thy) end; in (* main predicate definition function *) fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = let val ctxt = Proof_Context.init_global thy; val defs' = map (Thm.cterm_of ctxt #> Assumption.assume ctxt #> Drule.abs_def) defs; val (a_pred, a_intro, a_axioms, thy'') = if null exts then (NONE, NONE, [], thy) else let val abinding = if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding; val ((statement, intro, axioms), thy') = thy |> def_pred abinding parms defs' exts exts'; val (_, thy'') = thy' |> Sign.qualified_path true abinding |> Global_Theory.note_thms "" ((Binding.name introN, []), [([intro], [Locale.unfold_add])]) ||> Sign.restore_naming thy'; in (SOME statement, SOME intro, axioms, thy'') end; val (b_pred, b_intro, b_axioms, thy'''') = if null ints then (NONE, NONE, [], thy'') else let val ((statement, intro, axioms), thy''') = thy'' |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); val ctxt''' = Proof_Context.init_global thy'''; val (_, thy'''') = thy''' |> Sign.qualified_path true binding |> Global_Theory.note_thmss "" [((Binding.name introN, []), [([intro], [Locale.intro_add])]), ((Binding.name axiomsN, []), [(map (Drule.export_without_context o Element.conclude_witness ctxt''') axioms, [])])] ||> Sign.restore_naming thy'''; in (SOME statement, SOME intro, axioms, thy'''') end; in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end; end; local fun assumes_to_notes (Assumes asms) axms = fold_map (fn (a, spec) => fn axs => let val (ps, qs) = chop (length spec) axs in ((a, [(ps, [])]), qs) end) asms axms |> apfst (curry Notes "") | assumes_to_notes e axms = (e, axms); fun defines_to_notes ctxt (Defines defs) = Notes ("", map (fn (a, (def, _)) => (a, [([Assumption.assume ctxt (Thm.cterm_of ctxt def)], [(Attrib.internal o K) Locale.witness_add])])) defs) | defines_to_notes _ e = e; val is_hyp = fn Assumes _ => true | Defines _ => true | _ => false; fun gen_add_locale prep_decl binding raw_predicate_binding raw_import raw_body thy = let val name = Sign.full_name thy binding; val _ = Locale.defined thy name andalso error ("Duplicate definition of locale " ^ quote name); val ((fixed, deps, body_elems, _), (parms, ctxt')) = prep_decl raw_import I raw_body (Proof_Context.init_global thy); val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems; val extraTs = subtract (op =) (fold Term.add_tfreesT (map snd parms) []) (fold Term.add_tfrees exts' []); val _ = if null extraTs then () else warning ("Additional type variable(s) in locale specification " ^ Binding.print binding ^ ": " ^ commas (map (Syntax.string_of_typ ctxt' o TFree) (sort_by #1 extraTs))); val predicate_binding = if Binding.is_empty raw_predicate_binding then binding else raw_predicate_binding; val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') = define_preds predicate_binding parms text thy; val pred_ctxt = Proof_Context.init_global thy'; val a_satisfy = Element.satisfy_morphism a_axioms; val b_satisfy = Element.satisfy_morphism b_axioms; val params = fixed @ maps (fn Fixes fixes => map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fixes | _ => []) body_elems; val asm = if is_some b_statement then b_statement else a_statement; val hyp_spec = filter is_hyp body_elems; val notes = if is_some asm then [("", [((Binding.suffix_name ("_" ^ axiomsN) binding, []), [([Assumption.assume pred_ctxt (Thm.cterm_of pred_ctxt (the asm))], [(Attrib.internal o K) Locale.witness_add])])])] else []; val notes' = body_elems |> map (defines_to_notes pred_ctxt) |> map (Element.transform_ctxt a_satisfy) |> (fn elems => fold_map assumes_to_notes elems (map (Element.conclude_witness pred_ctxt) a_axioms)) |> fst |> map (Element.transform_ctxt b_satisfy) |> map_filter (fn Notes notes => SOME notes | _ => NONE); val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps; val axioms = map (Element.conclude_witness pred_ctxt) b_axioms; val loc_ctxt = thy' |> Locale.register_locale binding (extraTs, params) (asm, rev defs) (a_intro, b_intro) axioms hyp_spec [] (rev notes) (rev deps') |> Named_Target.init name |> fold (fn (kind, facts) => Local_Theory.notes_kind kind facts #> snd) notes'; in (name, loc_ctxt) end; in val add_locale = gen_add_locale cert_declaration; val add_locale_cmd = gen_add_locale read_declaration; end; (** Print the instances that would be activated by an interpretation of the expression in the current context (clean = false) or in an empty context (clean = true). **) fun pretty_dependencies ctxt clean expression = let val ((_, _, deps, _, export), expr_ctxt) = read_goal_expression expression ctxt; val export' = if clean then Morphism.identity else export; in Locale.pretty_dependencies expr_ctxt clean export' deps end; 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,408 +1,406 @@ (* Title: Pure/Thy/export_theory.ML Author: Makarius Export foundational theory content and locale/class structure. *) signature EXPORT_THEORY = sig val setup_presentation: (Thy_Info.presentation_context -> theory -> unit) -> unit 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); (* spec rules *) fun primrec_types ctxt const = Spec_Rules.retrieve ctxt (Const const) |> get_first (fn (Spec_Rules.Equational (Spec_Rules.Primrec types), _) => SOME (types, false) | (Spec_Rules.Equational (Spec_Rules.Primcorec types), _) => SOME (types, true) | _ => NONE) |> the_default ([], false); (* locales content *) 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); (* general setup *) fun setup_presentation f = Theory.setup (Thy_Info.add_presentation (fn context => fn thy => if Options.bool (#options context) "export_theory" then f context thy else ())); fun export_body thy name body = if XML.is_empty_body body then () else Export.export thy (Path.binding0 (Path.make ["theory", name])) body; (* presentation *) val _ = setup_presentation (fn {adjust_pos, ...} => fn thy => let val parents = Theory.parents_of thy; val rep_tsig = Type.rep_tsig (Sign.tsig_of thy); val thy_ctxt = Proof_Context.init_global thy; (* entities *) fun make_entity_markup name xname pos serial = let val props = Position.offset_properties_of (adjust_pos pos) @ Position.id_properties_of 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 export get_space decls = let val parent_spaces = map get_space parents; val space = get_space thy; in (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 body => cons (#serial (Name_Space.the_entry space name), XML.Elem (entity_markup space name, body)))) |> 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; fun export_type c (Type.LogicalType n) = SOME (encode_type (get_syntax_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE)) | export_type c (Type.Abbreviation (args, U, false)) = SOME (encode_type (get_syntax_type thy_ctxt c, args, SOME U)) | export_type _ _ = NONE; val _ = export_entities "types" export_type Sign.type_space (Name_Space.dest_table (#types rep_tsig)); (* consts *) val consts = Sign.consts_of thy; 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) (pair bool (pair (list string) bool))))) end; fun export_const c (T, abbrev) = let val syntax = get_syntax_const thy_ctxt c; val U = Logic.unvarifyT_global T; val U0 = Type.strip_sorts U; val recursion = primrec_types thy_ctxt (c, U); val abbrev' = abbrev |> Option.map (Proofterm.standard_vars_term Name.context #> map_types Type.strip_sorts); 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, recursion))))) end; val _ = export_entities "consts" (SOME oo export_const) Sign.const_space (#constants (Consts.dest consts)); (* 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; 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" (K (SOME o encode_axiom Name.context)) Theory.axiom_space (Theory.all_axioms_of thy); (* theorems and proof terms *) val clean_thm = Thm.check_hyps (Context.Theory thy) #> 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 = clean_thm (Thm.unconstrainT raw_thm); val proof0 = if Proofterm.export_standard_enabled () then Thm.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 Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm) (SOME proof0); - val _ = - if Proofterm.export_proof_boxes_required thy - then Proofterm.export_proof_boxes [proof] else (); + 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 thm = Global_Theory.get_thm_name thy (thm_name, Position.none); in XML.Elem (markup, encode_thm thm_id thm) 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; fun export_class name = (case try (Axclass.get_info thy) name of NONE => ([], []) | SOME {params, axioms, ...} => (params, map (Thm.plain_prop_of o clean_thm) axioms)) |> encode_class |> SOME; val _ = export_entities "classes" (fn name => fn () => export_class name) Sign.class_space (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig))))); (* sort algebra *) local 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; in 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)); end; val _ = if null classrel then () else export_body thy "classrel" (export_classrel classrel); val _ = if null arities then () else export_body thy "arities" (export_arities arities); (* 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; fun export_locale loc = 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)); val _ = export_entities "locales" (fn loc => fn () => SOME (export_locale loc)) Locale.locale_space (get_locales thy); (* 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 _ = 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"; (* constdefs *) val constdefs = Defs.dest_constdefs (map Theory.defs_of (Theory.parents_of thy)) (Theory.defs_of thy) |> sort_by #1; val encode_constdefs = let open XML.Encode in list (pair string string) end; val _ = if null constdefs then () else export_body thy "constdefs" (encode_constdefs constdefs); (* parents *) val _ = Export.export thy \<^path_binding>\theory/parents\ (XML.Encode.string (cat_lines (map Context.theory_long_name parents))); in () end); end; diff --git a/src/Pure/global_theory.ML b/src/Pure/global_theory.ML --- a/src/Pure/global_theory.ML +++ b/src/Pure/global_theory.ML @@ -1,369 +1,365 @@ (* Title: Pure/global_theory.ML Author: Makarius Global theory content: stored facts. *) signature GLOBAL_THEORY = sig val facts_of: theory -> Facts.T val check_fact: theory -> xstring * Position.T -> string val intern_fact: theory -> xstring -> string val defined_fact: theory -> string -> bool val alias_fact: binding -> string -> theory -> theory val hide_fact: bool -> string -> theory -> theory val dest_thms: bool -> theory list -> theory -> (Thm_Name.T * thm) list val dest_thm_names: theory -> (Proofterm.thm_id * Thm_Name.T) list val lookup_thm_id: theory -> Proofterm.thm_id -> Thm_Name.T option val lookup_thm: theory -> thm -> (Proofterm.thm_id * Thm_Name.T) option val get_thms: theory -> xstring -> thm list val get_thm: theory -> xstring -> thm val transfer_theories: theory -> thm -> thm val all_thms_of: theory -> bool -> (string * thm) list val get_thm_name: theory -> Thm_Name.T * Position.T -> thm val burrow_fact: ('a list -> 'b list) -> ('a list * 'c) list -> ('b list * 'c) list val burrow_facts: ('a list -> 'b list) -> ('c * ('a list * 'd) list) list -> ('c * ('b list * 'd) list) list val name_multi: string * Position.T -> thm list -> ((string * Position.T) * thm) list type name_flags val unnamed: name_flags val official1: name_flags val official2: name_flags val unofficial1: name_flags val unofficial2: name_flags val name_thm: name_flags -> string * Position.T -> thm -> thm val name_thms: name_flags -> string * Position.T -> thm list -> thm list val check_thms_lazy: thm list lazy -> thm list lazy val add_thms_lazy: string -> (binding * thm list lazy) -> theory -> theory val store_thm: binding * thm -> theory -> thm * theory val store_thm_open: binding * thm -> theory -> thm * theory val add_thms: ((binding * thm) * attribute list) list -> theory -> thm list * theory val add_thm: (binding * thm) * attribute list -> theory -> thm * theory val add_thmss: ((binding * thm list) * attribute list) list -> theory -> thm list list * theory val add_thms_dynamic': Context.generic -> binding * (Context.generic -> thm list) -> theory -> string * theory val add_thms_dynamic: binding * (Context.generic -> thm list) -> theory -> theory val note_thms: string -> Thm.binding * (thm list * attribute list) list -> theory -> (string * thm list) * theory val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list -> theory -> (string * thm list) list * theory val add_defs: bool -> ((binding * term) * attribute list) list -> theory -> thm list * theory val add_defs_unchecked: bool -> ((binding * term) * attribute list) list -> theory -> thm list * theory end; structure Global_Theory: GLOBAL_THEORY = struct (** theory data **) structure Data = Theory_Data ( type T = Facts.T * Thm_Name.T Inttab.table lazy option; val empty: T = (Facts.empty, NONE); fun extend (facts, _) = (facts, NONE); fun merge ((facts1, _), (facts2, _)) = (Facts.merge (facts1, facts2), NONE); ); (* facts *) val facts_of = #1 o Data.get; val map_facts = Data.map o apfst; fun check_fact thy = Facts.check (Context.Theory thy) (facts_of thy); val intern_fact = Facts.intern o facts_of; val defined_fact = Facts.defined o facts_of; fun alias_fact binding name thy = map_facts (Facts.alias (Sign.naming_of thy) binding name) thy; fun hide_fact fully name = map_facts (Facts.hide fully name); fun dest_thms verbose prev_thys thy = Facts.dest_static verbose (map facts_of prev_thys) (facts_of thy) |> maps (uncurry Thm_Name.make_list); (* thm_name vs. derivation_id *) val thm_names_of = #2 o Data.get; val map_thm_names = Data.map o apsnd; fun make_thm_names thy = (dest_thms true (Theory.parents_of thy) thy, Inttab.empty) |-> fold (fn (thm_name, thm) => fn thm_names => (case Thm.derivation_id (Thm.transfer thy thm) of NONE => thm_names | SOME {serial, theory_name} => if Context.theory_long_name thy <> theory_name then raise THM ("Bad theory name for derivation", 0, [thm]) else (case Inttab.lookup thm_names serial of SOME thm_name' => raise THM ("Duplicate use of derivation identity for " ^ Thm_Name.print thm_name ^ " vs. " ^ Thm_Name.print thm_name', 0, [thm]) | NONE => Inttab.update (serial, thm_name) thm_names))); fun get_thm_names thy = (case thm_names_of thy of NONE => Lazy.lazy (fn () => make_thm_names thy) | SOME lazy_tab => lazy_tab); fun dest_thm_names thy = let val theory_name = Context.theory_long_name thy; fun thm_id i = Proofterm.make_thm_id (i, theory_name); val thm_names = Lazy.force (get_thm_names thy); in Inttab.fold_rev (fn (i, thm_name) => cons (thm_id i, thm_name)) thm_names [] end; fun lookup_thm_id thy = let val theories = fold (fn thy' => Symtab.update (Context.theory_long_name thy', get_thm_names thy')) (Theory.nodes_of thy) Symtab.empty; fun lookup (thm_id: Proofterm.thm_id) = (case Symtab.lookup theories (#theory_name thm_id) of NONE => NONE | SOME lazy_tab => Inttab.lookup (Lazy.force lazy_tab) (#serial thm_id)); in lookup end; fun lookup_thm thy = let val lookup = lookup_thm_id thy in fn thm => (case Thm.derivation_id thm of NONE => NONE | SOME thm_id => (case lookup thm_id of NONE => NONE | SOME thm_name => SOME (thm_id, thm_name))) end; val _ = Theory.setup (Theory.at_end (fn thy => if is_some (thm_names_of thy) then NONE else let val lazy_tab = if Future.proofs_enabled 1 then Lazy.lazy (fn () => make_thm_names thy) else Lazy.value (make_thm_names thy); in SOME (map_thm_names (K (SOME lazy_tab)) thy) end)); (* retrieve theorems *) fun get_thms thy xname = #thms (Facts.retrieve (Context.Theory thy) (facts_of thy) (xname, Position.none)); fun get_thm thy xname = Facts.the_single (xname, Position.none) (get_thms thy xname); fun transfer_theories thy = let val theories = fold (fn thy' => Symtab.update (Context.theory_name thy', thy')) (Theory.nodes_of thy) Symtab.empty; fun transfer th = Thm.transfer (the_default thy (Symtab.lookup theories (Thm.theory_name th))) th; in transfer end; fun all_thms_of thy verbose = let val transfer = transfer_theories thy; val facts = facts_of thy; fun add (name, ths) = if not verbose andalso Facts.is_concealed facts name then I else append (map (`(Thm.get_name_hint) o transfer) ths); in Facts.fold_static add facts [] end; fun get_thm_name thy ((name, i), pos) = let val facts = facts_of thy; fun print_name () = Facts.markup_extern (Proof_Context.init_global thy) facts name |-> Markup.markup; in (case (Facts.lookup (Context.Theory thy) facts name, i) of (NONE, _) => error ("Undefined fact " ^ quote (print_name ()) ^ Position.here pos) | (SOME {thms = [thm], ...}, 0) => thm | (SOME {thms, ...}, 0) => Facts.err_single (print_name (), pos) thms | (SOME {thms, ...}, _) => if i > 0 andalso i <= length thms then nth thms (i - 1) else Facts.err_selection (print_name (), pos) i thms) |> Thm.transfer thy end; (** store theorems **) (* fact specifications *) fun burrow_fact f = split_list #>> burrow f #> op ~~; fun burrow_facts f = split_list ##> burrow (burrow_fact f) #> op ~~; (* name theorems *) abstype name_flags = No_Name_Flags | Name_Flags of {pre: bool, official: bool} with val unnamed = No_Name_Flags; val official1 = Name_Flags {pre = true, official = true}; val official2 = Name_Flags {pre = false, official = true}; val unofficial1 = Name_Flags {pre = true, official = false}; val unofficial2 = Name_Flags {pre = false, official = false}; fun name_thm name_flags (name, pos) = Thm.solve_constraints #> (fn thm => (case name_flags of No_Name_Flags => thm | Name_Flags {pre, official} => thm |> (official andalso (not pre orelse Thm.raw_derivation_name thm = "")) ? Thm.name_derivation (name, pos) |> (name <> "" andalso (not pre orelse not (Thm.has_name_hint thm))) ? Thm.put_name_hint name)); end; fun name_multi (name, pos) = Thm_Name.make_list name #> (map o apfst) (fn thm_name => (Thm_Name.flatten thm_name, pos)); fun name_thms name_flags name_pos = name_multi name_pos #> map (uncurry (name_thm name_flags)); (* apply theorems and attributes *) fun register_proofs thms thy = (thms, Thm.register_proofs (Lazy.value thms) thy); fun bind_name thy b = (Sign.full_name thy b, Binding.default_pos_of b); fun add_facts (b, fact) thy = let val (full_name, pos) = bind_name thy b; fun check fact = fact |> map_index (fn (i, thm) => let fun err msg = error ("Malformed global fact " ^ quote (full_name ^ (if length fact = 1 then "" else "(" ^ string_of_int (i + 1) ^ ")")) ^ Position.here pos ^ "\n" ^ msg); val prop = Thm.plain_prop_of thm handle THM _ => thm |> Thm.check_hyps (Context.Theory thy) |> Thm.full_prop_of; in ignore (Logic.unvarify_global (Term_Subst.zero_var_indexes prop)) handle TYPE (msg, _, _) => err msg | TERM (msg, _) => err msg | ERROR msg => err msg end); val arg = (b, Lazy.map_finished (tap check) fact); in thy |> map_facts (Facts.add_static (Context.Theory thy) {strict = true, index = false} arg #> #2) end; fun check_thms_lazy (thms: thm list lazy) = if Proofterm.proofs_enabled () orelse Options.default_bool "strict_facts" then Lazy.force_value thms else thms; fun add_thms_lazy kind (b, thms) thy = if Binding.is_empty b then Thm.register_proofs (check_thms_lazy thms) thy else let val name_pos = bind_name thy b; val thms' = check_thms_lazy thms |> Lazy.map_finished (name_thms official1 name_pos #> map (Thm.kind_rule kind)); in thy |> Thm.register_proofs thms' |> add_facts (b, thms') end; val app_facts = apfst flat oo fold_map (fn (thms, atts) => fn thy => fold_map (Thm.theory_attributes atts) (map (Thm.transfer thy) thms) thy); fun apply_facts name_flags1 name_flags2 (b, facts) thy = - if Binding.is_empty b then - let - val (thms, thy') = app_facts facts thy; - val _ = Thm.expose_proofs thy' thms; - in register_proofs thms thy' end + if Binding.is_empty b then app_facts facts thy |-> register_proofs else let val name_pos = bind_name thy b; val (thms', thy') = thy |> app_facts (map (apfst (name_thms name_flags1 name_pos)) facts) |>> name_thms name_flags2 name_pos |-> register_proofs; val thy'' = thy' |> add_facts (b, Lazy.value thms'); in (map (Thm.transfer thy'') thms', thy'') end; (* store_thm *) fun store_thm (b, th) = apply_facts official1 official2 (b, [([th], [])]) #>> the_single; fun store_thm_open (b, th) = apply_facts unofficial1 unofficial2 (b, [([th], [])]) #>> the_single; (* add_thms(s) *) val add_thmss = fold_map (fn ((b, thms), atts) => apply_facts official1 official2 (b, [(thms, atts)])); fun add_thms args = add_thmss (map (apfst (apsnd single)) args) #>> map the_single; val add_thm = yield_singleton add_thms; (* dynamic theorems *) fun add_thms_dynamic' context arg thy = let val (name, facts') = Facts.add_dynamic context arg (facts_of thy) in (name, map_facts (K facts') thy) end; fun add_thms_dynamic arg thy = add_thms_dynamic' (Context.Theory thy) arg thy |> snd; (* note_thmss *) fun note_thms kind ((b, more_atts), facts) thy = let val name = Sign.full_name thy b; val facts' = facts |> map (apsnd (fn atts => surround (Thm.kind kind) (atts @ more_atts))); val (thms', thy') = thy |> apply_facts official1 official2 (b, facts'); in ((name, thms'), thy') end; val note_thmss = fold_map o note_thms; (* old-style defs *) local fun add unchecked overloaded = fold_map (fn ((b, prop), atts) => fn thy => let val context = Defs.global_context thy; val ((_, def), thy') = Thm.add_def context unchecked overloaded (b, prop) thy; val thm = def |> Thm.forall_intr_frees |> Thm.forall_elim_vars 0 |> Thm.varifyT_global; in thy' |> apply_facts unnamed official2 (b, [([thm], atts)]) |>> the_single end); in val add_defs = add false; val add_defs_unchecked = add true; end; end; diff --git a/src/Pure/more_thm.ML b/src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML +++ b/src/Pure/more_thm.ML @@ -1,751 +1,755 @@ (* Title: Pure/more_thm.ML Author: Makarius Further operations on type ctyp/cterm/thm, outside the inference kernel. *) infix aconvc; signature BASIC_THM = sig include BASIC_THM val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T structure Ctermtab: TABLE structure Thmtab: TABLE val aconvc: cterm * cterm -> bool type attribute = Context.generic * thm -> Context.generic option * thm option end; signature THM = sig include THM structure Ctermtab: TABLE structure Thmtab: TABLE val eq_ctyp: ctyp * ctyp -> bool val aconvc: cterm * cterm -> bool val add_tvars: thm -> ctyp list -> ctyp list val add_frees: thm -> cterm list -> cterm list val add_vars: thm -> cterm list -> cterm list val dest_funT: ctyp -> ctyp * ctyp val strip_type: ctyp -> ctyp list * ctyp val all_name: Proof.context -> string * cterm -> cterm -> cterm val all: Proof.context -> cterm -> cterm -> cterm val mk_binop: cterm -> cterm -> cterm -> cterm val dest_binop: cterm -> cterm * cterm val dest_implies: cterm -> cterm * cterm val dest_equals: cterm -> cterm * cterm val dest_equals_lhs: cterm -> cterm val dest_equals_rhs: cterm -> cterm val lhs_of: thm -> cterm val rhs_of: thm -> cterm val fast_term_ord: cterm ord val term_ord: cterm ord val thm_ord: thm ord val cterm_cache: (cterm -> 'a) -> cterm -> 'a val thm_cache: (thm -> 'a) -> thm -> 'a val is_reflexive: thm -> bool val eq_thm: thm * thm -> bool val eq_thm_prop: thm * thm -> bool val eq_thm_strict: thm * thm -> bool val equiv_thm: theory -> thm * thm -> bool val class_triv: theory -> class -> thm val of_sort: ctyp * sort -> thm list val is_dummy: thm -> bool val add_thm: thm -> thm list -> thm list val del_thm: thm -> thm list -> thm list val merge_thms: thm list * thm list -> thm list val full_rules: thm Item_Net.T val intro_rules: thm Item_Net.T val elim_rules: thm Item_Net.T val declare_hyps: cterm -> Proof.context -> Proof.context val assume_hyps: cterm -> Proof.context -> thm * Proof.context val unchecked_hyps: Proof.context -> Proof.context val restore_hyps: Proof.context -> Proof.context -> Proof.context val undeclared_hyps: Context.generic -> thm -> term list val check_hyps: Context.generic -> thm -> thm val declare_term_sorts: term -> Proof.context -> Proof.context val extra_shyps': Proof.context -> thm -> sort list val check_shyps: Proof.context -> thm -> thm val weaken_sorts': Proof.context -> cterm -> cterm val elim_implies: thm -> thm -> thm val forall_intr_name: string * cterm -> thm -> thm val forall_elim_var: int -> thm -> thm val forall_elim_vars: int -> thm -> thm val instantiate_frees: ((string * sort) * ctyp) list * ((string * typ) * cterm) list -> thm -> thm val instantiate': ctyp option list -> cterm option list -> thm -> thm val forall_intr_frees: thm -> thm val unvarify_global: theory -> thm -> thm val unvarify_axiom: theory -> string -> thm val rename_params_rule: string list * int -> thm -> thm val rename_boundvars: term -> term -> thm -> thm val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory val add_axiom_global: binding * term -> theory -> (string * thm) * theory val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory val add_def_global: bool -> bool -> binding * term -> theory -> (string * thm) * theory type attribute = Context.generic * thm -> Context.generic option * thm option type binding = binding * attribute list val tag_rule: string * string -> thm -> thm val untag_rule: string -> thm -> thm val is_free_dummy: thm -> bool val tag_free_dummy: thm -> thm val def_name: string -> string val def_name_optional: string -> string -> string val def_binding: Binding.binding -> Binding.binding val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding val make_def_binding: bool -> Binding.binding -> Binding.binding val has_name_hint: thm -> bool val get_name_hint: thm -> string val put_name_hint: string -> thm -> thm val theoremK: string val legacy_get_kind: thm -> string val kind_rule: string -> thm -> thm val rule_attribute: thm list -> (Context.generic -> thm -> thm) -> attribute val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic val theory_attributes: attribute list -> thm -> theory -> thm * theory val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context val no_attributes: 'a -> 'a * 'b list val simple_fact: 'a -> ('a * 'b list) list val tag: string * string -> attribute val untag: string -> attribute val kind: string -> attribute val reconstruct_proof_of: thm -> Proofterm.proof val standard_proof_of: {full: bool, expand_name: Proofterm.thm_header -> string option} -> thm -> Proofterm.proof val register_proofs: thm list lazy -> theory -> theory val consolidate_theory: theory -> unit val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool} -> thm -> Pretty.T val pretty_thm: Proof.context -> thm -> Pretty.T val pretty_thm_item: Proof.context -> thm -> Pretty.T val pretty_thm_global: theory -> thm -> Pretty.T val string_of_thm: Proof.context -> thm -> string val string_of_thm_global: theory -> thm -> string end; structure Thm: THM = struct (** basic operations **) (* collecting ctyps and cterms *) val eq_ctyp = op = o apply2 Thm.typ_of; val op aconvc = op aconv o apply2 Thm.term_of; val add_tvars = Thm.fold_atomic_ctyps (fn a => is_TVar (Thm.typ_of a) ? insert eq_ctyp a); val add_frees = Thm.fold_atomic_cterms (fn a => is_Free (Thm.term_of a) ? insert (op aconvc) a); val add_vars = Thm.fold_atomic_cterms (fn a => is_Var (Thm.term_of a) ? insert (op aconvc) a); (* ctyp operations *) fun dest_funT cT = (case Thm.typ_of cT of Type ("fun", _) => let val [A, B] = Thm.dest_ctyp cT in (A, B) end | T => raise TYPE ("dest_funT", [T], [])); (* ctyp version of strip_type: maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *) fun strip_type cT = (case Thm.typ_of cT of Type ("fun", _) => let val (cT1, cT2) = dest_funT cT; val (cTs, cT') = strip_type cT2 in (cT1 :: cTs, cT') end | _ => ([], cT)); (* cterm operations *) fun all_name ctxt (x, t) A = let val T = Thm.typ_of_cterm t; val all_const = Thm.cterm_of ctxt (Const ("Pure.all", (T --> propT) --> propT)); in Thm.apply all_const (Thm.lambda_name (x, t) A) end; fun all ctxt t A = all_name ctxt ("", t) A; fun mk_binop c a b = Thm.apply (Thm.apply c a) b; fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); fun dest_implies ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_implies", [Thm.term_of ct])); fun dest_equals ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_equals", [Thm.term_of ct])); fun dest_equals_lhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct])); fun dest_equals_rhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct])); val lhs_of = dest_equals_lhs o Thm.cprop_of; val rhs_of = dest_equals_rhs o Thm.cprop_of; (* certified term order *) val fast_term_ord = Term_Ord.fast_term_ord o apply2 Thm.term_of; val term_ord = Term_Ord.term_ord o apply2 Thm.term_of; (* thm order: ignores theory context! *) val thm_ord = Term_Ord.fast_term_ord o apply2 Thm.prop_of <<< list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 Thm.tpairs_of <<< list_ord Term_Ord.fast_term_ord o apply2 Thm.hyps_of <<< list_ord Term_Ord.sort_ord o apply2 Thm.shyps_of; (* tables and caches *) structure Ctermtab = Table(type key = cterm val ord = fast_term_ord); structure Thmtab = Table(type key = thm val ord = thm_ord); fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f; fun thm_cache f = Cache.create Thmtab.empty Thmtab.lookup Thmtab.update f; (* equality *) fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) handle TERM _ => false; val eq_thm = is_equal o thm_ord; val eq_thm_prop = op aconv o apply2 Thm.full_prop_of; fun eq_thm_strict ths = eq_thm ths andalso Context.eq_thy_id (apply2 Thm.theory_id ths) andalso op = (apply2 Thm.maxidx_of ths) andalso op = (apply2 Thm.get_tags ths); (* pattern equivalence *) fun equiv_thm thy ths = Pattern.equiv thy (apply2 (Thm.full_prop_of o Thm.transfer thy) ths); (* type classes and sorts *) fun class_triv thy c = Thm.of_class (Thm.global_ctyp_of thy (TVar ((Name.aT, 0), [c])), c); fun of_sort (T, S) = map (fn c => Thm.of_class (T, c)) S; (* misc operations *) fun is_dummy thm = (case try Logic.dest_term (Thm.concl_of thm) of NONE => false | SOME t => Term.is_dummy_pattern (Term.head_of t)); (* collections of theorems in canonical order *) val add_thm = update eq_thm_prop; val del_thm = remove eq_thm_prop; val merge_thms = merge eq_thm_prop; val full_rules = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); val intro_rules = Item_Net.init eq_thm_prop (single o Thm.concl_of); val elim_rules = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); (** declared hyps and sort hyps **) structure Hyps = Proof_Data ( type T = {checked_hyps: bool, hyps: Termtab.set, shyps: sort Ord_List.T}; fun init _ : T = {checked_hyps = true, hyps = Termtab.empty, shyps = []}; ); fun map_hyps f = Hyps.map (fn {checked_hyps, hyps, shyps} => let val (checked_hyps', hyps', shyps') = f (checked_hyps, hyps, shyps) in {checked_hyps = checked_hyps', hyps = hyps', shyps = shyps'} end); (* hyps *) fun declare_hyps raw_ct ctxt = ctxt |> map_hyps (fn (checked_hyps, hyps, shyps) => let val ct = Thm.transfer_cterm (Proof_Context.theory_of ctxt) raw_ct; val hyps' = Termtab.update (Thm.term_of ct, ()) hyps; in (checked_hyps, hyps', shyps) end); fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt); val unchecked_hyps = map_hyps (fn (_, hyps, shyps) => (false, hyps, shyps)); fun restore_hyps ctxt = map_hyps (fn (_, hyps, shyps) => (#checked_hyps (Hyps.get ctxt), hyps, shyps)); fun undeclared_hyps context th = Thm.hyps_of th |> filter_out (case context of Context.Theory _ => K false | Context.Proof ctxt => (case Hyps.get ctxt of {checked_hyps = false, ...} => K true | {hyps, ...} => Termtab.defined hyps)); fun check_hyps context th = (case undeclared_hyps context th of [] => th | undeclared => error (Pretty.string_of (Pretty.big_list "Undeclared hyps:" (map (Pretty.item o single o Syntax.pretty_term (Syntax.init_pretty context)) undeclared)))); (* shyps *) fun declare_term_sorts t = map_hyps (fn (checked_hyps, hyps, shyps) => (checked_hyps, hyps, Sorts.insert_term t shyps)); fun extra_shyps' ctxt th = Sorts.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th); fun check_shyps ctxt raw_th = let val th = Thm.strip_shyps raw_th; val extra_shyps = extra_shyps' ctxt th; in if null extra_shyps then th else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" :: Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) extra_shyps)))) end; val weaken_sorts' = Thm.weaken_sorts o #shyps o Hyps.get; (** basic derived rules **) (*Elimination of implication A A \ B ------------ B *) fun elim_implies thA thAB = Thm.implies_elim thAB thA; (* forall_intr_name *) fun forall_intr_name (a, x) th = let val th' = Thm.forall_intr x th; val prop' = (case Thm.prop_of th' of all $ Abs (_, T, b) => all $ Abs (a, T, b)); in Thm.renamed_prop prop' th' end; (* forall_elim_var(s) *) local fun dest_all ct = (case Thm.term_of ct of Const ("Pure.all", _) $ Abs (a, _, _) => let val (x, ct') = Thm.dest_abs NONE (Thm.dest_arg ct) in SOME ((a, Thm.ctyp_of_cterm x), ct') end | _ => NONE); fun dest_all_list ct = (case dest_all ct of NONE => [] | SOME (v, ct') => v :: dest_all_list ct'); fun forall_elim_vars_list vars i th = let val used = (Thm.fold_terms o Term.fold_aterms) (fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I) th []; val vars' = (Name.variant_list used (map #1 vars), vars) |> ListPair.map (fn (x, (_, T)) => Thm.var ((x, i), T)); in fold Thm.forall_elim vars' th end; in fun forall_elim_vars i th = forall_elim_vars_list (dest_all_list (Thm.cprop_of th)) i th; fun forall_elim_var i th = let val vars = (case dest_all (Thm.cprop_of th) of SOME (v, _) => [v] | NONE => raise THM ("forall_elim_var", i, [th])); in forall_elim_vars_list vars i th end; end; (* instantiate frees *) fun instantiate_frees ([], []) th = th | instantiate_frees (instT, inst) th = let val idx = Thm.maxidx_of th + 1; fun index ((a, A), b) = (((a, idx), A), b); val insts = (map index instT, map index inst); val frees = (map (#1 o #1) instT, map (#1 o #1) inst); val hyps = Thm.chyps_of th; val inst_cterm = Thm.generalize_cterm frees idx #> Thm.instantiate_cterm insts; in th |> fold_rev Thm.implies_intr hyps |> Thm.generalize frees idx |> Thm.instantiate insts |> fold (elim_implies o Thm.assume o inst_cterm) hyps end; (* instantiate by left-to-right occurrence of variables *) fun instantiate' cTs cts thm = let fun err msg = raise TYPE ("instantiate': " ^ msg, map_filter (Option.map Thm.typ_of) cTs, map_filter (Option.map Thm.term_of) cts); fun zip_vars xs ys = zip_options xs ys handle ListPair.UnequalLengths => err "more instantiations than variables in thm"; val thm' = Thm.instantiate ((zip_vars (rev (Thm.fold_terms Term.add_tvars thm [])) cTs), []) thm; val thm'' = Thm.instantiate ([], zip_vars (rev (Thm.fold_terms Term.add_vars thm' [])) cts) thm'; in thm'' end; (* forall_intr_frees: generalization over all suitable Free variables *) fun forall_intr_frees th = let val fixed = fold Term.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th) @ Thm.hyps_of th) []; val frees = Thm.fold_atomic_cterms (fn a => (case Thm.term_of a of Free v => not (member (op =) fixed v) ? insert (op aconvc) a | _ => I)) th []; in fold Thm.forall_intr frees th end; (* unvarify_global: global schematic variables *) fun unvarify_global thy th = let val prop = Thm.full_prop_of th; val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th) handle TERM (msg, _) => raise THM (msg, 0, [th]); val instT = rev (Term.add_tvars prop []) |> map (fn v as ((a, _), S) => (v, TFree (a, S))); val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) => let val T' = Term_Subst.instantiateT instT T in (((a, i), T'), Thm.global_cterm_of thy (Free ((a, T')))) end); in Thm.instantiate (map (apsnd (Thm.global_ctyp_of thy)) instT, inst) th end; fun unvarify_axiom thy = unvarify_global thy o Thm.axiom thy; (* user renaming of parameters in a subgoal *) (*The names, if distinct, are used for the innermost parameters of subgoal i; preceding parameters may be renamed to make all parameters distinct.*) fun rename_params_rule (names, i) st = let val (_, Bs, Bi, C) = Thm.dest_state (st, i); val params = map #1 (Logic.strip_params Bi); val short = length params - length names; val names' = if short < 0 then error "More names than parameters in subgoal!" else Name.variant_list names (take short params) @ names; val free_names = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi []; val Bi' = Logic.list_rename_params names' Bi; in (case duplicates (op =) names of a :: _ => (warning ("Can't rename. Bound variables not distinct: " ^ a); st) | [] => (case inter (op =) names free_names of a :: _ => (warning ("Can't rename. Bound/Free variable clash: " ^ a); st) | [] => Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st)) end; (* preservation of bound variable names *) fun rename_boundvars pat obj th = (case Term.rename_abs pat obj (Thm.prop_of th) of NONE => th | SOME prop' => Thm.renamed_prop prop' th); (** specification primitives **) (* rules *) fun stripped_sorts thy t = let val tfrees = rev (Term.add_tfrees t []); val tfrees' = map (fn a => (a, [])) (Name.variant_list [] (map #1 tfrees)); val recover = map2 (fn (a', S') => fn (a, S) => (((a', 0), S'), Thm.global_ctyp_of thy (TVar ((a, 0), S)))) tfrees' tfrees; val strip = map (apply2 TFree) (tfrees ~~ tfrees'); val t' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip))) t; in (strip, recover, t') end; fun add_axiom ctxt (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val (strip, recover, prop') = stripped_sorts thy prop; val constraints = map (fn (TFree (_, S), T) => (T, S)) strip; val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.ctyp_of ctxt T, S)) strip; val thy' = thy |> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop')); val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = Thm.instantiate (recover, []) axm' |> unvarify_global thy' |> fold elim_implies of_sorts; in ((axm_name, thm), thy') end; fun add_axiom_global arg thy = add_axiom (Syntax.init_pretty_global thy) arg thy; fun add_def (context as (ctxt, _)) unchecked overloaded (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val prems = map (Thm.cterm_of ctxt) (Logic.strip_imp_prems prop); val (_, recover, concl') = stripped_sorts thy (Logic.strip_imp_concl prop); val thy' = Theory.add_def context unchecked overloaded (b, concl') thy; val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = Thm.instantiate (recover, []) axm' |> unvarify_global thy' |> fold_rev Thm.implies_intr prems; in ((axm_name, thm), thy') end; fun add_def_global unchecked overloaded arg thy = add_def (Defs.global_context thy) unchecked overloaded arg thy; (** theorem tags **) (* add / delete tags *) fun tag_rule tg = Thm.map_tags (insert (op =) tg); fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s')); (* free dummy thm -- for abstract closure *) val free_dummyN = "free_dummy"; fun is_free_dummy thm = Properties.defined (Thm.get_tags thm) free_dummyN; val tag_free_dummy = tag_rule (free_dummyN, ""); (* def_name *) fun def_name c = c ^ "_def"; fun def_name_optional c "" = def_name c | def_name_optional _ name = name; val def_binding = Binding.map_name def_name #> Binding.reset_pos; fun def_binding_optional b name = if Binding.is_empty name then def_binding b else name; fun make_def_binding cond b = if cond then def_binding b else Binding.empty; (* unofficial theorem names *) fun has_name_hint thm = AList.defined (op =) (Thm.get_tags thm) Markup.nameN; fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN); fun get_name_hint thm = if has_name_hint thm then the_name_hint thm else "??.unknown"; fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name); (* theorem kinds *) val theoremK = "theorem"; fun legacy_get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN); fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN; (** attributes **) (*attributes subsume any kind of rules or context modifiers*) type attribute = Context.generic * thm -> Context.generic option * thm option; type binding = binding * attribute list; fun rule_attribute ths f (x, th) = (NONE, (case find_first is_free_dummy (th :: ths) of SOME th' => SOME th' | NONE => SOME (f x th))); fun declaration_attribute f (x, th) = (if is_free_dummy th then NONE else SOME (f th x), NONE); fun mixed_attribute f (x, th) = let val (x', th') = f (x, th) in (SOME x', SOME th') end; fun apply_attribute (att: attribute) th x = let val (x', th') = att (x, check_hyps x (Thm.transfer'' x th)) in (the_default th th', the_default x x') end; fun attribute_declaration att th x = #2 (apply_attribute att th x); fun apply_attributes mk dest = let fun app [] th x = (th, x) | app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts; in app end; val theory_attributes = apply_attributes Context.Theory Context.the_theory; val proof_attributes = apply_attributes Context.Proof Context.the_proof; fun no_attributes x = (x, []); fun simple_fact x = [(x, [])]; fun tag tg = rule_attribute [] (K (tag_rule tg)); fun untag s = rule_attribute [] (K (untag_rule s)); fun kind k = rule_attribute [] (K (k <> "" ? kind_rule k)); (** proof terms **) fun reconstruct_proof_of thm = Proofterm.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm); fun standard_proof_of {full, expand_name} thm = let val thy = Thm.theory_of_thm thm in reconstruct_proof_of thm |> Proofterm.expand_proof thy expand_name |> Proofterm.rew_proof thy |> Proofterm.no_thm_proofs |> not full ? Proofterm.shrink_proof end; (** forked proofs **) structure Proofs = Theory_Data ( type T = thm list lazy list; val empty = []; fun extend _ = empty; fun merge _ = empty; ); fun register_proofs ths = (Proofs.map o cons) (Lazy.map_finished (map Thm.trim_context) ths); fun consolidate_theory thy = - rev (Proofs.get thy) - |> maps (map (Thm.transfer thy) o Lazy.force) - |> Thm.consolidate; + let + val thms = + rev (Proofs.get thy) + |> maps (map (Thm.transfer thy) o Lazy.force); + val _ = Thm.consolidate thms; + val _ = Thm.expose_proofs thy thms; + in () end; (** print theorems **) (* options *) val show_consts = Config.declare_option_bool ("show_consts", \<^here>); val show_hyps = Config.declare_bool ("show_hyps", \<^here>) (K false); val show_tags = Config.declare_bool ("show_tags", \<^here>) (K false); (* pretty_thm etc. *) fun pretty_tag (name, arg) = Pretty.strs [name, quote arg]; val pretty_tags = Pretty.list "[" "]" o map pretty_tag; fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps'} raw_th = let val show_tags = Config.get ctxt show_tags; val show_hyps = Config.get ctxt show_hyps; val th = raw_th |> perhaps (try (Thm.transfer' ctxt)) |> perhaps (try Thm.strip_shyps); val hyps = if show_hyps then Thm.hyps_of th else undeclared_hyps (Context.Proof ctxt) th; val extra_shyps = extra_shyps' ctxt th; val tags = Thm.get_tags th; val tpairs = Thm.tpairs_of th; val q = if quote then Pretty.quote else I; val prt_term = q o Syntax.pretty_term ctxt; val hlen = length extra_shyps + length hyps + length tpairs; val hsymbs = if hlen = 0 then [] else if show_hyps orelse show_hyps' then [Pretty.brk 2, Pretty.list "[" "]" (map (q o Syntax.pretty_flexpair ctxt) tpairs @ map prt_term hyps @ map (Syntax.pretty_sort ctxt) extra_shyps)] else [Pretty.brk 2, Pretty.str ("[" ^ replicate_string hlen "." ^ "]")]; val tsymbs = if null tags orelse not show_tags then [] else [Pretty.brk 1, pretty_tags tags]; in Pretty.block (prt_term (Thm.prop_of th) :: (hsymbs @ tsymbs)) end; fun pretty_thm ctxt = pretty_thm_raw ctxt {quote = false, show_hyps = true}; fun pretty_thm_item ctxt th = Pretty.item [pretty_thm ctxt th]; fun pretty_thm_global thy = pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false}; val string_of_thm = Pretty.string_of oo pretty_thm; val string_of_thm_global = Pretty.string_of oo pretty_thm_global; open Thm; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Pure/proofterm.ML b/src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML +++ b/src/Pure/proofterm.ML @@ -1,2319 +1,2305 @@ (* 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 | OfClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: (string * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} type oracle = string * 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 consolidate: proof_body list -> unit 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 -> (string * sort) list -> 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: string list * string list -> int -> term -> proof -> proof val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list -> 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 list -> unit + 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 | OfClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: (string * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} and thm_body = - Thm_Body of {export_proof: unit lazy, open_proof: proof -> proof, body: proof_body future} + 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, consolidate: unit lazy}; + body: proof_body future, export: unit lazy, consolidate: unit lazy}; type oracle = string * term option; val oracle_ord = prod_ord fast_string_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}; -val no_export_proof = Lazy.value (); - -fun thm_body body = - Thm_Body {export_proof = no_export_proof, open_proof = I, body = Future.value body}; -fun thm_body_export_proof (Thm_Body {export_proof, ...}) = export_proof; +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 = +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 = - Thm_Node {theory_name = theory_name, name = name, prop = prop, body = body, - consolidate = +fun make_thm_node theory_name name prop body export = + let + val body' = body + |> Options.default_bool "prune_proofs" ? (Future.map o map_proof_of) (K MinProof); + val consolidate = Lazy.lazy_name "Proofterm.make_thm_node" (fn () => - let val PBody {thms, ...} = Future.join body - in consolidate (join_thms thms) end)}; + 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); + (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 {export_proof, open_proof, body})) = + | 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 {export_proof = export_proof, open_proof = open_proof, body = 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 OfClass (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 string (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 OfClass (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) => OfClass (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 string (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)) end; + 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 (OfClass 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) => OfClass (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 (OfClass (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]) (OfClass (_, c)) = OfClass (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 *) fun del_conflicting_tvars envT T = Term_Subst.instantiateT (map_filter (fn ixnS as (_, S) => (Type.lookup envT ixnS; NONE) handle TYPE _ => SOME (ixnS, Logic.dummy_tfree S)) (Term.add_tvarsT T [])) T; fun del_conflicting_vars env t = Term_Subst.instantiate (map_filter (fn ixnS as (_, S) => (Type.lookup (Envir.type_env env) ixnS; NONE) handle TYPE _ => SOME (ixnS, Logic.dummy_tfree S)) (Term.add_tvars t []), map_filter (fn (ixnT as (_, T)) => (Envir.lookup env ixnT; NONE) handle TYPE _ => SOME (ixnT, Free ("dummy", T))) (Term.add_vars t [])) t; 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 (OfClass (T, c)) = OfClass (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; (* 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 = Term.fold_types (Term.fold_atyps (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t []; 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 null frees then [] else fold_aterms (fn Free (x, T) => member (op =) frees x ? insert (op =) (x, T) | _ => I) (Term_Subst.generalize (tfrees, []) idx prop) []; in prf |> Same.commit (map_proof_terms_same (Term_Subst.generalize_same (tfrees, []) 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, map (apsnd 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' _ _ (OfClass (T, c)) = OfClass (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_sorts prf = let fun get S2 (T, S1) = if Sorts.sort_le algebra (S1, S2) then SOME T else NONE; val extra = map (`Logic.dummy_tfree) extra_sorts; val replacements = present @ extra @ witnessed; fun replace T = if exists (fn (T', _) => T' = T) present then raise Same.SAME else (case get_first (get (Type.sort_of_atyp T)) replacements 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 OfClass 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 OfClass _) = ([], 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 (OfClass (T1, c1), OfClass (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 _ _ (OfClass (T, c)) = OfClass (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 | (OfClass (_, c), OfClass (_, 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 %% OfClass _) = 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 OfClass (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 _ (OfClass (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 (map #2 ps @ [body0]); + 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 TVar (("'", idx), Type.sort_of_atyp T) 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_not excl x = ((is_Free x orelse is_Var x) andalso not (member (op =) excl x)) ? insert (op =) x; val visible = fold_aterms (add_not []) term []; val hidden = fold_proof_terms (fold_aterms (add_not 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 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 proofs = +fun export_proof_boxes bodies = let - fun export_boxes (AbsP (_, _, prf)) = export_boxes prf - | export_boxes (Abst (_, _, prf)) = export_boxes prf - | export_boxes (prf1 %% prf2) = export_boxes prf1 #> export_boxes prf2 - | export_boxes (prf % _) = export_boxes prf - | export_boxes (PThm ({serial = i, ...}, thm_body)) = - (fn boxes => - if Inttab.defined boxes i then boxes - else - let - val prf' = thm_body_proof_raw thm_body; - val export = thm_body_export_proof thm_body; - val boxes' = Inttab.update (i, export) boxes; - in export_boxes prf' boxes' end) - | export_boxes _ = I; - val boxes = (proofs, Inttab.empty) |-> fold export_boxes |> Inttab.dest; - in List.app (Lazy.force o #2) boxes end; + 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 = (bodies, Inttab.empty) |-> fold export_body |> 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 export thy i prop prf = - if export_enabled () then - let - val _ = export_proof_boxes [prf]; - val _ = export_proof thy i prop prf; - in () end - else (); - -fun prune proof = - if Options.default_bool "prune_proofs" then MinProof - else proof; - 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 publish i = map_proof_of (rew_proof thy #> tap (export thy i prop1) #> prune); - val open_proof = not named ? rew_proof thy; - 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 ? publish i; + val postproc = map_proof_of (unconstrainT #> named ? rew_proof thy); in (i, fulfill_proof_future thy promises postproc body0) end; val (i, body') = (*non-deterministic, depends on unknown promises*) (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 (publish i)) end + in (i, body' |> ser <> i ? Future.map (map_proof_of (rew_proof thy))) end else new_prf () | _ => new_prf ()); - val export_proof = - if named orelse not (export_enabled ()) then no_export_proof - else + 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 Exn.reraise exn) + else no_export; val theory_name = Context.theory_long_name thy; - val thm = (i, make_thm_node theory_name name prop1 body'); + val thm = (i, make_thm_node theory_name name prop1 body' export); val header = thm_header i ([pos, Position.thread_data ()]) theory_name name prop1 NONE; - val thm_body = Thm_Body {export_proof = export_proof, open_proof = open_proof, body = body'}; + val thm_body = Thm_Body {open_proof = open_proof, body = body'}; val head = PThm (header, 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 OfClass (#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) (fold boxes_of proofs Inttab.empty) [] 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/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2344 +1,2341 @@ (* 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 extra_shyps: thm -> sort list val proof_bodies_of: thm list -> proof_body list val proof_body_of: thm -> proof_body val proof_of: thm -> 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 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 expose_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: string list * string list -> int -> thm -> thm val generalize_cterm: string list * string list -> int -> cterm -> cterm val generalize_ctyp: string list -> 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 strip_shyps: thm -> thm val unconstrainT: thm -> thm val varifyT_global': (string * sort) list -> 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 = if Context.subthy_id (Context.theory_id thy, theory_id th) then th else transfer thy th; fun join_transfer_context (ctxt, th) = if Context.subthy_id (Context.theory_id (Proof_Context.theory_of ctxt), theory_id th) then (Context.raw_transfer (theory_of_thm th) ctxt, th) else (ctxt, transfer' 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; (*dangling sort constraints of a thm*) fun extra_shyps (th as Thm (_, {shyps, ...})) = Sorts.subtract (fold_terms Sorts.insert_term th []) shyps; (** 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; val consolidate = ignore o proof_bodies_of; fun expose_proofs thy thms = if Proofterm.export_proof_boxes_required thy then - Proofterm.export_proof_boxes (map (proof_of o transfer thy) thms) + 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 "oracle", 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; (*** 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 = solve_constraints #> (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); -fun expose_derivation pos thm = - close_derivation pos thm - |> tap (expose_proof (theory_of_thm 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, SOME prop), Proofterm.oracle_proof name prop) | 1 => ((name, SOME prop), MinProof) | 0 => ((name, 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 ([], []) _ th = th | generalize (tfrees, frees) idx th = 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 null tfrees then K false else Term.exists_subtype (fn TFree (a, _) => member (op =) tfrees a | _ => false); fun bad_term (Free (x, T)) = bad_type T orelse member (op =) 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 ([], []) _ ct = ct | generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) = 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 [] _ cT = cT | generalize_ctyp tfrees idx (Ctyp {cert, T, maxidx, sorts}) = 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 pretty_typing thy t T = Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy T]; fun add_inst (v as (_, T), cu) (cert, sorts) = let val Cterm {t = u, T = U, cert = cert2, sorts = sorts_u, maxidx = maxidx_u, ...} = cu; val cert' = Context.join_certificate (cert, cert2); val sorts' = Sorts.union sorts_u sorts; in if T = U then ((v, (u, maxidx_u)), (cert', sorts')) else let val msg = (case cert' of Context.Certificate thy' => Pretty.string_of (Pretty.block [Pretty.str "instantiate: type conflict", Pretty.fbrk, pretty_typing thy' (Var v) T, Pretty.fbrk, pretty_typing thy' u U]) | Context.Certificate_Id _ => "instantiate: type conflict"); in raise TYPE (msg, [T, U], [Var v, u]) end end; fun add_instT (v as (_, S), cU) (cert, sorts) = let val Ctyp {T = U, cert = cert2, sorts = sorts_U, maxidx = maxidx_U, ...} = cU; val cert' = Context.join_certificate (cert, cert2); val thy' = Context.certificate_theory cert' handle ERROR msg => raise CONTEXT (msg, [cU], [], [], NONE); val sorts' = Sorts.union sorts_U sorts; in if Sign.of_sort thy' (U, S) then ((v, (U, maxidx_U)), (cert', sorts')) else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy' S, [U], []) 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 (inst', (instT', (cert', shyps'))) = (cert, shyps) |> fold_map add_inst inst ||> fold_map add_instT instT 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' = fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) instT' constraints; in Thm (deriv_rule1 (fn d => Proofterm.instantiate (map (apsnd #1) instT', map (apsnd #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 (inst', (instT', (cert', sorts'))) = (cert, sorts) |> fold_map add_inst inst ||> fold_map add_instT instT; 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.OfClass (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; (*Remove extra sorts that are witnessed by type signature information*) fun strip_shyps (thm as Thm (_, {shyps = [], ...})) = thm | strip_shyps (thm as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = let val thy = theory_of_thm thm; val algebra = Sign.classes_of thy; 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 extra' = fold (Sorts.remove_sort o #2) witnessed extra |> Sorts.minimal_sorts algebra; 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; (*Internalize sort constraints of type variables*) val unconstrainT = solve_constraints #> (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.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' []; (*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.empty |> 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 - |> expose_derivation \<^here> + |> 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)) |> 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 - |> expose_derivation \<^here> + |> 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;