diff --git a/src/Doc/Implementation/Proof.thy b/src/Doc/Implementation/Proof.thy --- a/src/Doc/Implementation/Proof.thy +++ b/src/Doc/Implementation/Proof.thy @@ -1,473 +1,473 @@ (*:maxLineLen=78:*) theory Proof imports Base begin chapter \Structured proofs\ section \Variables \label{sec:variables}\ text \ Any variable that is not explicitly bound by \\\-abstraction is considered as ``free''. Logically, free variables act like outermost universal quantification at the sequent level: \A\<^sub>1(x), \, A\<^sub>n(x) \ B(x)\ means that the result holds \<^emph>\for all\ values of \x\. Free variables for terms (not types) can be fully internalized into the logic: \\ B(x)\ and \\ \x. B(x)\ are interchangeable, provided that \x\ does not occur elsewhere in the context. Inspecting \\ \x. B(x)\ more closely, we see that inside the quantifier, \x\ is essentially ``arbitrary, but fixed'', while from outside it appears as a place-holder for instantiation (thanks to \\\ elimination). The Pure logic represents the idea of variables being either inside or outside the current scope by providing separate syntactic categories for \<^emph>\fixed variables\ (e.g.\ \x\) vs.\ \<^emph>\schematic variables\ (e.g.\ \?x\). Incidently, a universal result \\ \x. B(x)\ has the HHF normal form \\ B(?x)\, which represents its generality without requiring an explicit quantifier. The same principle works for type variables: \\ B(?\)\ represents the idea of ``\\ \\. B(\)\'' without demanding a truly polymorphic framework. \<^medskip> Additional care is required to treat type variables in a way that facilitates type-inference. In principle, term variables depend on type variables, which means that type variables would have to be declared first. For example, a raw type-theoretic framework would demand the context to be constructed in stages as follows: \\ = \: type, x: \, a: A(x\<^sub>\)\. We allow a slightly less formalistic mode of operation: term variables \x\ are fixed without specifying a type yet (essentially \<^emph>\all\ potential occurrences of some instance \x\<^sub>\\ are fixed); the first occurrence of \x\ within a specific term assigns its most general type, which is then maintained consistently in the context. The above example becomes \\ = x: term, \: type, A(x\<^sub>\)\, where type \\\ is fixed \<^emph>\after\ term \x\, and the constraint \x :: \\ is an implicit consequence of the occurrence of \x\<^sub>\\ in the subsequent proposition. This twist of dependencies is also accommodated by the reverse operation of exporting results from a context: a type variable \\\ is considered fixed as long as it occurs in some fixed term variable of the context. For example, exporting \x: term, \: type \ x\<^sub>\ \ x\<^sub>\\ produces in the first step \x: term \ x\<^sub>\ \ x\<^sub>\\ for fixed \\\, and only in the second step \\ ?x\<^sub>?\<^sub>\ \ ?x\<^sub>?\<^sub>\\ for schematic \?x\ and \?\\. The following Isar source text illustrates this scenario. \ notepad begin { fix x \ \all potential occurrences of some \x::\\ are fixed\ { have "x::'a \ x" \ \implicit type assignment by concrete occurrence\ by (rule reflexive) } thm this \ \result still with fixed type \'a\\ } thm this \ \fully general result for arbitrary \?x::?'a\\ end text \ The Isabelle/Isar proof context manages the details of term vs.\ type variables, with high-level principles for moving the frontier between fixed and schematic variables. The \add_fixes\ operation explicitly declares fixed variables; the \declare_term\ operation absorbs a term into a context by fixing new type variables and adding syntactic constraints. The \export\ operation is able to perform the main work of generalizing term and type variables as sketched above, assuming that fixing variables and terms have been declared properly. There \import\ operation makes a generalized fact a genuine part of the context, by inventing fixed variables for the schematic ones. The effect can be reversed by using \export\ later, potentially with an extended context; the result is equivalent to the original modulo renaming of schematic variables. The \focus\ operation provides a variant of \import\ for nested propositions (with explicit quantification): \\x\<^sub>1 \ x\<^sub>n. B(x\<^sub>1, \, x\<^sub>n)\ is decomposed by inventing fixed variables \x\<^sub>1, \, x\<^sub>n\ for the body. \ text %mlref \ \begin{mldecls} @{define_ML Variable.add_fixes: " string list -> Proof.context -> string list * Proof.context"} \\ @{define_ML Variable.variant_fixes: " string list -> Proof.context -> string list * Proof.context"} \\ @{define_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\ @{define_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\ @{define_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\ @{define_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\ @{define_ML Variable.import: "bool -> thm list -> Proof.context -> - ((((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list) * thm list) + ((ctyp Term_Subst.TVars.table * cterm Term_Subst.Vars.table) * thm list) * Proof.context"} \\ @{define_ML Variable.focus: "binding list option -> term -> Proof.context -> ((string * (string * typ)) list * term) * Proof.context"} \\ \end{mldecls} \<^descr> \<^ML>\Variable.add_fixes\~\xs ctxt\ fixes term variables \xs\, returning the resulting internal names. By default, the internal representation coincides with the external one, which also means that the given variables must not be fixed already. There is a different policy within a local proof body: the given names are just hints for newly invented Skolem variables. \<^descr> \<^ML>\Variable.variant_fixes\ is similar to \<^ML>\Variable.add_fixes\, but always produces fresh variants of the given names. \<^descr> \<^ML>\Variable.declare_term\~\t ctxt\ declares term \t\ to belong to the context. This automatically fixes new type variables, but not term variables. Syntactic constraints for type and term variables are declared uniformly, though. \<^descr> \<^ML>\Variable.declare_constraints\~\t ctxt\ declares syntactic constraints from term \t\, without making it part of the context yet. \<^descr> \<^ML>\Variable.export\~\inner outer thms\ generalizes fixed type and term variables in \thms\ according to the difference of the \inner\ and \outer\ context, following the principles sketched above. \<^descr> \<^ML>\Variable.polymorphic\~\ctxt ts\ generalizes type variables in \ts\ as far as possible, even those occurring in fixed term variables. The default policy of type-inference is to fix newly introduced type variables, which is essentially reversed with \<^ML>\Variable.polymorphic\: here the given terms are detached from the context as far as possible. \<^descr> \<^ML>\Variable.import\~\open thms ctxt\ invents fixed type and term variables for the schematic ones occurring in \thms\. The \open\ flag indicates whether the fixed names should be accessible to the user, otherwise newly introduced names are marked as ``internal'' (\secref{sec:names}). \<^descr> \<^ML>\Variable.focus\~\bindings B\ decomposes the outermost \\\ prefix of proposition \B\, using the given name bindings. \ text %mlex \ The following example shows how to work with fixed term and type parameters and with type-inference. \ ML_val \ (*static compile-time context -- for testing only*) val ctxt0 = \<^context>; (*locally fixed parameters -- no type assignment yet*) val ([x, y], ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y"]; (*t1: most general fixed type; t1': most general arbitrary type*) val t1 = Syntax.read_term ctxt1 "x"; val t1' = singleton (Variable.polymorphic ctxt1) t1; (*term u enforces specific type assignment*) val u = Syntax.read_term ctxt1 "(x::nat) \ y"; (*official declaration of u -- propagates constraints etc.*) val ctxt2 = ctxt1 |> Variable.declare_term u; val t2 = Syntax.read_term ctxt2 "x"; (*x::nat is enforced*) \ text \ In the above example, the starting context is derived from the toplevel theory, which means that fixed variables are internalized literally: \x\ is mapped again to \x\, and attempting to fix it again in the subsequent context is an error. Alternatively, fixed parameters can be renamed explicitly as follows: \ ML_val \ val ctxt0 = \<^context>; val ([x1, x2, x3], ctxt1) = ctxt0 |> Variable.variant_fixes ["x", "x", "x"]; \ text \ The following ML code can now work with the invented names of \x1\, \x2\, \x3\, without depending on the details on the system policy for introducing these variants. Recall that within a proof body the system always invents fresh ``Skolem constants'', e.g.\ as follows: \ notepad begin ML_prf %"ML" \val ctxt0 = \<^context>; val ([x1], ctxt1) = ctxt0 |> Variable.add_fixes ["x"]; val ([x2], ctxt2) = ctxt1 |> Variable.add_fixes ["x"]; val ([x3], ctxt3) = ctxt2 |> Variable.add_fixes ["x"]; val ([y1, y2], ctxt4) = ctxt3 |> Variable.variant_fixes ["y", "y"];\ end text \ In this situation \<^ML>\Variable.add_fixes\ and \<^ML>\Variable.variant_fixes\ are very similar, but identical name proposals given in a row are only accepted by the second version. \ section \Assumptions \label{sec:assumptions}\ text \ An \<^emph>\assumption\ is a proposition that it is postulated in the current context. Local conclusions may use assumptions as additional facts, but this imposes implicit hypotheses that weaken the overall statement. Assumptions are restricted to fixed non-schematic statements, i.e.\ all generality needs to be expressed by explicit quantifiers. Nevertheless, the result will be in HHF normal form with outermost quantifiers stripped. For example, by assuming \\x :: \. P x\ we get \\x :: \. P x \ P ?x\ for schematic \?x\ of fixed type \\\. Local derivations accumulate more and more explicit references to hypotheses: \A\<^sub>1, \, A\<^sub>n \ B\ where \A\<^sub>1, \, A\<^sub>n\ needs to be covered by the assumptions of the current context. \<^medskip> The \add_assms\ operation augments the context by local assumptions, which are parameterized by an arbitrary \export\ rule (see below). The \export\ operation moves facts from a (larger) inner context into a (smaller) outer context, by discharging the difference of the assumptions as specified by the associated export rules. Note that the discharged portion is determined by the difference of contexts, not the facts being exported! There is a separate flag to indicate a goal context, where the result is meant to refine an enclosing sub-goal of a structured proof state. \<^medskip> The most basic export rule discharges assumptions directly by means of the \\\ introduction rule: \[ \infer[(\\\intro\)]{\\ - A \ A \ B\}{\\ \ B\} \] The variant for goal refinements marks the newly introduced premises, which causes the canonical Isar goal refinement scheme to enforce unification with local premises within the goal: \[ \infer[(\#\\intro\)]{\\ - A \ #A \ B\}{\\ \ B\} \] \<^medskip> Alternative versions of assumptions may perform arbitrary transformations on export, as long as the corresponding portion of hypotheses is removed from the given facts. For example, a local definition works by fixing \x\ and assuming \x \ t\, with the following export rule to reverse the effect: \[ \infer[(\\\expand\)]{\\ - (x \ t) \ B t\}{\\ \ B x\} \] This works, because the assumption \x \ t\ was introduced in a context with \x\ being fresh, so \x\ does not occur in \\\ here. \ text %mlref \ \begin{mldecls} @{define_ML_type Assumption.export} \\ @{define_ML Assumption.assume: "Proof.context -> cterm -> thm"} \\ @{define_ML Assumption.add_assms: "Assumption.export -> cterm list -> Proof.context -> thm list * Proof.context"} \\ @{define_ML Assumption.add_assumes: " cterm list -> Proof.context -> thm list * Proof.context"} \\ @{define_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Assumption.export\ represents arbitrary export rules, which is any function of type \<^ML_type>\bool -> cterm list -> thm -> thm\, where the \<^ML_type>\bool\ indicates goal mode, and the \<^ML_type>\cterm list\ the collection of assumptions to be discharged simultaneously. \<^descr> \<^ML>\Assumption.assume\~\ctxt A\ turns proposition \A\ into a primitive assumption \A \ A'\, where the conclusion \A'\ is in HHF normal form. \<^descr> \<^ML>\Assumption.add_assms\~\r As\ augments the context by assumptions \As\ with export rule \r\. The resulting facts are hypothetical theorems as produced by the raw \<^ML>\Assumption.assume\. \<^descr> \<^ML>\Assumption.add_assumes\~\As\ is a special case of \<^ML>\Assumption.add_assms\ where the export rule performs \\\intro\ or \#\\intro\, depending on goal mode. \<^descr> \<^ML>\Assumption.export\~\is_goal inner outer thm\ exports result \thm\ from the the \inner\ context back into the \outer\ one; \is_goal = true\ means this is a goal context. The result is in HHF normal form. Note that \<^ML>\Proof_Context.export\ combines \<^ML>\Variable.export\ and \<^ML>\Assumption.export\ in the canonical way. \ text %mlex \ The following example demonstrates how rules can be derived by building up a context of assumptions first, and exporting some local fact afterwards. We refer to \<^theory>\Pure\ equality here for testing purposes. \ ML_val \ (*static compile-time context -- for testing only*) val ctxt0 = \<^context>; val ([eq], ctxt1) = ctxt0 |> Assumption.add_assumes [\<^cprop>\x \ y\]; val eq' = Thm.symmetric eq; (*back to original context -- discharges assumption*) val r = Assumption.export false ctxt1 ctxt0 eq'; \ text \ Note that the variables of the resulting rule are not generalized. This would have required to fix them properly in the context beforehand, and export wrt.\ variables afterwards (cf.\ \<^ML>\Variable.export\ or the combined \<^ML>\Proof_Context.export\). \ section \Structured goals and results \label{sec:struct-goals}\ text \ Local results are established by monotonic reasoning from facts within a context. This allows common combinations of theorems, e.g.\ via \\/\\ elimination, resolution rules, or equational reasoning, see \secref{sec:thms}. Unaccounted context manipulations should be avoided, notably raw \\/\\ introduction or ad-hoc references to free variables or assumptions not present in the proof context. \<^medskip> The \SUBPROOF\ combinator allows to structure a tactical proof recursively by decomposing a selected sub-goal: \(\x. A(x) \ B(x)) \ \\ is turned into \B(x) \ \\ after fixing \x\ and assuming \A(x)\. This means the tactic needs to solve the conclusion, but may use the premise as a local fact, for locally fixed variables. The family of \FOCUS\ combinators is similar to \SUBPROOF\, but allows to retain schematic variables and pending subgoals in the resulting goal state. The \prove\ operation provides an interface for structured backwards reasoning under program control, with some explicit sanity checks of the result. The goal context can be augmented by additional fixed variables (cf.\ \secref{sec:variables}) and assumptions (cf.\ \secref{sec:assumptions}), which will be available as local facts during the proof and discharged into implications in the result. Type and term variables are generalized as usual, according to the context. The \obtain\ operation produces results by eliminating existing facts by means of a given tactic. This acts like a dual conclusion: the proof demonstrates that the context may be augmented by parameters and assumptions, without affecting any conclusions that do not mention these parameters. See also @{cite "isabelle-isar-ref"} for the user-level @{command obtain} and @{command guess} elements. Final results, which may not refer to the parameters in the conclusion, need to exported explicitly into the original context.\ text %mlref \ \begin{mldecls} @{define_ML SUBPROOF: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{define_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{define_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{define_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{define_ML Subgoal.focus: "Proof.context -> int -> binding list option -> thm -> Subgoal.focus * thm"} \\ @{define_ML Subgoal.focus_prems: "Proof.context -> int -> binding list option -> thm -> Subgoal.focus * thm"} \\ @{define_ML Subgoal.focus_params: "Proof.context -> int -> binding list option -> thm -> Subgoal.focus * thm"} \\ \end{mldecls} \begin{mldecls} @{define_ML Goal.prove: "Proof.context -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm"} \\ @{define_ML Goal.prove_common: "Proof.context -> int option -> string list -> term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) -> thm list"} \\ \end{mldecls} \begin{mldecls} @{define_ML Obtain.result: "(Proof.context -> tactic) -> thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\ \end{mldecls} \<^descr> \<^ML>\SUBPROOF\~\tac ctxt i\ decomposes the structure of the specified sub-goal, producing an extended context and a reduced goal, which needs to be solved by the given tactic. All schematic parameters of the goal are imported into the context as fixed ones, which may not be instantiated in the sub-proof. \<^descr> \<^ML>\Subgoal.FOCUS\, \<^ML>\Subgoal.FOCUS_PREMS\, and \<^ML>\Subgoal.FOCUS_PARAMS\ are similar to \<^ML>\SUBPROOF\, but are slightly more flexible: only the specified parts of the subgoal are imported into the context, and the body tactic may introduce new subgoals and schematic variables. \<^descr> \<^ML>\Subgoal.focus\, \<^ML>\Subgoal.focus_prems\, \<^ML>\Subgoal.focus_params\ extract the focus information from a goal state in the same way as the corresponding tacticals above. This is occasionally useful to experiment without writing actual tactics yet. \<^descr> \<^ML>\Goal.prove\~\ctxt xs As C tac\ states goal \C\ in the context augmented by fixed variables \xs\ and assumptions \As\, and applies tactic \tac\ to solve it. The latter may depend on the local assumptions being presented as facts. The result is in HHF normal form. \<^descr> \<^ML>\Goal.prove_common\~\ctxt fork_pri\ is the common form to state and prove a simultaneous goal statement, where \<^ML>\Goal.prove\ is a convenient shorthand that is most frequently used in applications. The given list of simultaneous conclusions is encoded in the goal state by means of Pure conjunction: \<^ML>\Goal.conjunction_tac\ will turn this into a collection of individual subgoals, but note that the original multi-goal state is usually required for advanced induction. It is possible to provide an optional priority for a forked proof, typically \<^ML>\SOME ~1\, while \<^ML>\NONE\ means the proof is immediate (sequential) as for \<^ML>\Goal.prove\. Note that a forked proof does not exhibit any failures in the usual way via exceptions in ML, but accumulates error situations under the execution id of the running transaction. Thus the system is able to expose error messages ultimately to the end-user, even though the subsequent ML code misses them. \<^descr> \<^ML>\Obtain.result\~\tac thms ctxt\ eliminates the given facts using a tactic, which results in additional fixed variables and assumptions in the context. Final results need to be exported explicitly. \ text %mlex \ The following minimal example illustrates how to access the focus information of a structured goal state. \ notepad begin fix A B C :: "'a \ bool" have "\x. A x \ B x \ C x" ML_val \val {goal, context = goal_ctxt, ...} = @{Isar.goal}; val (focus as {params, asms, concl, ...}, goal') = Subgoal.focus goal_ctxt 1 (SOME [\<^binding>\x\]) goal; val [A, B] = #prems focus; val [(_, x)] = #params focus;\ sorry end text \ \<^medskip> The next example demonstrates forward-elimination in a local context, using \<^ML>\Obtain.result\. \ notepad begin assume ex: "\x. B x" ML_prf %"ML" \val ctxt0 = \<^context>; val (([(_, x)], [B]), ctxt1) = ctxt0 |> Obtain.result (fn _ => eresolve_tac ctxt0 @{thms exE} 1) [@{thm ex}];\ ML_prf %"ML" \singleton (Proof_Context.export ctxt1 ctxt0) @{thm refl};\ ML_prf %"ML" \Proof_Context.export ctxt1 ctxt0 [Thm.reflexive x] handle ERROR msg => (warning msg; []);\ end end diff --git a/src/Doc/Typeclass_Hierarchy/Setup.thy b/src/Doc/Typeclass_Hierarchy/Setup.thy --- a/src/Doc/Typeclass_Hierarchy/Setup.thy +++ b/src/Doc/Typeclass_Hierarchy/Setup.thy @@ -1,39 +1,39 @@ theory Setup imports Complex_Main "HOL-Library.Multiset" "HOL-Library.Lattice_Syntax" begin ML_file \../antiquote_setup.ML\ ML_file \../more_antiquote.ML\ attribute_setup all = - \Scan.succeed (Thm.rule_attribute [] (K Drule.forall_intr_vars))\ + \Scan.succeed (Thm.rule_attribute [] (K Thm.forall_intr_vars))\ "quantified schematic vars" setup \Config.put_global Printer.show_type_emphasis false\ setup \ let fun strip_all t = if Logic.is_all t then case snd (dest_comb t) of Abs (w, T, t') => strip_all t' |>> cons (w, T) else ([], t); fun frugal (w, T as TFree (v, sort)) visited = if member (op =) visited v then ((w, dummyT), visited) else ((w, T), v :: visited) | frugal (w, T) visited = ((w, dummyT), visited); fun frugal_sorts t = let val (vTs, body) = strip_all t val (vTs', _) = fold_map frugal vTs []; in Logic.list_all (vTs', map_types (K dummyT) body) end; in Term_Style.setup \<^binding>\frugal_sorts\ (Scan.succeed (K frugal_sorts)) end \ declare [[show_sorts]] end diff --git a/src/HOL/Analysis/metric_arith.ML b/src/HOL/Analysis/metric_arith.ML --- a/src/HOL/Analysis/metric_arith.ML +++ b/src/HOL/Analysis/metric_arith.ML @@ -1,324 +1,323 @@ signature METRIC_ARITH = sig val metric_arith_tac : Proof.context -> int -> tactic val trace: bool Config.T end structure MetricArith : METRIC_ARITH = struct fun default d x = case x of SOME y => SOME y | NONE => d (* apply f to both cterms in ct_pair, merge results *) fun app_union_ct_pair f ct_pair = uncurry (union (op aconvc)) (apply2 f ct_pair) val trace = Attrib.setup_config_bool \<^binding>\metric_trace\ (K false) fun trace_tac ctxt msg = if Config.get ctxt trace then print_tac ctxt msg else all_tac fun argo_trace_ctxt ctxt = if Config.get ctxt trace then Config.map (Argo_Tactic.trace) (K "basic") ctxt else ctxt fun IF_UNSOLVED' tac i = IF_UNSOLVED (tac i) fun REPEAT' tac i = REPEAT (tac i) -fun frees ct = Drule.cterm_add_frees ct [] -fun free_in v ct = member (op aconvc) (frees ct) v +fun free_in v ct = member (op aconvc) (Misc_Legacy.cterm_frees ct) v (* build a cterm set with elements cts of type ty *) fun mk_ct_set ctxt ty = map Thm.term_of #> HOLogic.mk_set ty #> Thm.cterm_of ctxt fun prenex_tac ctxt = let val prenex_simps = Proof_Context.get_thms ctxt @{named_theorems metric_prenex} val prenex_ctxt = put_simpset HOL_basic_ss ctxt addsimps prenex_simps in simp_tac prenex_ctxt THEN' K (trace_tac ctxt "Prenex form") end fun nnf_tac ctxt = let val nnf_simps = Proof_Context.get_thms ctxt @{named_theorems metric_nnf} val nnf_ctxt = put_simpset HOL_basic_ss ctxt addsimps nnf_simps in simp_tac nnf_ctxt THEN' K (trace_tac ctxt "NNF form") end fun unfold_tac ctxt = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps ( Proof_Context.get_thms ctxt @{named_theorems metric_unfold})) fun pre_arith_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps ( Proof_Context.get_thms ctxt @{named_theorems metric_pre_arith})) THEN' K (trace_tac ctxt "Prepared for decision procedure") fun dist_refl_sym_tac ctxt = let val refl_sym_simps = @{thms dist_self dist_commute add_0_right add_0_left simp_thms} val refl_sym_ctxt = put_simpset HOL_basic_ss ctxt addsimps refl_sym_simps in simp_tac refl_sym_ctxt THEN' K (trace_tac ctxt ("Simplified using " ^ @{make_string} refl_sym_simps)) end fun is_exists ct = case Thm.term_of ct of Const (\<^const_name>\HOL.Ex\,_)$_ => true | Const (\<^const_name>\Trueprop\,_)$_ => is_exists (Thm.dest_arg ct) | _ => false fun is_forall ct = case Thm.term_of ct of Const (\<^const_name>\HOL.All\,_)$_ => true | Const (\<^const_name>\Trueprop\,_)$_ => is_forall (Thm.dest_arg ct) | _ => false fun dist_ty mty = mty --> mty --> \<^typ>\real\ (* find all free points in ct of type metric_ty *) fun find_points ctxt metric_ty ct = let fun find ct = (if Thm.typ_of_cterm ct = metric_ty then [ct] else []) @ ( case Thm.term_of ct of _ $ _ => app_union_ct_pair find (Thm.dest_comb ct) | Abs (_, _, _) => (* ensure the point doesn't contain the bound variable *) let val (var, bod) = Thm.dest_abs NONE ct in filter (free_in var #> not) (find bod) end | _ => []) val points = find ct in case points of [] => (* if no point can be found, invent one *) let val free_name = Term.variant_frees (Thm.term_of ct) [("x", metric_ty)] in map (Free #> Thm.cterm_of ctxt) free_name end | _ => points end (* find all cterms "dist x y" in ct, where x and y have type metric_ty *) fun find_dist metric_ty ct = let val dty = dist_ty metric_ty fun find ct = case Thm.term_of ct of Const (\<^const_name>\dist\, ty) $ _ $ _ => if ty = dty then [ct] else [] | _ $ _ => app_union_ct_pair find (Thm.dest_comb ct) | Abs (_, _, _) => let val (var, bod) = Thm.dest_abs NONE ct in filter (free_in var #> not) (find bod) end | _ => [] in find ct end (* find all "x=y", where x has type metric_ty *) fun find_eq metric_ty ct = let fun find ct = case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, ty) $ _ $ _ => if fst (dest_funT ty) = metric_ty then [ct] else app_union_ct_pair find (Thm.dest_binop ct) | _ $ _ => app_union_ct_pair find (Thm.dest_comb ct) | Abs (_, _, _) => let val (var, bod) = Thm.dest_abs NONE ct in filter (free_in var #> not) (find bod) end | _ => [] in find ct end (* rewrite ct of the form "dist x y" using maxdist_thm *) fun maxdist_conv ctxt fset_ct ct = let val (xct, yct) = Thm.dest_binop ct val solve_prems = rule_by_tactic ctxt (ALLGOALS (simp_tac (put_simpset HOL_ss ctxt addsimps @{thms finite.emptyI finite_insert empty_iff insert_iff}))) val image_simp = Simplifier.rewrite (put_simpset HOL_ss ctxt addsimps @{thms image_empty image_insert}) val dist_refl_sym_simp = Simplifier.rewrite (put_simpset HOL_ss ctxt addsimps @{thms dist_commute dist_self}) val algebra_simp = Simplifier.rewrite (put_simpset HOL_ss ctxt addsimps @{thms diff_self diff_0_right diff_0 abs_zero abs_minus_cancel abs_minus_commute}) val insert_simp = Simplifier.rewrite (put_simpset HOL_ss ctxt addsimps @{thms insert_absorb2 insert_commute}) val sup_simp = Simplifier.rewrite (put_simpset HOL_ss ctxt addsimps @{thms cSup_singleton Sup_insert_insert}) val real_abs_dist_simp = Simplifier.rewrite (put_simpset HOL_ss ctxt addsimps @{thms real_abs_dist}) val maxdist_thm = @{thm maxdist_thm} |> infer_instantiate' ctxt [SOME fset_ct, SOME xct, SOME yct] |> solve_prems in ((Conv.rewr_conv maxdist_thm) then_conv (* SUP to Sup *) image_simp then_conv dist_refl_sym_simp then_conv algebra_simp then_conv (* eliminate duplicate terms in set *) insert_simp then_conv (* Sup to max *) sup_simp then_conv real_abs_dist_simp) ct end (* rewrite ct of the form "x=y" using metric_eq_thm *) fun metric_eq_conv ctxt fset_ct ct = let val (xct, yct) = Thm.dest_binop ct val solve_prems = rule_by_tactic ctxt (ALLGOALS (simp_tac (put_simpset HOL_ss ctxt addsimps @{thms empty_iff insert_iff}))) val ball_simp = Simplifier.rewrite (put_simpset HOL_ss ctxt addsimps @{thms Set.ball_empty ball_insert}) val dist_refl_sym_simp = Simplifier.rewrite (put_simpset HOL_ss ctxt addsimps @{thms dist_commute dist_self}) val metric_eq_thm = @{thm metric_eq_thm} |> infer_instantiate' ctxt [SOME xct, SOME fset_ct, SOME yct] |> solve_prems in ((Conv.rewr_conv metric_eq_thm) then_conv (* convert \x\{x\<^sub>1,...,x\<^sub>n}. P x to P x\<^sub>1 \ ... \ P x\<^sub>n *) ball_simp then_conv dist_refl_sym_simp) ct end (* build list of theorems "0 \ dist x y" for all dist terms in ct *) fun augment_dist_pos ctxt metric_ty ct = let fun inst dist_ct = let val (xct, yct) = Thm.dest_binop dist_ct in infer_instantiate' ctxt [SOME xct, SOME yct] @{thm zero_le_dist} end in map inst (find_dist metric_ty ct) end fun top_sweep_rewrs_tac ctxt thms = CONVERSION (Conv.top_sweep_conv (K (Conv.rewrs_conv thms)) ctxt) (* apply maxdist_conv and metric_eq_conv to the goal, thereby embedding the goal in (\\<^sup>n,dist\<^sub>\) *) fun embedding_tac ctxt metric_ty i goal = let val ct = (Thm.cprem_of goal 1) val points = find_points ctxt metric_ty ct val fset_ct = mk_ct_set ctxt metric_ty points (* embed all subterms of the form "dist x y" in (\\<^sup>n,dist\<^sub>\) *) val eq1 = map (maxdist_conv ctxt fset_ct) (find_dist metric_ty ct) (* replace point equality by equality of components in \\<^sup>n *) val eq2 = map (metric_eq_conv ctxt fset_ct) (find_eq metric_ty ct) in ( K (trace_tac ctxt "Embedding into \\<^sup>n") THEN' top_sweep_rewrs_tac ctxt (eq1 @ eq2))i goal end (* decision procedure for linear real arithmetic *) fun lin_real_arith_tac ctxt metric_ty i goal = let val dist_thms = augment_dist_pos ctxt metric_ty (Thm.cprem_of goal 1) val ctxt' = argo_trace_ctxt ctxt in (Argo_Tactic.argo_tac ctxt' dist_thms) i goal end fun basic_metric_arith_tac ctxt metric_ty = HEADGOAL (dist_refl_sym_tac ctxt THEN' IF_UNSOLVED' (embedding_tac ctxt metric_ty) THEN' IF_UNSOLVED' (pre_arith_tac ctxt) THEN' IF_UNSOLVED' (lin_real_arith_tac ctxt metric_ty)) (* tries to infer the metric space from ct from dist terms, if no dist terms are present, equality terms will be used *) fun guess_metric ctxt ct = let fun find_dist ct = case Thm.term_of ct of Const (\<^const_name>\dist\, ty) $ _ $ _ => SOME (fst (dest_funT ty)) | _ $ _ => let val (s, t) = Thm.dest_comb ct in default (find_dist t) (find_dist s) end | Abs (_, _, _) => find_dist (snd (Thm.dest_abs NONE ct)) | _ => NONE fun find_eq ct = case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, ty) $ x $ _ => let val (l, r) = Thm.dest_binop ct in if Sign.of_sort (Proof_Context.theory_of ctxt) (type_of x, \<^sort>\metric_space\) then SOME (fst (dest_funT ty)) else default (find_dist r) (find_eq l) end | _ $ _ => let val (s, t) = Thm.dest_comb ct in default (find_eq t) (find_eq s) end | Abs (_, _, _) => find_eq (snd (Thm.dest_abs NONE ct)) | _ => NONE in case default (find_eq ct) (find_dist ct) of SOME ty => ty | NONE => error "No Metric Space was found" end (* eliminate \ by proving the goal for a single witness from the metric space *) fun elim_exists ctxt goal = let val ct = Thm.cprem_of goal 1 val metric_ty = guess_metric ctxt ct val points = find_points ctxt metric_ty ct fun try_point ctxt pt = let val ex_rule = infer_instantiate' ctxt [NONE, SOME pt] @{thm exI} in HEADGOAL (resolve_tac ctxt [ex_rule] ORELSE' (* variable doesn't occur in body *) resolve_tac ctxt @{thms exI}) THEN trace_tac ctxt ("Removed existential quantifier, try " ^ @{make_string} pt) THEN try_points ctxt end and try_points ctxt goal = ( if is_exists (Thm.cprem_of goal 1) then FIRST (map (try_point ctxt) points) else if is_forall (Thm.cprem_of goal 1) then HEADGOAL (resolve_tac ctxt @{thms HOL.allI} THEN' Subgoal.FOCUS (fn {context = ctxt', ...} => trace_tac ctxt "Removed universal quantifier" THEN try_points ctxt') ctxt) else basic_metric_arith_tac ctxt metric_ty) goal in try_points ctxt goal end fun metric_arith_tac ctxt = (* unfold common definitions to get rid of sets *) unfold_tac ctxt THEN' (* remove all meta-level connectives *) IF_UNSOLVED' (Object_Logic.full_atomize_tac ctxt) THEN' (* convert goal to prenex form *) IF_UNSOLVED' (prenex_tac ctxt) THEN' (* and NNF to ? *) IF_UNSOLVED' (nnf_tac ctxt) THEN' (* turn all universally quantified variables into free variables, by focusing the subgoal *) REPEAT' (resolve_tac ctxt @{thms HOL.allI}) THEN' IF_UNSOLVED' (SUBPROOF (fn {context=ctxt', ...} => trace_tac ctxt' "Focused on subgoal" THEN elim_exists ctxt') ctxt) end diff --git a/src/HOL/Decision_Procs/ferrante_rackoff.ML b/src/HOL/Decision_Procs/ferrante_rackoff.ML --- a/src/HOL/Decision_Procs/ferrante_rackoff.ML +++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML @@ -1,233 +1,233 @@ (* Title: HOL/Decision_Procs/ferrante_rackoff.ML Author: Amine Chaieb, TU Muenchen Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders. Proof-synthesis and tactic. *) signature FERRANTE_RACKOFF = sig val dlo_conv: Proof.context -> conv val dlo_tac: Proof.context -> int -> tactic end; structure FerranteRackoff: FERRANTE_RACKOFF = struct open Ferrante_Rackoff_Data; open Conv; type entry = {minf: thm list, pinf: thm list, nmi: thm list, npi: thm list, ld: thm list, qe: thm, atoms : cterm list} * {isolate_conv: cterm list -> cterm -> thm, whatis : cterm -> cterm -> ord, simpset : simpset}; fun get_p1 th = funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE) (funpow 2 Thm.dest_arg (Thm.cprop_of th)) |> Thm.dest_arg fun ferrack_conv ctxt (entr as ({minf = minf, pinf = pinf, nmi = nmi, npi = npi, ld = ld, qe = qe, atoms = atoms}, {isolate_conv = icv, whatis = wi, simpset = simpset}):entry) = let fun uset (vars as (x::vs)) p = case Thm.term_of p of Const(\<^const_name>\HOL.conj\, _)$ _ $ _ => let val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb val (lS,lth) = uset vars l val (rS, rth) = uset vars r in (lS@rS, Drule.binop_cong_rule b lth rth) end | Const(\<^const_name>\HOL.disj\, _)$ _ $ _ => let val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb val (lS,lth) = uset vars l val (rS, rth) = uset vars r in (lS@rS, Drule.binop_cong_rule b lth rth) end | _ => let val th = icv vars p val p' = Thm.rhs_of th val c = wi x p' val S = (if member (op =) [Lt, Le, Eq] c then single o Thm.dest_arg else if member (op =) [Gt, Ge] c then single o Thm.dest_arg1 else if c = NEq then single o Thm.dest_arg o Thm.dest_arg else K []) p' in (S,th) end val ((p1_v,p2_v),(mp1_v,mp2_v)) = funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE) (funpow 4 Thm.dest_arg (Thm.cprop_of (hd minf))) |> Thm.dest_binop |> apply2 Thm.dest_binop |> apfst (apply2 Thm.dest_fun) |> apply2 (apply2 (dest_Var o Thm.term_of)) fun myfwd (th1, th2, th3, th4, th5) p1 p2 [(th_1,th_2,th_3,th_4,th_5), (th_1',th_2',th_3',th_4',th_5')] = let val (mp1, mp2) = (get_p1 th_1, get_p1 th_1') val (pp1, pp2) = (get_p1 th_2, get_p1 th_2') fun fw mi th th' th'' = let val th0 = if mi then Drule.instantiate_normalize ([],[(p1_v, p1),(p2_v, p2),(mp1_v, mp1), (mp2_v, mp2)]) th else Drule.instantiate_normalize ([],[(p1_v, p1),(p2_v, p2),(mp1_v, pp1), (mp2_v, pp2)]) th in Thm.implies_elim (Thm.implies_elim th0 th') th'' end in (fw true th1 th_1 th_1', fw false th2 th_2 th_2', fw true th3 th_3 th_3', fw false th4 th_4 th_4', fw true th5 th_5 th_5') end val U_v = dest_Var (Thm.term_of (Thm.dest_arg (Thm.dest_arg (Thm.dest_arg1 (Thm.cprop_of qe))))) fun main vs p = let val ((xn,ce),(x,fm)) = (case Thm.term_of p of Const(\<^const_name>\Ex\,_)$Abs(xn,xT,_) => Thm.dest_comb p ||> Thm.dest_abs (SOME xn) |>> pair xn | _ => raise CTERM ("main QE only treats existential quantifiers!", [p])) val cT = Thm.ctyp_of_cterm x val (u,nth) = uset (x::vs) fm |>> distinct (op aconvc) val nthx = Thm.abstract_rule xn x nth val q = Thm.rhs_of nth val qx = Thm.rhs_of nthx val enth = Drule.arg_cong_rule ce nthx val [th0,th1] = map (Thm.instantiate' [SOME cT] []) @{thms "finite.intros"} fun ins x th = Thm.implies_elim (Thm.instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg) (Thm.cprop_of th), SOME x] th1) th val fU = fold ins u th0 val cU = funpow 2 Thm.dest_arg (Thm.cprop_of fU) local val insI1 = Thm.instantiate' [SOME cT] [] @{thm "insertI1"} val insI2 = Thm.instantiate' [SOME cT] [] @{thm "insertI2"} in fun provein x S = case Thm.term_of S of Const(\<^const_name>\Orderings.bot\, _) => raise CTERM ("provein : not a member!", [S]) | Const(\<^const_name>\insert\, _) $ y $_ => let val (cy,S') = Thm.dest_binop S in if Thm.term_of x aconv y then Thm.instantiate' [] [SOME x, SOME S'] insI1 else Thm.implies_elim (Thm.instantiate' [] [SOME x, SOME S', SOME cy] insI2) (provein x S') end end val tabU = fold (fn t => fn tab => Termtab.update (Thm.term_of t, provein t cU) tab) u Termtab.empty val U = the o Termtab.lookup tabU o Thm.term_of val [minf_conj, minf_disj, minf_eq, minf_neq, minf_lt, minf_le, minf_gt, minf_ge, minf_P] = minf val [pinf_conj, pinf_disj, pinf_eq, pinf_neq, pinf_lt, pinf_le, pinf_gt, pinf_ge, pinf_P] = pinf val [nmi_conj, nmi_disj, nmi_eq, nmi_neq, nmi_lt, nmi_le, nmi_gt, nmi_ge, nmi_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) nmi val [npi_conj, npi_disj, npi_eq, npi_neq, npi_lt, npi_le, npi_gt, npi_ge, npi_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) npi val [ld_conj, ld_disj, ld_eq, ld_neq, ld_lt, ld_le, ld_gt, ld_ge, ld_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) ld fun decomp_mpinf fm = case Thm.term_of fm of Const(\<^const_name>\HOL.conj\,_)$_$_ => let val (p,q) = Thm.dest_binop fm in ([p,q], myfwd (minf_conj,pinf_conj, nmi_conj, npi_conj,ld_conj) (Thm.lambda x p) (Thm.lambda x q)) end | Const(\<^const_name>\HOL.disj\,_)$_$_ => let val (p,q) = Thm.dest_binop fm in ([p,q],myfwd (minf_disj, pinf_disj, nmi_disj, npi_disj,ld_disj) (Thm.lambda x p) (Thm.lambda x q)) end | _ => (let val c = wi x fm val t = (if c=Nox then I else if member (op =) [Lt, Le, Eq] c then Thm.dest_arg else if member (op =) [Gt, Ge] c then Thm.dest_arg1 else if c = NEq then (Thm.dest_arg o Thm.dest_arg) else raise Fail "decomp_mpinf: Impossible case!!") fm val [mi_th, pi_th, nmi_th, npi_th, ld_th] = if c = Nox then map (Thm.instantiate' [] [SOME fm]) [minf_P, pinf_P, nmi_P, npi_P, ld_P] else let val [mi_th,pi_th,nmi_th,npi_th,ld_th] = map (Thm.instantiate' [] [SOME t]) (case c of Lt => [minf_lt, pinf_lt, nmi_lt, npi_lt, ld_lt] | Le => [minf_le, pinf_le, nmi_le, npi_le, ld_le] | Gt => [minf_gt, pinf_gt, nmi_gt, npi_gt, ld_gt] | Ge => [minf_ge, pinf_ge, nmi_ge, npi_ge, ld_ge] | Eq => [minf_eq, pinf_eq, nmi_eq, npi_eq, ld_eq] | NEq => [minf_neq, pinf_neq, nmi_neq, npi_neq, ld_neq]) val tU = U t fun Ufw th = Thm.implies_elim th tU in [mi_th, pi_th, Ufw nmi_th, Ufw npi_th, Ufw ld_th] end in ([], K (mi_th, pi_th, nmi_th, npi_th, ld_th)) end) val (minf_th, pinf_th, nmi_th, npi_th, ld_th) = divide_and_conquer decomp_mpinf q val qe_th = Drule.implies_elim_list ((fconv_rule (Thm.beta_conversion true)) (Thm.instantiate' [] (map SOME [cU, qx, get_p1 minf_th, get_p1 pinf_th]) qe)) [fU, ld_th, nmi_th, npi_th, minf_th, pinf_th] val bex_conv = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms bex_simps(1-5)}) val result_th = fconv_rule (arg_conv bex_conv) (Thm.transitive enth qe_th) in result_th end in main end; val grab_atom_bop = let fun h bounds tm = (case Thm.term_of tm of Const (\<^const_name>\HOL.eq\, T) $ _ $ _ => if domain_type T = HOLogic.boolT then find_args bounds tm else Thm.dest_fun2 tm | Const (\<^const_name>\Not\, _) $ _ => h bounds (Thm.dest_arg tm) | Const (\<^const_name>\All\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\Ex\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\HOL.conj\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\HOL.disj\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\HOL.implies\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Pure.imp\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Pure.all\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\Trueprop\, _) $ _ => h bounds (Thm.dest_arg tm) | _ => Thm.dest_fun2 tm) and find_args bounds tm = (h bounds (Thm.dest_arg tm) handle CTERM _ => Thm.dest_arg1 tm) and find_body bounds b = let val (_, b') = Thm.dest_abs (SOME (Name.bound bounds)) b in h (bounds + 1) b' end; in h end; fun raw_ferrack_qe_conv ctxt (thy, {isolate_conv, whatis, simpset = ss}) tm = let val ss' = merge_ss (simpset_of (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps all_simps not_all all_not_ex ex_disj_distrib}), ss); val pcv = Simplifier.rewrite (put_simpset ss' ctxt); val postcv = Simplifier.rewrite (put_simpset ss ctxt); val nnf = K (nnf_conv ctxt then_conv postcv) - val qe_conv = Qelim.gen_qelim_conv ctxt pcv postcv pcv cons (Drule.cterm_add_frees tm []) + val qe_conv = Qelim.gen_qelim_conv ctxt pcv postcv pcv cons (Misc_Legacy.cterm_frees tm) (isolate_conv ctxt) nnf (fn vs => ferrack_conv ctxt (thy,{isolate_conv = isolate_conv ctxt, whatis = whatis, simpset = ss}) vs then_conv postcv) in (Simplifier.rewrite (put_simpset ss ctxt) then_conv qe_conv) tm end; fun dlo_instance ctxt tm = Ferrante_Rackoff_Data.match ctxt (grab_atom_bop 0 tm); fun dlo_conv ctxt tm = (case dlo_instance ctxt tm of NONE => raise CTERM ("ferrackqe_conv: no corresponding instance in context!", [tm]) | SOME instance => raw_ferrack_qe_conv ctxt instance tm); fun dlo_tac ctxt = CSUBGOAL (fn (p, i) => (case dlo_instance ctxt p of NONE => no_tac | SOME instance => Object_Logic.full_atomize_tac ctxt i THEN simp_tac (put_simpset (#simpset (snd instance)) ctxt) i THEN (* FIXME already part of raw_ferrack_qe_conv? *) CONVERSION (Object_Logic.judgment_conv ctxt (raw_ferrack_qe_conv ctxt instance)) i THEN simp_tac ctxt i)); (* FIXME really? *) end; diff --git a/src/HOL/Decision_Procs/langford.ML b/src/HOL/Decision_Procs/langford.ML --- a/src/HOL/Decision_Procs/langford.ML +++ b/src/HOL/Decision_Procs/langford.ML @@ -1,263 +1,263 @@ (* Title: HOL/Decision_Procs/langford.ML Author: Amine Chaieb, TU Muenchen *) signature LANGFORD = sig val dlo_tac : Proof.context -> int -> tactic val dlo_conv : Proof.context -> cterm -> thm end structure Langford: LANGFORD = struct val dest_set = let fun h acc ct = (case Thm.term_of ct of Const (\<^const_name>\Orderings.bot\, _) => acc | Const (\<^const_name>\insert\, _) $ _ $ t => h (Thm.dest_arg1 ct :: acc) (Thm.dest_arg ct)); in h [] end; fun prove_finite cT u = let val [th0, th1] = map (Thm.instantiate' [SOME cT] []) @{thms finite.intros} fun ins x th = Thm.implies_elim (Thm.instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg) (Thm.cprop_of th), SOME x] th1) th in fold ins u th0 end; fun simp_rule ctxt = Conv.fconv_rule (Conv.arg_conv (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms ball_simps simp_thms}))); fun basic_dloqe ctxt stupid dlo_qeth dlo_qeth_nolb dlo_qeth_noub gather ep = (case Thm.term_of ep of Const (\<^const_name>\Ex\, _) $ _ => let val p = Thm.dest_arg ep val ths = simplify (put_simpset HOL_basic_ss ctxt addsimps gather) (Thm.instantiate' [] [SOME p] stupid) val (L, U) = let val (_, q) = Thm.dest_abs NONE (Thm.dest_arg (Thm.rhs_of ths)) in (Thm.dest_arg1 q |> Thm.dest_arg1, Thm.dest_arg q |> Thm.dest_arg1) end fun proveneF S = let val (a, A) = Thm.dest_comb S |>> Thm.dest_arg val cT = Thm.ctyp_of_cterm a val ne = Thm.instantiate' [SOME cT] [SOME a, SOME A] @{thm insert_not_empty} val f = prove_finite cT (dest_set S) in (ne, f) end val qe = (case (Thm.term_of L, Thm.term_of U) of (Const (\<^const_name>\Orderings.bot\, _),_) => let val (neU, fU) = proveneF U in simp_rule ctxt (Thm.transitive ths (dlo_qeth_nolb OF [neU, fU])) end | (_, Const (\<^const_name>\Orderings.bot\, _)) => let val (neL,fL) = proveneF L in simp_rule ctxt (Thm.transitive ths (dlo_qeth_noub OF [neL, fL])) end | _ => let val (neL, fL) = proveneF L val (neU, fU) = proveneF U in simp_rule ctxt (Thm.transitive ths (dlo_qeth OF [neL, neU, fL, fU])) end) in qe end | _ => error "dlo_qe : Not an existential formula"); val all_conjuncts = let fun h acc ct = (case Thm.term_of ct of \<^term>\HOL.conj\ $ _ $ _ => h (h acc (Thm.dest_arg ct)) (Thm.dest_arg1 ct) | _ => ct :: acc) in h [] end; fun conjuncts ct = (case Thm.term_of ct of \<^term>\HOL.conj\ $ _ $ _ => Thm.dest_arg1 ct :: conjuncts (Thm.dest_arg ct) | _ => [ct]); fun fold1 f = foldr1 (uncurry f); (* FIXME !? *) val list_conj = fold1 (fn c => fn c' => Thm.apply (Thm.apply \<^cterm>\HOL.conj\ c) c'); fun mk_conj_tab th = let fun h acc th = (case Thm.prop_of th of \<^term>\Trueprop\ $ (\<^term>\HOL.conj\ $ p $ q) => h (h acc (th RS conjunct2)) (th RS conjunct1) | \<^term>\Trueprop\ $ p => (p, th) :: acc) in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end; fun is_conj (\<^term>\HOL.conj\$_$_) = true | is_conj _ = false; fun prove_conj tab cjs = (case cjs of [c] => if is_conj (Thm.term_of c) then prove_conj tab (conjuncts c) else tab c | c :: cs => conjI OF [prove_conj tab [c], prove_conj tab cs]); fun conj_aci_rule eq = let val (l, r) = Thm.dest_equals eq fun tabl c = the (Termtab.lookup (mk_conj_tab (Thm.assume l)) (Thm.term_of c)) fun tabr c = the (Termtab.lookup (mk_conj_tab (Thm.assume r)) (Thm.term_of c)) val ll = Thm.dest_arg l val rr = Thm.dest_arg r val thl = prove_conj tabl (conjuncts rr) |> Drule.implies_intr_hyps val thr = prove_conj tabr (conjuncts ll) |> Drule.implies_intr_hyps val eqI = Thm.instantiate' [] [SOME ll, SOME rr] @{thm iffI} in Thm.implies_elim (Thm.implies_elim eqI thl) thr |> mk_meta_eq end; fun contains x ct = member (op aconv) (Misc_Legacy.term_frees (Thm.term_of ct)) (Thm.term_of x); fun is_eqx x eq = (case Thm.term_of eq of Const (\<^const_name>\HOL.eq\, _) $ l $ r => l aconv Thm.term_of x orelse r aconv Thm.term_of x | _ => false); local fun proc ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\Ex\, _) $ Abs (xn, _, _) => let val e = Thm.dest_fun ct val (x,p) = Thm.dest_abs (SOME xn) (Thm.dest_arg ct) val Pp = Thm.apply \<^cterm>\Trueprop\ p val (eqs,neqs) = List.partition (is_eqx x) (all_conjuncts p) in (case eqs of [] => let val (dx, ndx) = List.partition (contains x) neqs in case ndx of [] => NONE | _ => conj_aci_rule (Thm.mk_binop \<^cterm>\(\) :: prop => _\ Pp (Thm.apply \<^cterm>\Trueprop\ (list_conj (ndx @ dx)))) |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e |> Conv.fconv_rule (Conv.arg_conv (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps}))) |> SOME end | _ => conj_aci_rule (Thm.mk_binop \<^cterm>\(\) :: prop => _\ Pp (Thm.apply \<^cterm>\Trueprop\ (list_conj (eqs @ neqs)))) |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e |> Conv.fconv_rule (Conv.arg_conv (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps}))) |> SOME) end | _ => NONE); in val reduce_ex_simproc = Simplifier.make_simproc \<^context> "reduce_ex_simproc" {lhss = [\<^term>\\x. P x\], proc = K proc}; end; fun raw_dlo_conv ctxt dlo_ss ({qe_bnds, qe_nolb, qe_noub, gst, gs, ...}: Langford_Data.entry) = let val ctxt' = Context_Position.set_visible false (put_simpset dlo_ss ctxt) addsimps @{thms dnf_simps} addsimprocs [reduce_ex_simproc] val dnfex_conv = Simplifier.rewrite ctxt' val pcv = Simplifier.rewrite (put_simpset dlo_ss ctxt addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib}) in fn p => Qelim.gen_qelim_conv ctxt pcv pcv dnfex_conv cons - (Drule.cterm_add_frees p []) (K Thm.reflexive) (K Thm.reflexive) + (Misc_Legacy.cterm_frees p) (K Thm.reflexive) (K Thm.reflexive) (K (basic_dloqe ctxt gst qe_bnds qe_nolb qe_noub gs)) p end; val grab_atom_bop = let fun h bounds tm = (case Thm.term_of tm of Const (\<^const_name>\HOL.eq\, T) $ _ $ _ => if domain_type T = HOLogic.boolT then find_args bounds tm else Thm.dest_fun2 tm | Const (\<^const_name>\Not\, _) $ _ => h bounds (Thm.dest_arg tm) | Const (\<^const_name>\All\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\Pure.all\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\Ex\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\HOL.conj\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\HOL.disj\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\HOL.implies\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Pure.imp\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Trueprop\, _) $ _ => h bounds (Thm.dest_arg tm) | _ => Thm.dest_fun2 tm) and find_args bounds tm = (h bounds (Thm.dest_arg tm) handle CTERM _ => h bounds (Thm.dest_arg1 tm)) and find_body bounds b = let val (_, b') = Thm.dest_abs (SOME (Name.bound bounds)) b in h (bounds + 1) b' end; in h end; fun dlo_instance ctxt tm = (fst (Langford_Data.get ctxt), Langford_Data.match ctxt (grab_atom_bop 0 tm)); fun dlo_conv ctxt tm = (case dlo_instance ctxt tm of (_, NONE) => raise CTERM ("dlo_conv (langford): no corresponding instance in context!", [tm]) | (ss, SOME instance) => raw_dlo_conv ctxt ss instance tm); fun generalize_tac ctxt f = CSUBGOAL (fn (p, _) => PRIMITIVE (fn st => let fun all x t = Thm.apply (Thm.cterm_of ctxt (Logic.all_const (Thm.typ_of_cterm x))) (Thm.lambda x t) val ts = sort Thm.fast_term_ord (f p) val p' = fold_rev all ts p in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end)); fun cfrees ats ct = let val ins = insert (op aconvc) fun h acc t = (case Thm.term_of t of _ $ _ $ _ => if member (op aconvc) ats (Thm.dest_fun2 t) then ins (Thm.dest_arg t) (ins (Thm.dest_arg1 t) acc) else h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) | _ $ _ => h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) | Abs _ => Thm.dest_abs NONE t ||> h acc |> uncurry (remove (op aconvc)) | Free _ => if member (op aconvc) ats t then acc else ins t acc | Var _ => if member (op aconvc) ats t then acc else ins t acc | _ => acc) in h [] ct end fun dlo_tac ctxt = CSUBGOAL (fn (p, i) => (case dlo_instance ctxt p of (ss, NONE) => simp_tac (put_simpset ss ctxt) i | (ss, SOME instance) => Object_Logic.full_atomize_tac ctxt i THEN simp_tac (put_simpset ss ctxt) i THEN (CONVERSION Thm.eta_long_conversion) i THEN (TRY o generalize_tac ctxt (cfrees (#atoms instance))) i THEN Object_Logic.full_atomize_tac ctxt i THEN CONVERSION (Object_Logic.judgment_conv ctxt (raw_dlo_conv ctxt ss instance)) i THEN (simp_tac (put_simpset ss ctxt) i))); end; diff --git a/src/HOL/Eisbach/eisbach_rule_insts.ML b/src/HOL/Eisbach/eisbach_rule_insts.ML --- a/src/HOL/Eisbach/eisbach_rule_insts.ML +++ b/src/HOL/Eisbach/eisbach_rule_insts.ML @@ -1,233 +1,233 @@ (* Title: HOL/Eisbach/eisbach_rule_insts.ML Author: Daniel Matichuk, NICTA/UNSW Eisbach-aware variants of the "where" and "of" attributes. Alternate syntax for rule_insts.ML participates in token closures by examining the behaviour of Rule_Insts.where_rule and instantiating token values accordingly. Instantiations in re-interpretation are done with infer_instantiate. *) structure Eisbach_Rule_Insts: sig end = struct fun restore_tags thm = Thm.map_tags (K (Thm.get_tags thm)); val mk_term_type_internal = Logic.protect o Logic.mk_term o Logic.mk_type; fun term_type_cases f g t = (case \<^try>\Logic.dest_type (Logic.dest_term (Logic.unprotect t))\ of SOME T => f T | NONE => (case \<^try>\Logic.dest_term t\ of SOME t => g t | NONE => raise Fail "Lost encoded instantiation")) fun add_thm_insts ctxt thm = let - val tyvars = Thm.fold_terms Term.add_tvars thm []; + val tyvars = Thm.fold_terms {hyps = false} Term.add_tvars thm []; val tyvars' = tyvars |> map (mk_term_type_internal o TVar); - val tvars = Thm.fold_terms Term.add_vars thm []; + val tvars = Thm.fold_terms {hyps = false} Term.add_vars thm []; val tvars' = tvars |> map (Logic.mk_term o Var); val conj = Logic.mk_conjunction_list (tyvars' @ tvars') |> Thm.cterm_of ctxt |> Drule.mk_term; in ((tyvars, tvars), Conjunction.intr thm conj) end; fun get_thm_insts thm = let val (thm', insts) = Conjunction.elim thm; val insts' = insts |> Drule.dest_term |> Thm.term_of |> Logic.dest_conjunction_list |> (fn f => fold (fn t => fn (tys, ts) => term_type_cases (fn T => (T :: tys, ts)) (fn t => (tys, t :: ts)) t) f ([], [])) ||> rev |>> rev; in (thm', insts') end; fun instantiate_xis ctxt insts thm = let - val tyvars = Thm.fold_terms Term.add_tvars thm []; - val tvars = Thm.fold_terms Term.add_vars thm []; + val tyvars = Thm.fold_terms {hyps = false} Term.add_tvars thm []; + val tvars = Thm.fold_terms {hyps = false} Term.add_vars thm []; fun add_inst (xi, t) (Ts, ts) = (case AList.lookup (op =) tyvars xi of SOME S => (((xi, S), Thm.ctyp_of ctxt (Logic.dest_type t)) :: Ts, ts) | NONE => (case AList.lookup (op =) tvars xi of SOME _ => (Ts, (xi, Thm.cterm_of ctxt t) :: ts) | NONE => error "indexname not found in thm")); val (instT, inst) = fold add_inst insts ([], []); in (Thm.instantiate (instT, []) thm |> infer_instantiate ctxt inst COMP_INCR asm_rl) |> Thm.adjust_maxidx_thm ~1 |> restore_tags thm end; datatype rule_inst = Named_Insts of ((indexname * string) * (term -> unit)) list * (binding * string option * mixfix) list | Term_Insts of (indexname * term) list | Unchecked_Term_Insts of term option list * term option list; fun mk_pair (t, t') = Logic.mk_conjunction (Logic.mk_term t, Logic.mk_term t'); fun dest_pair t = apply2 Logic.dest_term (Logic.dest_conjunction t); fun embed_indexname ((xi, s), f) = let fun wrap_xi xi t = mk_pair (Var (xi, fastype_of t), t); in ((xi, s), f o wrap_xi xi) end; fun unembed_indexname t = dest_pair t |> apfst (Term.dest_Var #> fst); fun read_where_insts (insts, fixes) = let val insts' = if forall (fn (_, v) => Parse_Tools.is_real_val v) insts then Term_Insts (map (unembed_indexname o Parse_Tools.the_real_val o snd) insts) else Named_Insts (map (fn (xi, p) => embed_indexname ((xi, Parse_Tools.the_parse_val p), Parse_Tools.the_parse_fun p)) insts, fixes); in insts' end; fun of_rule thm (args, concl_args) = let fun zip_vars _ [] = [] | zip_vars (_ :: xs) (NONE :: rest) = zip_vars xs rest | zip_vars ((x, _) :: xs) (SOME t :: rest) = (x, t) :: zip_vars xs rest | zip_vars [] _ = error "More instantiations than variables in theorem"; val insts = zip_vars (rev (Term.add_vars (Thm.full_prop_of thm) [])) args @ zip_vars (rev (Term.add_vars (Thm.concl_of thm) [])) concl_args; in insts end; val inst = Args.maybe Parse_Tools.name_term; val concl = Args.$$$ "concl" -- Args.colon; fun close_unchecked_insts context ((insts, concl_inst), fixes) = let val ctxt = Context.proof_of context; val ctxt1 = ctxt |> Proof_Context.add_fixes_cmd fixes |> #2; val insts' = insts @ concl_inst; val term_insts = map (the_list o (Option.map Parse_Tools.the_parse_val)) insts' |> burrow (Syntax.read_terms ctxt1 #> Variable.export_terms ctxt1 ctxt) |> map (try the_single); val _ = (insts', term_insts) |> ListPair.app (fn (SOME p, SOME t) => Parse_Tools.the_parse_fun p t | _ => ()); val (insts'', concl_insts'') = chop (length insts) term_insts; in Unchecked_Term_Insts (insts'', concl_insts'') end; fun read_of_insts checked context ((insts, concl_insts), fixes) = if forall (fn SOME t => Parse_Tools.is_real_val t | NONE => true) (insts @ concl_insts) then if checked then (fn _ => Term_Insts (map (unembed_indexname o Parse_Tools.the_real_val) (map_filter I (insts @ concl_insts)))) else (fn _ => Unchecked_Term_Insts (map (Option.map Parse_Tools.the_real_val) insts, map (Option.map Parse_Tools.the_real_val) concl_insts)) else if checked then (fn thm => Named_Insts (apply2 (map (Option.map (fn p => (Parse_Tools.the_parse_val p, Parse_Tools.the_parse_fun p)))) (insts, concl_insts) |> of_rule thm |> map ((fn (xi, (nm, f)) => embed_indexname ((xi, nm), f))), fixes)) else let val result = close_unchecked_insts context ((insts, concl_insts), fixes); in fn _ => result end; fun read_instantiate_closed ctxt (Named_Insts (insts, fixes)) thm = let val insts' = map (fn ((v, t), _) => ((v, Position.none), t)) insts; val (thm_insts, thm') = add_thm_insts ctxt thm; val (thm'', thm_insts') = Rule_Insts.where_rule ctxt insts' fixes thm' |> get_thm_insts; val tyinst = ListPair.zip (fst thm_insts, fst thm_insts') |> map (fn ((xi, _), typ) => (xi, typ)); val tinst = ListPair.zip (snd thm_insts, snd thm_insts') |> map (fn ((xi, _), t) => (xi, t)); val _ = map (fn ((xi, _), f) => (case AList.lookup (op =) tyinst xi of SOME typ => f (Logic.mk_type typ) | NONE => (case AList.lookup (op =) tinst xi of SOME t => f t | NONE => error "Lost indexname in instantiated theorem"))) insts; in (thm'' |> restore_tags thm) end | read_instantiate_closed ctxt (Unchecked_Term_Insts insts) thm = let val (xis, ts) = ListPair.unzip (of_rule thm insts); val ctxt' = Variable.declare_maxidx (Thm.maxidx_of thm) ctxt; val (ts', ctxt'') = Variable.import_terms false ts ctxt'; val ts'' = Variable.export_terms ctxt'' ctxt ts'; val insts' = ListPair.zip (xis, ts''); in instantiate_xis ctxt insts' thm end | read_instantiate_closed ctxt (Term_Insts insts) thm = instantiate_xis ctxt insts thm; val _ = Theory.setup (Attrib.setup \<^binding>\where\ (Scan.lift (Parse.and_list1 (Args.var -- (Args.$$$ "=" |-- Parse_Tools.name_term)) -- Parse.for_fixes) >> (fn args => let val args' = read_where_insts args in fn (context, thm) => (NONE, SOME (read_instantiate_closed (Context.proof_of context) args' thm)) end)) "named instantiation of theorem"); val _ = Theory.setup (Attrib.setup \<^binding>\of\ (Scan.lift (Args.mode "unchecked" -- (Scan.repeat (Scan.unless concl inst) -- Scan.optional (concl |-- Scan.repeat inst) [] -- Parse.for_fixes)) -- Scan.state >> (fn ((unchecked, args), context) => let val read_insts = read_of_insts (not unchecked) context args; in fn (context, thm) => let val thm' = if Thm.is_free_dummy thm andalso unchecked then Drule.free_dummy_thm else read_instantiate_closed (Context.proof_of context) (read_insts thm) thm in (NONE, SOME thm') end end)) "positional instantiation of theorem"); end; diff --git a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML --- a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML +++ b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML @@ -1,1034 +1,1036 @@ (* Title: HOL/Library/Sum_of_Squares/sum_of_squares.ML Author: Amine Chaieb, University of Cambridge Author: Philipp Meyer, TU Muenchen A tactic for proving nonlinear inequalities. *) signature SUM_OF_SQUARES = sig datatype proof_method = Certificate of RealArith.pss_tree | Prover of string -> string val sos_tac: (RealArith.pss_tree -> unit) -> proof_method -> Proof.context -> int -> tactic val trace: bool Config.T val debug: bool Config.T val trace_message: Proof.context -> (unit -> string) -> unit val debug_message: Proof.context -> (unit -> string) -> unit exception Failure of string; end structure Sum_of_Squares: SUM_OF_SQUARES = struct val max = Integer.max; val denominator_rat = Rat.dest #> snd #> Rat.of_int; fun int_of_rat a = (case Rat.dest a of (i, 1) => i | _ => error "int_of_rat: not an int"); fun lcm_rat x y = Rat.of_int (Integer.lcm (int_of_rat x) (int_of_rat y)); fun rat_pow r i = let fun pow r i = if i = 0 then @1 else let val d = pow r (i div 2) in d * d * (if i mod 2 = 0 then @1 else r) end in if i < 0 then pow (Rat.inv r) (~ i) else pow r i end; fun round_rat r = let val (a,b) = Rat.dest (abs r) val d = a div b val s = if r < @0 then ~ o Rat.of_int else Rat.of_int val x2 = 2 * (a - (b * d)) in s (if x2 >= b then d + 1 else d) end val trace = Attrib.setup_config_bool \<^binding>\sos_trace\ (K false); val debug = Attrib.setup_config_bool \<^binding>\sos_debug\ (K false); fun trace_message ctxt msg = if Config.get ctxt trace orelse Config.get ctxt debug then tracing (msg ()) else (); fun debug_message ctxt msg = if Config.get ctxt debug then tracing (msg ()) else (); exception Sanity; exception Unsolvable; exception Failure of string; datatype proof_method = Certificate of RealArith.pss_tree | Prover of (string -> string) (* Turn a rational into a decimal string with d sig digits. *) local fun normalize y = if abs y < @1/10 then normalize (@10 * y) - 1 else if abs y >= @1 then normalize (y / @10) + 1 else 0 in fun decimalize d x = if x = @0 then "0.0" else let val y = abs x val e = normalize y val z = rat_pow @10 (~ e) * y + @1 val k = int_of_rat (round_rat (rat_pow @10 d * z)) in (if x < @0 then "-0." else "0.") ^ implode (tl (raw_explode(string_of_int k))) ^ (if e = 0 then "" else "e" ^ string_of_int e) end end; (* Iterations over numbers, and lists indexed by numbers. *) fun itern k l f a = (case l of [] => a | h::t => itern (k + 1) t f (f h k a)); fun iter (m,n) f a = if n < m then a else iter (m + 1, n) f (f m a); (* The main types. *) type vector = int * Rat.rat FuncUtil.Intfunc.table; type matrix = (int * int) * Rat.rat FuncUtil.Intpairfunc.table; fun iszero (_, r) = r = @0; (* Vectors. Conventionally indexed 1..n. *) fun vector_0 n = (n, FuncUtil.Intfunc.empty): vector; fun dim (v: vector) = fst v; fun vector_cmul c (v: vector) = let val n = dim v in if c = @0 then vector_0 n else (n,FuncUtil.Intfunc.map (fn _ => fn x => c * x) (snd v)) end; fun vector_of_list l = let val n = length l in (n, fold_rev FuncUtil.Intfunc.update (1 upto n ~~ l) FuncUtil.Intfunc.empty): vector end; (* Matrices; again rows and columns indexed from 1. *) fun dimensions (m: matrix) = fst m; fun row k (m: matrix) : vector = let val (_, j) = dimensions m in (j, FuncUtil.Intpairfunc.fold (fn ((i, j), c) => fn a => if i = k then FuncUtil.Intfunc.update (j, c) a else a) (snd m) FuncUtil.Intfunc.empty) end; (* Monomials. *) fun monomial_eval assig m = FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => a * rat_pow (FuncUtil.Ctermfunc.apply assig x) k) m @1; val monomial_1 = FuncUtil.Ctermfunc.empty; fun monomial_var x = FuncUtil.Ctermfunc.onefunc (x, 1); val monomial_mul = FuncUtil.Ctermfunc.combine Integer.add (K false); fun monomial_multidegree m = FuncUtil.Ctermfunc.fold (fn (_, k) => fn a => k + a) m 0; fun monomial_variables m = FuncUtil.Ctermfunc.dom m; (* Polynomials. *) fun eval assig p = FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => a + c * monomial_eval assig m) p @0; val poly_0 = FuncUtil.Monomialfunc.empty; fun poly_isconst p = FuncUtil.Monomialfunc.fold (fn (m, _) => fn a => FuncUtil.Ctermfunc.is_empty m andalso a) p true; fun poly_var x = FuncUtil.Monomialfunc.onefunc (monomial_var x, @1); fun poly_const c = if c = @0 then poly_0 else FuncUtil.Monomialfunc.onefunc (monomial_1, c); fun poly_cmul c p = if c = @0 then poly_0 else FuncUtil.Monomialfunc.map (fn _ => fn x => c * x) p; fun poly_neg p = FuncUtil.Monomialfunc.map (K ~) p; fun poly_add p1 p2 = FuncUtil.Monomialfunc.combine (curry op +) (fn x => x = @0) p1 p2; fun poly_sub p1 p2 = poly_add p1 (poly_neg p2); fun poly_cmmul (c,m) p = if c = @0 then poly_0 else if FuncUtil.Ctermfunc.is_empty m then FuncUtil.Monomialfunc.map (fn _ => fn d => c * d) p else FuncUtil.Monomialfunc.fold (fn (m', d) => fn a => (FuncUtil.Monomialfunc.update (monomial_mul m m', c * d) a)) p poly_0; fun poly_mul p1 p2 = FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => poly_add (poly_cmmul (c,m) p2) a) p1 poly_0; fun poly_square p = poly_mul p p; fun poly_pow p k = if k = 0 then poly_const @1 else if k = 1 then p else let val q = poly_square(poly_pow p (k div 2)) in if k mod 2 = 1 then poly_mul p q else q end; fun multidegree p = FuncUtil.Monomialfunc.fold (fn (m, _) => fn a => max (monomial_multidegree m) a) p 0; fun poly_variables p = sort Thm.fast_term_ord (FuncUtil.Monomialfunc.fold_rev (fn (m, _) => union (is_equal o Thm.fast_term_ord) (monomial_variables m)) p []); (* Conversion from HOL term. *) local val neg_tm = \<^cterm>\uminus :: real \ _\ val add_tm = \<^cterm>\(+) :: real \ _\ val sub_tm = \<^cterm>\(-) :: real \ _\ val mul_tm = \<^cterm>\(*) :: real \ _\ val inv_tm = \<^cterm>\inverse :: real \ _\ val div_tm = \<^cterm>\(/) :: real \ _\ val pow_tm = \<^cterm>\(^) :: real \ _\ val zero_tm = \<^cterm>\0:: real\ val is_numeral = can (HOLogic.dest_number o Thm.term_of) fun poly_of_term tm = if tm aconvc zero_tm then poly_0 else if RealArith.is_ratconst tm then poly_const(RealArith.dest_ratconst tm) else (let val (lop, r) = Thm.dest_comb tm in if lop aconvc neg_tm then poly_neg(poly_of_term r) else if lop aconvc inv_tm then let val p = poly_of_term r in if poly_isconst p then poly_const(Rat.inv (eval FuncUtil.Ctermfunc.empty p)) else error "poly_of_term: inverse of non-constant polyomial" end else (let val (opr,l) = Thm.dest_comb lop in if opr aconvc pow_tm andalso is_numeral r then poly_pow (poly_of_term l) ((snd o HOLogic.dest_number o Thm.term_of) r) else if opr aconvc add_tm then poly_add (poly_of_term l) (poly_of_term r) else if opr aconvc sub_tm then poly_sub (poly_of_term l) (poly_of_term r) else if opr aconvc mul_tm then poly_mul (poly_of_term l) (poly_of_term r) else if opr aconvc div_tm then let val p = poly_of_term l val q = poly_of_term r in if poly_isconst q then poly_cmul (Rat.inv (eval FuncUtil.Ctermfunc.empty q)) p else error "poly_of_term: division by non-constant polynomial" end else poly_var tm end handle CTERM ("dest_comb",_) => poly_var tm) end handle CTERM ("dest_comb",_) => poly_var tm) in val poly_of_term = fn tm => if type_of (Thm.term_of tm) = \<^typ>\real\ then poly_of_term tm else error "poly_of_term: term does not have real type" end; (* String of vector (just a list of space-separated numbers). *) fun sdpa_of_vector (v: vector) = let val n = dim v val strs = map (decimalize 20 o (fn i => FuncUtil.Intfunc.tryapplyd (snd v) i @0)) (1 upto n) in space_implode " " strs ^ "\n" end; fun triple_int_ord ((a, b, c), (a', b', c')) = prod_ord int_ord (prod_ord int_ord int_ord) ((a, (b, c)), (a', (b', c'))); structure Inttriplefunc = FuncFun(type key = int * int * int val ord = triple_int_ord); (* Parse back csdp output. *) local val decimal_digits = Scan.many1 Symbol.is_ascii_digit val decimal_nat = decimal_digits >> (#1 o Library.read_int); val decimal_int = decimal_nat >> Rat.of_int; val decimal_sig = decimal_int -- Scan.option (Scan.$$ "." |-- decimal_digits) >> (fn (a, NONE) => a | (a, SOME bs) => a + Rat.of_int (#1 (Library.read_int bs)) / rat_pow @10 (length bs)); fun signed neg parse = $$ "-" |-- parse >> neg || $$ "+" |-- parse || parse; val exponent = ($$ "e" || $$ "E") |-- signed ~ decimal_nat; val decimal = signed ~ decimal_sig -- Scan.optional exponent 0 >> (fn (a, b) => a * rat_pow @10 b); val csdp_output = decimal -- Scan.repeat (Scan.$$ " " |-- Scan.option decimal) --| Scan.many Symbol.not_eof >> (fn (a, bs) => vector_of_list (a :: map_filter I bs)); in fun parse_csdpoutput s = Symbol.scanner "Malformed CSDP output" csdp_output (raw_explode s); end; (* Try some apparently sensible scaling first. Note that this is purely to *) (* get a cleaner translation to floating-point, and doesn't affect any of *) (* the results, in principle. In practice it seems a lot better when there *) (* are extreme numbers in the original problem. *) (* Version for (int*int*int) keys *) local fun max_rat x y = if x < y then y else x fun common_denominator fld amat acc = fld (fn (_,c) => fn a => lcm_rat (denominator_rat c) a) amat acc fun maximal_element fld amat acc = fld (fn (_,c) => fn maxa => max_rat maxa (abs c)) amat acc fun float_of_rat x = let val (a,b) = Rat.dest x in Real.fromInt a / Real.fromInt b end; fun int_of_float x = (trunc x handle Overflow => 0 | Domain => 0) in fun tri_scale_then solver (obj:vector) mats = let val cd1 = fold_rev (common_denominator Inttriplefunc.fold) mats @1 val cd2 = common_denominator FuncUtil.Intfunc.fold (snd obj) @1 val mats' = map (Inttriplefunc.map (fn _ => fn x => cd1 * x)) mats val obj' = vector_cmul cd2 obj val max1 = fold_rev (maximal_element Inttriplefunc.fold) mats' @0 val max2 = maximal_element FuncUtil.Intfunc.fold (snd obj') @0 val scal1 = rat_pow @2 (20 - int_of_float(Math.ln (float_of_rat max1) / Math.ln 2.0)) val scal2 = rat_pow @2 (20 - int_of_float(Math.ln (float_of_rat max2) / Math.ln 2.0)) val mats'' = map (Inttriplefunc.map (fn _ => fn x => x * scal1)) mats' val obj'' = vector_cmul scal2 obj' in solver obj'' mats'' end end; (* Round a vector to "nice" rationals. *) fun nice_rational n x = round_rat (n * x) / n; fun nice_vector n ((d,v) : vector) = (d, FuncUtil.Intfunc.fold (fn (i,c) => fn a => let val y = nice_rational n c in if c = @0 then a else FuncUtil.Intfunc.update (i,y) a end) v FuncUtil.Intfunc.empty): vector fun dest_ord f x = is_equal (f x); (* Stuff for "equations" ((int*int*int)->num functions). *) fun tri_equation_cmul c eq = if c = @0 then Inttriplefunc.empty else Inttriplefunc.map (fn _ => fn d => c * d) eq; fun tri_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +) (fn x => x = @0) eq1 eq2; fun tri_equation_eval assig eq = let fun value v = Inttriplefunc.apply assig v in Inttriplefunc.fold (fn (v, c) => fn a => a + value v * c) eq @0 end; (* Eliminate all variables, in an essentially arbitrary order. *) fun tri_eliminate_all_equations one = let fun choose_variable eq = let val (v,_) = Inttriplefunc.choose eq in if is_equal (triple_int_ord(v,one)) then let val eq' = Inttriplefunc.delete_safe v eq in if Inttriplefunc.is_empty eq' then error "choose_variable" else fst (Inttriplefunc.choose eq') end else v end fun eliminate dun eqs = (case eqs of [] => dun | eq :: oeqs => if Inttriplefunc.is_empty eq then eliminate dun oeqs else let val v = choose_variable eq val a = Inttriplefunc.apply eq v val eq' = tri_equation_cmul ((Rat.of_int ~1) / a) (Inttriplefunc.delete_safe v eq) fun elim e = let val b = Inttriplefunc.tryapplyd e v @0 in if b = @0 then e else tri_equation_add e (tri_equation_cmul (~ b / a) eq) end in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map (K elim) dun)) (map elim oeqs) end) in fn eqs => let val assig = eliminate Inttriplefunc.empty eqs val vs = Inttriplefunc.fold (fn (_, f) => fn a => remove (dest_ord triple_int_ord) one (Inttriplefunc.dom f) @ a) assig [] in (distinct (dest_ord triple_int_ord) vs,assig) end end; (* Multiply equation-parametrized poly by regular poly and add accumulator. *) fun tri_epoly_pmul p q acc = FuncUtil.Monomialfunc.fold (fn (m1, c) => fn a => FuncUtil.Monomialfunc.fold (fn (m2, e) => fn b => let val m = monomial_mul m1 m2 val es = FuncUtil.Monomialfunc.tryapplyd b m Inttriplefunc.empty in FuncUtil.Monomialfunc.update (m,tri_equation_add (tri_equation_cmul c e) es) b end) q a) p acc; (* Hence produce the "relevant" monomials: those whose squares lie in the *) (* Newton polytope of the monomials in the input. (This is enough according *) (* to Reznik: "Extremal PSD forms with few terms", Duke Math. Journal, *) (* vol 45, pp. 363--374, 1978. *) (* *) (* These are ordered in sort of decreasing degree. In particular the *) (* constant monomial is last; this gives an order in diagonalization of the *) (* quadratic form that will tend to display constants. *) (* Diagonalize (Cholesky/LDU) the matrix corresponding to a quadratic form. *) local fun diagonalize n i m = if FuncUtil.Intpairfunc.is_empty (snd m) then [] else let val a11 = FuncUtil.Intpairfunc.tryapplyd (snd m) (i,i) @0 in if a11 < @0 then raise Failure "diagonalize: not PSD" else if a11 = @0 then if FuncUtil.Intfunc.is_empty (snd (row i m)) then diagonalize n (i + 1) m else raise Failure "diagonalize: not PSD ___ " else let val v = row i m val v' = (fst v, FuncUtil.Intfunc.fold (fn (i, c) => fn a => let val y = c / a11 in if y = @0 then a else FuncUtil.Intfunc.update (i,y) a end) (snd v) FuncUtil.Intfunc.empty) fun upt0 x y a = if y = @0 then a else FuncUtil.Intpairfunc.update (x,y) a val m' = ((n, n), iter (i + 1, n) (fn j => iter (i + 1, n) (fn k => (upt0 (j, k) (FuncUtil.Intpairfunc.tryapplyd (snd m) (j, k) @0 - FuncUtil.Intfunc.tryapplyd (snd v) j @0 * FuncUtil.Intfunc.tryapplyd (snd v') k @0)))) FuncUtil.Intpairfunc.empty) in (a11, v') :: diagonalize n (i + 1) m' end end in fun diag m = let val nn = dimensions m val n = fst nn in if snd nn <> n then error "diagonalize: non-square matrix" else diagonalize n 1 m end end; (* Enumeration of monomials with given multidegree bound. *) fun enumerate_monomials d vars = if d < 0 then [] else if d = 0 then [FuncUtil.Ctermfunc.empty] else if null vars then [monomial_1] else let val alts = map_range (fn k => let val oths = enumerate_monomials (d - k) (tl vars) in map (fn ks => if k = 0 then ks else FuncUtil.Ctermfunc.update (hd vars, k) ks) oths end) (d + 1) in flat alts end; (* Enumerate products of distinct input polys with degree <= d. *) (* We ignore any constant input polynomials. *) (* Give the output polynomial and a record of how it was derived. *) fun enumerate_products d pols = if d = 0 then [(poly_const @1, RealArith.Rational_lt @1)] else if d < 0 then [] else (case pols of [] => [(poly_const @1, RealArith.Rational_lt @1)] | (p, b) :: ps => let val e = multidegree p in if e = 0 then enumerate_products d ps else enumerate_products d ps @ map (fn (q, c) => (poly_mul p q, RealArith.Product (b, c))) (enumerate_products (d - e) ps) end) (* Convert regular polynomial. Note that we treat (0,0,0) as -1. *) fun epoly_of_poly p = FuncUtil.Monomialfunc.fold (fn (m, c) => fn a => FuncUtil.Monomialfunc.update (m, Inttriplefunc.onefunc ((0, 0, 0), ~ c)) a) p FuncUtil.Monomialfunc.empty; (* String for block diagonal matrix numbered k. *) fun sdpa_of_blockdiagonal k m = let val pfx = string_of_int k ^" " val ents = Inttriplefunc.fold (fn ((b, i, j), c) => fn a => if i > j then a else ((b, i, j), c) :: a) m [] val entss = sort (triple_int_ord o apply2 fst) ents in fold_rev (fn ((b,i,j),c) => fn a => pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ " " ^ decimalize 20 c ^ "\n" ^ a) entss "" end; (* SDPA for problem using block diagonal (i.e. multiple SDPs) *) fun sdpa_of_blockproblem nblocks blocksizes obj mats = let val m = length mats - 1 in string_of_int m ^ "\n" ^ string_of_int nblocks ^ "\n" ^ (space_implode " " (map string_of_int blocksizes)) ^ "\n" ^ sdpa_of_vector obj ^ fold_rev (fn (k, m) => fn a => sdpa_of_blockdiagonal (k - 1) m ^ a) (1 upto length mats ~~ mats) "" end; (* Run prover on a problem in block diagonal form. *) fun run_blockproblem prover nblocks blocksizes obj mats = parse_csdpoutput (prover (sdpa_of_blockproblem nblocks blocksizes obj mats)) (* 3D versions of matrix operations to consider blocks separately. *) val bmatrix_add = Inttriplefunc.combine (curry op +) (fn x => x = @0); fun bmatrix_cmul c bm = if c = @0 then Inttriplefunc.empty else Inttriplefunc.map (fn _ => fn x => c * x) bm; val bmatrix_neg = bmatrix_cmul (Rat.of_int ~1); (* Smash a block matrix into components. *) fun blocks blocksizes bm = map (fn (bs, b0) => let val m = Inttriplefunc.fold (fn ((b, i, j), c) => fn a => if b = b0 then FuncUtil.Intpairfunc.update ((i, j), c) a else a) bm FuncUtil.Intpairfunc.empty val _ = FuncUtil.Intpairfunc.fold (fn ((i, j), _) => fn a => max a (max i j)) m 0 in (((bs, bs), m): matrix) end) (blocksizes ~~ (1 upto length blocksizes)); (* FIXME : Get rid of this !!!*) local fun tryfind_with msg _ [] = raise Failure msg | tryfind_with _ f (x::xs) = (f x handle Failure s => tryfind_with s f xs); in fun tryfind f = tryfind_with "tryfind" f end (* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *) fun real_positivnullstellensatz_general ctxt prover linf d eqs leqs pol = let val vars = fold_rev (union (op aconvc) o poly_variables) (pol :: eqs @ map fst leqs) [] val monoid = if linf then (poly_const @1, RealArith.Rational_lt @1):: (filter (fn (p,_) => multidegree p <= d) leqs) else enumerate_products d leqs val nblocks = length monoid fun mk_idmultiplier k p = let val e = d - multidegree p val mons = enumerate_monomials e vars val nons = mons ~~ (1 upto length mons) in (mons, fold_rev (fn (m, n) => FuncUtil.Monomialfunc.update (m, Inttriplefunc.onefunc ((~k, ~n, n), @1))) nons FuncUtil.Monomialfunc.empty) end fun mk_sqmultiplier k (p,_) = let val e = (d - multidegree p) div 2 val mons = enumerate_monomials e vars val nons = mons ~~ (1 upto length mons) in (mons, fold_rev (fn (m1, n1) => fold_rev (fn (m2, n2) => fn a => let val m = monomial_mul m1 m2 in if n1 > n2 then a else let val c = if n1 = n2 then @1 else @2 val e = FuncUtil.Monomialfunc.tryapplyd a m Inttriplefunc.empty in FuncUtil.Monomialfunc.update (m, tri_equation_add (Inttriplefunc.onefunc ((k, n1, n2), c)) e) a end end) nons) nons FuncUtil.Monomialfunc.empty) end val (sqmonlist,sqs) = split_list (map2 mk_sqmultiplier (1 upto length monoid) monoid) val (_(*idmonlist*),ids) = split_list (map2 mk_idmultiplier (1 upto length eqs) eqs) val blocksizes = map length sqmonlist val bigsum = fold_rev (fn (p, q) => fn a => tri_epoly_pmul p q a) (eqs ~~ ids) (fold_rev (fn ((p, _), s) => fn a => tri_epoly_pmul p s a) (monoid ~~ sqs) (epoly_of_poly (poly_neg pol))) val eqns = FuncUtil.Monomialfunc.fold (fn (_, e) => fn a => e :: a) bigsum [] val (pvs, assig) = tri_eliminate_all_equations (0, 0, 0) eqns val qvars = (0, 0, 0) :: pvs val allassig = fold_rev (fn v => Inttriplefunc.update (v, (Inttriplefunc.onefunc (v, @1)))) pvs assig fun mk_matrix v = Inttriplefunc.fold (fn ((b, i, j), ass) => fn m => if b < 0 then m else let val c = Inttriplefunc.tryapplyd ass v @0 in if c = @0 then m else Inttriplefunc.update ((b, j, i), c) (Inttriplefunc.update ((b, i, j), c) m) end) allassig Inttriplefunc.empty val diagents = Inttriplefunc.fold (fn ((b, i, j), e) => fn a => if b > 0 andalso i = j then tri_equation_add e a else a) allassig Inttriplefunc.empty val mats = map mk_matrix qvars val obj = (length pvs, itern 1 pvs (fn v => fn i => FuncUtil.Intfunc.updatep iszero (i,Inttriplefunc.tryapplyd diagents v @0)) FuncUtil.Intfunc.empty) val raw_vec = if null pvs then vector_0 0 else tri_scale_then (run_blockproblem prover nblocks blocksizes) obj mats fun int_element (_, v) i = FuncUtil.Intfunc.tryapplyd v i @0 fun find_rounding d = let val _ = debug_message ctxt (fn () => "Trying rounding with limit "^Rat.string_of_rat d ^ "\n") val vec = nice_vector d raw_vec val blockmat = iter (1, dim vec) (fn i => fn a => bmatrix_add (bmatrix_cmul (int_element vec i) (nth mats i)) a) (bmatrix_neg (nth mats 0)) val allmats = blocks blocksizes blockmat in (vec, map diag allmats) end val (vec, ratdias) = if null pvs then find_rounding @1 else tryfind find_rounding (map Rat.of_int (1 upto 31) @ map (rat_pow @2) (5 upto 66)) val newassigs = fold_rev (fn k => Inttriplefunc.update (nth pvs (k - 1), int_element vec k)) (1 upto dim vec) (Inttriplefunc.onefunc ((0, 0, 0), Rat.of_int ~1)) val finalassigs = Inttriplefunc.fold (fn (v, e) => fn a => Inttriplefunc.update (v, tri_equation_eval newassigs e) a) allassig newassigs fun poly_of_epoly p = FuncUtil.Monomialfunc.fold (fn (v, e) => fn a => FuncUtil.Monomialfunc.updatep iszero (v, tri_equation_eval finalassigs e) a) p FuncUtil.Monomialfunc.empty fun mk_sos mons = let fun mk_sq (c, m) = (c, fold_rev (fn k => fn a => FuncUtil.Monomialfunc.updatep iszero (nth mons (k - 1), int_element m k) a) (1 upto length mons) FuncUtil.Monomialfunc.empty) in map mk_sq end val sqs = map2 mk_sos sqmonlist ratdias val cfs = map poly_of_epoly ids val msq = filter (fn (_, b) => not (null b)) (map2 pair monoid sqs) fun eval_sq sqs = fold_rev (fn (c, q) => poly_add (poly_cmul c (poly_mul q q))) sqs poly_0 val sanity = fold_rev (fn ((p, _), s) => poly_add (poly_mul p (eval_sq s))) msq (fold_rev (fn (p, q) => poly_add (poly_mul p q)) (cfs ~~ eqs) (poly_neg pol)) in if not(FuncUtil.Monomialfunc.is_empty sanity) then raise Sanity else (cfs, map (fn (a, b) => (snd a, b)) msq) end (* Iterative deepening. *) fun deepen ctxt f n = (trace_message ctxt (fn () => "Searching with depth limit " ^ string_of_int n); (f n handle Failure s => (trace_message ctxt (fn () => "failed with message: " ^ s); deepen ctxt f (n + 1)))); (* Map back polynomials and their composites to a positivstellensatz. *) fun cterm_of_sqterm (c, p) = RealArith.Product (RealArith.Rational_lt c, RealArith.Square p); fun cterm_of_sos (pr,sqs) = if null sqs then pr else RealArith.Product (pr, foldr1 RealArith.Sum (map cterm_of_sqterm sqs)); (* Interface to HOL. *) local open Conv val concl = Thm.dest_arg o Thm.cprop_of in (* FIXME: Replace tryfind by get_first !! *) fun real_nonlinear_prover proof_method ctxt = let val {add = _, mul = _, neg = _, pow = _, sub = _, main = real_poly_conv} = Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt (the (Semiring_Normalizer.match ctxt \<^cterm>\(0::real) + 1\)) Thm.term_ord fun mainf cert_choice translator (eqs, les, lts) = let val eq0 = map (poly_of_term o Thm.dest_arg1 o concl) eqs val le0 = map (poly_of_term o Thm.dest_arg o concl) les val lt0 = map (poly_of_term o Thm.dest_arg o concl) lts val eqp0 = map_index (fn (i, t) => (t,RealArith.Axiom_eq i)) eq0 val lep0 = map_index (fn (i, t) => (t,RealArith.Axiom_le i)) le0 val ltp0 = map_index (fn (i, t) => (t,RealArith.Axiom_lt i)) lt0 val (keq,eq) = List.partition (fn (p, _) => multidegree p = 0) eqp0 val (klep,lep) = List.partition (fn (p, _) => multidegree p = 0) lep0 val (kltp,ltp) = List.partition (fn (p, _) => multidegree p = 0) ltp0 fun trivial_axiom (p, ax) = (case ax of RealArith.Axiom_eq n => if eval FuncUtil.Ctermfunc.empty p <> @0 then nth eqs n else raise Failure "trivial_axiom: Not a trivial axiom" | RealArith.Axiom_le n => if eval FuncUtil.Ctermfunc.empty p < @0 then nth les n else raise Failure "trivial_axiom: Not a trivial axiom" | RealArith.Axiom_lt n => if eval FuncUtil.Ctermfunc.empty p <= @0 then nth lts n else raise Failure "trivial_axiom: Not a trivial axiom" | _ => error "trivial_axiom: Not a trivial axiom") in let val th = tryfind trivial_axiom (keq @ klep @ kltp) in (fconv_rule (arg_conv (arg1_conv (real_poly_conv ctxt)) then_conv Numeral_Simprocs.field_comp_conv ctxt) th, RealArith.Trivial) end handle Failure _ => let val proof = (case proof_method of Certificate certs => (* choose certificate *) let fun chose_cert [] (RealArith.Cert c) = c | chose_cert (RealArith.Left::s) (RealArith.Branch (l, _)) = chose_cert s l | chose_cert (RealArith.Right::s) (RealArith.Branch (_, r)) = chose_cert s r | chose_cert _ _ = error "certificate tree in invalid form" in chose_cert cert_choice certs end | Prover prover => (* call prover *) let val pol = fold_rev poly_mul (map fst ltp) (poly_const @1) val leq = lep @ ltp fun tryall d = let val e = multidegree pol val k = if e = 0 then 0 else d div e val eq' = map fst eq in tryfind (fn i => (d, i, real_positivnullstellensatz_general ctxt prover false d eq' leq (poly_neg(poly_pow pol i)))) (0 upto k) end val (_,i,(cert_ideal,cert_cone)) = deepen ctxt tryall 0 val proofs_ideal = map2 (fn q => fn (_,ax) => RealArith.Eqmul(q,ax)) cert_ideal eq val proofs_cone = map cterm_of_sos cert_cone val proof_ne = if null ltp then RealArith.Rational_lt @1 else let val p = foldr1 RealArith.Product (map snd ltp) in funpow i (fn q => RealArith.Product (p, q)) (RealArith.Rational_lt @1) end in foldr1 RealArith.Sum (proof_ne :: proofs_ideal @ proofs_cone) end) in (translator (eqs,les,lts) proof, RealArith.Cert proof) end end in mainf end end (* FIXME : This is very bad!!!*) fun subst_conv eqs t = let val t' = fold (Thm.lambda o Thm.lhs_of) eqs t in Conv.fconv_rule (Thm.beta_conversion true) (fold (fn a => fn b => Thm.combination b a) eqs (Thm.reflexive t')) end (* A wrapper that tries to substitute away variables first. *) local open Conv val concl = Thm.dest_arg o Thm.cprop_of val shuffle1 = fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: field_simps)}) val shuffle2 = fconv_rule (rewr_conv @{lemma "(x + a == y) == (x == y - (a::real))" by (atomize (full)) (simp add: field_simps)}) fun substitutable_monomial fvs tm = (case Thm.term_of tm of Free (_, \<^typ>\real\) => if not (member (op aconvc) fvs tm) then (@1, tm) else raise Failure "substitutable_monomial" | \<^term>\(*) :: real \ _\ $ _ $ (Free _) => if RealArith.is_ratconst (Thm.dest_arg1 tm) andalso not (member (op aconvc) fvs (Thm.dest_arg tm)) then (RealArith.dest_ratconst (Thm.dest_arg1 tm), Thm.dest_arg tm) else raise Failure "substitutable_monomial" | \<^term>\(+) :: real \ _\$_$_ => - (substitutable_monomial (Drule.cterm_add_frees (Thm.dest_arg tm) fvs) (Thm.dest_arg1 tm) + (substitutable_monomial (merge (op aconvc) (fvs, Misc_Legacy.cterm_frees (Thm.dest_arg tm))) + (Thm.dest_arg1 tm) handle Failure _ => - substitutable_monomial (Drule.cterm_add_frees (Thm.dest_arg1 tm) fvs) (Thm.dest_arg tm)) + substitutable_monomial (merge (op aconvc) (fvs, Misc_Legacy.cterm_frees (Thm.dest_arg1 tm))) + (Thm.dest_arg tm)) | _ => raise Failure "substitutable_monomial") fun isolate_variable v th = let val w = Thm.dest_arg1 (Thm.cprop_of th) in if v aconvc w then th else (case Thm.term_of w of \<^term>\(+) :: real \ _\ $ _ $ _ => if Thm.dest_arg1 w aconvc v then shuffle2 th else isolate_variable v (shuffle1 th) | _ => error "isolate variable : This should not happen?") end in fun real_nonlinear_subst_prover prover ctxt = let val {add = _, mul = real_poly_mul_conv, neg = _, pow = _, sub = _, main = real_poly_conv} = Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt (the (Semiring_Normalizer.match ctxt \<^cterm>\(0::real) + 1\)) Thm.term_ord fun make_substitution th = let val (c,v) = substitutable_monomial [] (Thm.dest_arg1(concl th)) val th1 = Drule.arg_cong_rule (Thm.apply \<^cterm>\(*) :: real \ _\ (RealArith.cterm_of_rat (Rat.inv c))) (mk_meta_eq th) val th2 = fconv_rule (binop_conv (real_poly_mul_conv ctxt)) th1 in fconv_rule (arg_conv (real_poly_conv ctxt)) (isolate_variable v th2) end fun oprconv cv ct = let val g = Thm.dest_fun2 ct in if g aconvc \<^cterm>\(\) :: real \ _\ orelse g aconvc \<^cterm>\(<) :: real \ _\ then arg_conv cv ct else arg1_conv cv ct end fun mainf cert_choice translator = let fun substfirst (eqs, les, lts) = (let val eth = tryfind make_substitution eqs val modify = fconv_rule (arg_conv (oprconv(subst_conv [eth] then_conv (real_poly_conv ctxt)))) in substfirst (filter_out (fn t => (Thm.dest_arg1 o Thm.dest_arg o Thm.cprop_of) t aconvc \<^cterm>\0::real\) (map modify eqs), map modify les, map modify lts) end handle Failure _ => real_nonlinear_prover prover ctxt cert_choice translator (rev eqs, rev les, rev lts)) in substfirst end in mainf end (* Overall function. *) fun real_sos prover ctxt = RealArith.gen_prover_real_arith ctxt (real_nonlinear_subst_prover prover ctxt) end; val known_sos_constants = [\<^term>\(\)\, \<^term>\Trueprop\, \<^term>\HOL.False\, \<^term>\HOL.implies\, \<^term>\HOL.conj\, \<^term>\HOL.disj\, \<^term>\Not\, \<^term>\(=) :: bool \ _\, \<^term>\All :: (real \ _) \ _\, \<^term>\Ex :: (real \ _) \ _\, \<^term>\(=) :: real \ _\, \<^term>\(<) :: real \ _\, \<^term>\(\) :: real \ _\, \<^term>\(+) :: real \ _\, \<^term>\(-) :: real \ _\, \<^term>\(*) :: real \ _\, \<^term>\uminus :: real \ _\, \<^term>\(/) :: real \ _\, \<^term>\inverse :: real \ _\, \<^term>\(^) :: real \ _\, \<^term>\abs :: real \ _\, \<^term>\min :: real \ _\, \<^term>\max :: real \ _\, \<^term>\0::real\, \<^term>\1::real\, \<^term>\numeral :: num \ nat\, \<^term>\numeral :: num \ real\, \<^term>\Num.Bit0\, \<^term>\Num.Bit1\, \<^term>\Num.One\]; fun check_sos kcts ct = let val t = Thm.term_of ct val _ = if not (null (Term.add_tfrees t []) andalso null (Term.add_tvars t [])) then error "SOS: not sos. Additional type varables" else () val fs = Term.add_frees t [] val _ = if exists (fn ((_,T)) => not (T = \<^typ>\real\)) fs then error "SOS: not sos. Variables with type not real" else () val vs = Term.add_vars t [] val _ = if exists (fn ((_,T)) => not (T = \<^typ>\real\)) vs then error "SOS: not sos. Variables with type not real" else () val ukcs = subtract (fn (t,p) => Const p aconv t) kcts (Term.add_consts t []) val _ = if null ukcs then () else error ("SOSO: Unknown constants in Subgoal:" ^ commas (map fst ukcs)) in () end fun core_sos_tac print_cert prover = SUBPROOF (fn {concl, context = ctxt, ...} => let val _ = check_sos known_sos_constants concl val (th, certificates) = real_sos prover ctxt (Thm.dest_arg concl) val _ = print_cert certificates in resolve_tac ctxt [th] 1 end); fun default_SOME _ NONE v = SOME v | default_SOME _ (SOME v) _ = SOME v; fun lift_SOME f NONE a = f a | lift_SOME _ (SOME a) _ = SOME a; local val is_numeral = can (HOLogic.dest_number o Thm.term_of) in fun get_denom b ct = (case Thm.term_of ct of \<^term>\(/) :: real \ _\ $ _ $ _ => if is_numeral (Thm.dest_arg ct) then get_denom b (Thm.dest_arg1 ct) else default_SOME (get_denom b) (get_denom b (Thm.dest_arg ct)) (Thm.dest_arg ct, b) | \<^term>\(<) :: real \ _\ $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct) | \<^term>\(\) :: real \ _\ $ _ $ _ => lift_SOME (get_denom true) (get_denom true (Thm.dest_arg ct)) (Thm.dest_arg1 ct) | _ $ _ => lift_SOME (get_denom b) (get_denom b (Thm.dest_fun ct)) (Thm.dest_arg ct) | _ => NONE) end; val sos_ss = @{context} |> fold Simplifier.add_simp @{thms field_simps} |> Simplifier.del_simp @{thm mult_nonneg_nonneg} |> Simplifier.simpset_of; fun elim_one_denom_tac ctxt = CSUBGOAL (fn (P, i) => (case get_denom false P of NONE => no_tac | SOME (d, ord) => let val simp_ctxt = ctxt |> Simplifier.put_simpset sos_ss val th = Thm.instantiate' [] [SOME d, SOME (Thm.dest_arg P)] (if ord then @{lemma "(d=0 \ P) \ (d>0 \ P) \ (d<(0::real) \ P) \ P" by auto} else @{lemma "(d=0 \ P) \ (d \ (0::real) \ P) \ P" by blast}) in resolve_tac ctxt [th] i THEN Simplifier.asm_full_simp_tac simp_ctxt i end)); fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i); fun sos_tac print_cert prover ctxt = Object_Logic.full_atomize_tac ctxt THEN' elim_denom_tac ctxt THEN' core_sos_tac print_cert prover ctxt; end; diff --git a/src/HOL/Tools/Argo/argo_tactic.ML b/src/HOL/Tools/Argo/argo_tactic.ML --- a/src/HOL/Tools/Argo/argo_tactic.ML +++ b/src/HOL/Tools/Argo/argo_tactic.ML @@ -1,732 +1,732 @@ (* Title: HOL/Tools/Argo/argo_tactic.ML Author: Sascha Boehme HOL method and tactic for the Argo solver. *) signature ARGO_TACTIC = sig val trace: string Config.T val timeout: real Config.T (* extending the tactic *) type trans_context = Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option type extension = { trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans', trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans', term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option, replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv, replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option} val add_extension: extension -> theory -> theory (* proof utilities *) val discharges: thm -> thm list -> thm list val flatten_conv: conv -> thm -> conv (* interface to the tactic as well as the underlying checker and prover *) datatype result = Satisfiable of term -> bool option | Unsatisfiable val check: term list -> Proof.context -> result * Proof.context val prove: thm list -> Proof.context -> thm option * Proof.context val argo_tac: Proof.context -> thm list -> int -> tactic end structure Argo_Tactic: ARGO_TACTIC = struct (* readable fresh names for terms *) fun fresh_name n = Name.variant (case Long_Name.base_name n of "" => "x" | n' => n') fun fresh_type_name (Type (n, _)) = fresh_name n | fresh_type_name (TFree (n, _)) = fresh_name n | fresh_type_name (TVar ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i) fun fresh_term_name (Const (n, _)) = fresh_name n | fresh_term_name (Free (n, _)) = fresh_name n | fresh_term_name (Var ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i) | fresh_term_name _ = fresh_name "" (* tracing *) datatype mode = None | Basic | Full fun string_of_mode None = "none" | string_of_mode Basic = "basic" | string_of_mode Full = "full" fun requires_mode None = [] | requires_mode Basic = [Basic, Full] | requires_mode Full = [Full] val trace = Attrib.setup_config_string \<^binding>\argo_trace\ (K (string_of_mode None)) fun allows_mode ctxt = exists (equal (Config.get ctxt trace) o string_of_mode) o requires_mode fun output mode ctxt msg = if allows_mode ctxt mode then Output.tracing ("Argo: " ^ msg) else () val tracing = output Basic val full_tracing = output Full fun with_mode mode ctxt f = if allows_mode ctxt mode then f ctxt else () val with_tracing = with_mode Basic val with_full_tracing = with_mode Full (* timeout *) val timeout = Attrib.setup_config_real \<^binding>\argo_timeout\ (K 10.0) val timeout_seconds = seconds o Config.apply timeout fun with_timeout ctxt f x = Timeout.apply (timeout_seconds ctxt) f x (* extending the tactic *) type trans_context = Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option type extension = { trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans', trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans', term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option, replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv, replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option} fun eq_extension ((serial1, _), (serial2, _)) = (serial1 = serial2) structure Extensions = Theory_Data ( type T = (serial * extension) list val empty = [] val extend = I val merge = merge eq_extension ) fun add_extension ext = Extensions.map (insert eq_extension (serial (), ext)) fun get_extensions ctxt = Extensions.get (Proof_Context.theory_of ctxt) fun apply_first ctxt f = get_first (fn (_, e) => f e) (get_extensions ctxt) fun ext_trans sel ctxt f x tcx = apply_first ctxt (fn ext => sel ext f x tcx) val ext_trans_type = ext_trans (fn {trans_type, ...}: extension => trans_type) val ext_trans_term = ext_trans (fn {trans_term, ...}: extension => trans_term) fun ext_term_of ctxt f e = apply_first ctxt (fn {term_of, ...}: extension => term_of f e) fun ext_replay_rewr ctxt r = get_extensions ctxt |> map (fn (_, {replay_rewr, ...}: extension) => replay_rewr ctxt r) |> Conv.first_conv fun ext_replay cprop_of ctxt rule prems = (case apply_first ctxt (fn {replay, ...}: extension => replay cprop_of ctxt rule prems) of SOME thm => thm | NONE => raise THM ("failed to replay " ^ quote (Argo_Proof.string_of_rule rule), 0, [])) (* translating input terms *) fun add_new_type T (names, types, terms) = let val (n, names') = fresh_type_name T names val ty = Argo_Expr.Type n in (ty, (names', Typtab.update (T, ty) types, terms)) end fun add_type T (tcx as (_, types, _)) = (case Typtab.lookup types T of SOME ty => (ty, tcx) | NONE => add_new_type T tcx) fun trans_type _ \<^typ>\HOL.bool\ = pair Argo_Expr.Bool | trans_type ctxt (Type (\<^type_name>\fun\, [T1, T2])) = trans_type ctxt T1 ##>> trans_type ctxt T2 #>> Argo_Expr.Func | trans_type ctxt T = (fn tcx => (case ext_trans_type ctxt (trans_type ctxt) T tcx of SOME result => result | NONE => add_type T tcx)) fun add_new_term ctxt t T tcx = let val (ty, (names, types, terms)) = trans_type ctxt T tcx val (n, names') = fresh_term_name t names val c = (n, ty) in (Argo_Expr.mk_con c, (names', types, Termtab.update (t, c) terms)) end fun add_term ctxt t (tcx as (_, _, terms)) = (case Termtab.lookup terms t of SOME c => (Argo_Expr.mk_con c, tcx) | NONE => add_new_term ctxt t (Term.fastype_of t) tcx) fun mk_eq \<^typ>\HOL.bool\ = Argo_Expr.mk_iff | mk_eq _ = Argo_Expr.mk_eq fun trans_term _ \<^const>\HOL.True\ = pair Argo_Expr.true_expr | trans_term _ \<^const>\HOL.False\ = pair Argo_Expr.false_expr | trans_term ctxt (\<^const>\HOL.Not\ $ t) = trans_term ctxt t #>> Argo_Expr.mk_not | trans_term ctxt (\<^const>\HOL.conj\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_and2 | trans_term ctxt (\<^const>\HOL.disj\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_or2 | trans_term ctxt (\<^const>\HOL.implies\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_imp | trans_term ctxt (Const (\<^const_name>\HOL.If\, _) $ t1 $ t2 $ t3) = trans_term ctxt t1 ##>> trans_term ctxt t2 ##>> trans_term ctxt t3 #>> (fn ((u1, u2), u3) => Argo_Expr.mk_ite u1 u2 u3) | trans_term ctxt (Const (\<^const_name>\HOL.eq\, T) $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry (mk_eq (Term.domain_type T)) | trans_term ctxt (t as (t1 $ t2)) = (fn tcx => (case ext_trans_term ctxt (trans_term ctxt) t tcx of SOME result => result | NONE => tcx |> trans_term ctxt t1 ||>> trans_term ctxt t2 |>> uncurry Argo_Expr.mk_app)) | trans_term ctxt t = (fn tcx => (case ext_trans_term ctxt (trans_term ctxt) t tcx of SOME result => result | NONE => add_term ctxt t tcx)) fun translate ctxt prop = trans_term ctxt (HOLogic.dest_Trueprop prop) (* invoking the solver *) type data = { names: Name.context, types: Argo_Expr.typ Typtab.table, terms: (string * Argo_Expr.typ) Termtab.table, argo: Argo_Solver.context} fun mk_data names types terms argo: data = {names=names, types=types, terms=terms, argo=argo} val empty_data = mk_data Name.context Typtab.empty Termtab.empty Argo_Solver.context structure Solver_Data = Proof_Data ( type T = data option fun init _ = SOME empty_data ) datatype ('m, 'p) solver_result = Model of 'm | Proof of 'p fun raw_solve es argo = Model (Argo_Solver.assert es argo) handle Argo_Proof.UNSAT proof => Proof proof fun value_of terms model t = (case Termtab.lookup terms t of SOME c => model c | _ => NONE) fun trace_props props ctxt = tracing ctxt (Pretty.string_of (Pretty.big_list "using these propositions:" (map (Syntax.pretty_term ctxt) props))) fun trace_result ctxt ({elapsed, ...}: Timing.timing) msg = tracing ctxt ("found a " ^ msg ^ " in " ^ string_of_int (Time.toMilliseconds elapsed) ^ " ms") fun solve props ctxt = (case Solver_Data.get ctxt of NONE => error "bad Argo solver context" | SOME {names, types, terms, argo} => let val _ = with_tracing ctxt (trace_props props) val (es, (names', types', terms')) = fold_map (translate ctxt) props (names, types, terms) val _ = tracing ctxt "starting the prover" in (case Timing.timing (raw_solve es) argo of (time, Proof proof) => let val _ = trace_result ctxt time "proof" in (Proof (terms', proof), Solver_Data.put NONE ctxt) end | (time, Model argo') => let val _ = trace_result ctxt time "model" val model = value_of terms' (Argo_Solver.model_of argo') in (Model model, Solver_Data.put (SOME (mk_data names' types' terms' argo')) ctxt) end) end) (* reverse translation *) structure Contab = Table(type key = string * Argo_Expr.typ val ord = Argo_Expr.con_ord) fun mk_nary f ts = uncurry (fold_rev (curry f)) (split_last ts) fun mk_nary' d _ [] = d | mk_nary' _ f ts = mk_nary f ts fun mk_ite t1 t2 t3 = let val T = Term.fastype_of t2 val ite = Const (\<^const_name>\HOL.If\, [\<^typ>\HOL.bool\, T, T] ---> T) in ite $ t1 $ t2 $ t3 end fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = \<^const>\HOL.True\ | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = \<^const>\HOL.False\ | term_of cx (Argo_Expr.E (Argo_Expr.Not, [e])) = HOLogic.mk_not (term_of cx e) | term_of cx (Argo_Expr.E (Argo_Expr.And, es)) = mk_nary' \<^const>\HOL.True\ HOLogic.mk_conj (map (term_of cx) es) | term_of cx (Argo_Expr.E (Argo_Expr.Or, es)) = mk_nary' \<^const>\HOL.False\ HOLogic.mk_disj (map (term_of cx) es) | term_of cx (Argo_Expr.E (Argo_Expr.Imp, [e1, e2])) = HOLogic.mk_imp (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.Iff, [e1, e2])) = HOLogic.mk_eq (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.Ite, [e1, e2, e3])) = mk_ite (term_of cx e1) (term_of cx e2) (term_of cx e3) | term_of cx (Argo_Expr.E (Argo_Expr.Eq, [e1, e2])) = HOLogic.mk_eq (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.App, [e1, e2])) = term_of cx e1 $ term_of cx e2 | term_of (_, cons) (Argo_Expr.E (Argo_Expr.Con (c as (n, _)), _)) = (case Contab.lookup cons c of SOME t => t | NONE => error ("Unexpected expression named " ^ quote n)) | term_of (cx as (ctxt, _)) e = (case ext_term_of ctxt (term_of cx) e of SOME t => t | NONE => raise Fail "bad expression") fun as_prop ct = Thm.apply \<^cterm>\HOL.Trueprop\ ct fun cterm_of ctxt cons e = Thm.cterm_of ctxt (term_of (ctxt, cons) e) fun cprop_of ctxt cons e = as_prop (cterm_of ctxt cons e) (* generic proof tools *) fun discharge thm rule = thm INCR_COMP rule fun discharge2 thm1 thm2 rule = discharge thm2 (discharge thm1 rule) fun discharges thm rules = [thm] RL rules fun under_assumption f ct = let val cprop = as_prop ct in Thm.implies_intr cprop (f (Thm.assume cprop)) end fun instantiate cv ct = Thm.instantiate ([], [(Term.dest_Var (Thm.term_of cv), ct)]) (* proof replay for tautologies *) fun prove_taut ctxt ns t = Goal.prove ctxt ns [] (HOLogic.mk_Trueprop t) (fn {context, ...} => HEADGOAL (Classical.fast_tac context)) fun with_frees ctxt n mk = let val ns = map (fn i => "P" ^ string_of_int i) (0 upto (n - 1)) val ts = map (Free o rpair \<^typ>\bool\) ns val t = mk_nary HOLogic.mk_disj (mk ts) in prove_taut ctxt ns t end fun taut_and1_term ts = mk_nary HOLogic.mk_conj ts :: map HOLogic.mk_not ts fun taut_and2_term i ts = [HOLogic.mk_not (mk_nary HOLogic.mk_conj ts), nth ts i] fun taut_or1_term i ts = [mk_nary HOLogic.mk_disj ts, HOLogic.mk_not (nth ts i)] fun taut_or2_term ts = HOLogic.mk_not (mk_nary HOLogic.mk_disj ts) :: ts val iff_1_taut = @{lemma "P = Q \ P \ Q" by fast} val iff_2_taut = @{lemma "P = Q \ (\P) \ (\Q)" by fast} val iff_3_taut = @{lemma "\(P = Q) \ (\P) \ Q" by fast} val iff_4_taut = @{lemma "\(P = Q) \ P \ (\Q)" by fast} val ite_then_taut = @{lemma "\P \ (if P then t else u) = t" by auto} val ite_else_taut = @{lemma "P \ (if P then t else u) = u" by auto} fun taut_rule_of ctxt (Argo_Proof.Taut_And_1 n) = with_frees ctxt n taut_and1_term | taut_rule_of ctxt (Argo_Proof.Taut_And_2 (i, n)) = with_frees ctxt n (taut_and2_term i) | taut_rule_of ctxt (Argo_Proof.Taut_Or_1 (i, n)) = with_frees ctxt n (taut_or1_term i) | taut_rule_of ctxt (Argo_Proof.Taut_Or_2 n) = with_frees ctxt n taut_or2_term | taut_rule_of _ Argo_Proof.Taut_Iff_1 = iff_1_taut | taut_rule_of _ Argo_Proof.Taut_Iff_2 = iff_2_taut | taut_rule_of _ Argo_Proof.Taut_Iff_3 = iff_3_taut | taut_rule_of _ Argo_Proof.Taut_Iff_4 = iff_4_taut | taut_rule_of _ Argo_Proof.Taut_Ite_Then = ite_then_taut | taut_rule_of _ Argo_Proof.Taut_Ite_Else = ite_else_taut fun replay_taut ctxt k ct = let val rule = taut_rule_of ctxt k in Thm.instantiate (Thm.match (Thm.cprop_of rule, ct)) rule end (* proof replay for conjunct extraction *) fun replay_conjunct 0 1 thm = thm | replay_conjunct 0 _ thm = discharge thm @{thm conjunct1} | replay_conjunct 1 2 thm = discharge thm @{thm conjunct2} | replay_conjunct i n thm = replay_conjunct (i - 1) (n - 1) (discharge thm @{thm conjunct2}) (* proof replay for rewrite steps *) fun mk_rewr thm = thm RS @{thm eq_reflection} fun not_nary_conv rule i ct = if i > 1 then (Conv.rewr_conv rule then_conv Conv.arg_conv (not_nary_conv rule (i - 1))) ct else Conv.all_conv ct val flatten_and_thm = @{lemma "(P1 \ P2) \ P3 \ P1 \ P2 \ P3" by simp} val flatten_or_thm = @{lemma "(P1 \ P2) \ P3 \ P1 \ P2 \ P3" by simp} fun flatten_conv cv rule ct = ( Conv.try_conv (Conv.arg_conv cv) then_conv Conv.try_conv (Conv.rewr_conv rule then_conv cv)) ct fun flat_conj_conv ct = (case Thm.term_of ct of \<^const>\HOL.conj\ $ _ $ _ => flatten_conv flat_conj_conv flatten_and_thm ct | _ => Conv.all_conv ct) fun flat_disj_conv ct = (case Thm.term_of ct of \<^const>\HOL.disj\ $ _ $ _ => flatten_conv flat_disj_conv flatten_or_thm ct | _ => Conv.all_conv ct) fun explode rule1 rule2 thm = explode rule1 rule2 (thm RS rule1) @ explode rule1 rule2 (thm RS rule2) handle THM _ => [thm] val explode_conj = explode @{thm conjunct1} @{thm conjunct2} val explode_ndis = explode @{lemma "\(P \ Q) \ \P" by auto} @{lemma "\(P \ Q) \ \Q" by auto} fun pick_false i thms = nth thms i fun pick_dual rule (i1, i2) thms = rule OF [nth thms i1, nth thms i2] handle THM _ => rule OF [nth thms i2, nth thms i1] val pick_dual_conj = pick_dual @{lemma "\P \ P \ False" by auto} val pick_dual_ndis = pick_dual @{lemma "\P \ P \ \True" by auto} fun join thm0 rule is thms = let val l = length thms val thms' = fold (fn i => cons (if 0 <= i andalso i < l then nth thms i else thm0)) is [] in fold (fn thm => fn thm' => discharge2 thm thm' rule) (tl thms') (hd thms') end val join_conj = join @{lemma "True" by auto} @{lemma "P \ Q \ P \ Q" by auto} val join_ndis = join @{lemma "\False" by auto} @{lemma "\P \ \Q \ \(P \ Q)" by auto} val false_thm = @{lemma "False \ P" by auto} val ntrue_thm = @{lemma "\True \ P" by auto} val iff_conj_thm = @{lemma "(P \ Q) \ (Q \ P) \ P = Q" by auto} val iff_ndis_thm = @{lemma "(\P \ \Q) \ (\Q \ \P) \ P = Q" by auto} fun iff_intro rule lf rf ct = let val lhs = under_assumption lf ct val rhs = rf (Thm.dest_arg (snd (Thm.dest_implies (Thm.cprop_of lhs)))) in mk_rewr (discharge2 lhs rhs rule) end fun with_conj f g ct = iff_intro iff_conj_thm (f o explode_conj) g ct fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply \<^cterm>\HOL.Not\ ct) fun swap_indices n iss = map (fn i => find_index (fn is => member (op =) is i) iss) (0 upto (n - 1)) fun sort_nary w f g (n, iss) = w (f (map hd iss)) (under_assumption (f (swap_indices n iss) o g)) val sort_conj = sort_nary with_conj join_conj explode_conj val sort_ndis = sort_nary with_ndis join_ndis explode_ndis val not_true_thm = mk_rewr @{lemma "(\True) = False" by auto} val not_false_thm = mk_rewr @{lemma "(\False) = True" by auto} val not_not_thm = mk_rewr @{lemma "(\\P) = P" by auto} val not_and_thm = mk_rewr @{lemma "(\(P \ Q)) = (\P \ \Q)" by auto} val not_or_thm = mk_rewr @{lemma "(\(P \ Q)) = (\P \ \Q)" by auto} val not_iff_thms = map mk_rewr @{lemma "(\((\P) = Q)) = (P = Q)" "(\(P = (\Q))) = (P = Q)" "(\(P = Q)) = ((\P) = Q)" by auto} val iff_true_thms = map mk_rewr @{lemma "(True = P) = P" "(P = True) = P" by auto} val iff_false_thms = map mk_rewr @{lemma "(False = P) = (\P)" "(P = False) = (\P)" by auto} val iff_not_not_thm = mk_rewr @{lemma "((\P) = (\Q)) = (P = Q)" by auto} val iff_refl_thm = mk_rewr @{lemma "(P = P) = True" by auto} val iff_symm_thm = mk_rewr @{lemma "(P = Q) = (Q = P)" by auto} val iff_dual_thms = map mk_rewr @{lemma "(P = (\P)) = False" "((\P) = P) = False" by auto} val imp_thm = mk_rewr @{lemma "(P \ Q) = (\P \ Q)" by auto} val ite_prop_thm = mk_rewr @{lemma "(If P Q R) = ((\P \ Q) \ (P \ R) \ (Q \ R))" by auto} val ite_true_thm = mk_rewr @{lemma "(If True t u) = t" by auto} val ite_false_thm = mk_rewr @{lemma "(If False t u) = u" by auto} val ite_eq_thm = mk_rewr @{lemma "(If P t t) = t" by auto} val eq_refl_thm = mk_rewr @{lemma "(t = t) = True" by auto} val eq_symm_thm = mk_rewr @{lemma "(t1 = t2) = (t2 = t1)" by auto} fun replay_rewr _ Argo_Proof.Rewr_Not_True = Conv.rewr_conv not_true_thm | replay_rewr _ Argo_Proof.Rewr_Not_False = Conv.rewr_conv not_false_thm | replay_rewr _ Argo_Proof.Rewr_Not_Not = Conv.rewr_conv not_not_thm | replay_rewr _ (Argo_Proof.Rewr_Not_And i) = not_nary_conv not_and_thm i | replay_rewr _ (Argo_Proof.Rewr_Not_Or i) = not_nary_conv not_or_thm i | replay_rewr _ Argo_Proof.Rewr_Not_Iff = Conv.rewrs_conv not_iff_thms | replay_rewr _ (Argo_Proof.Rewr_And_False i) = with_conj (pick_false i) (K false_thm) | replay_rewr _ (Argo_Proof.Rewr_And_Dual ip) = with_conj (pick_dual_conj ip) (K false_thm) | replay_rewr _ (Argo_Proof.Rewr_And_Sort is) = flat_conj_conv then_conv sort_conj is | replay_rewr _ (Argo_Proof.Rewr_Or_True i) = with_ndis (pick_false i) (K ntrue_thm) | replay_rewr _ (Argo_Proof.Rewr_Or_Dual ip) = with_ndis (pick_dual_ndis ip) (K ntrue_thm) | replay_rewr _ (Argo_Proof.Rewr_Or_Sort is) = flat_disj_conv then_conv sort_ndis is | replay_rewr _ Argo_Proof.Rewr_Iff_True = Conv.rewrs_conv iff_true_thms | replay_rewr _ Argo_Proof.Rewr_Iff_False = Conv.rewrs_conv iff_false_thms | replay_rewr _ Argo_Proof.Rewr_Iff_Not_Not = Conv.rewr_conv iff_not_not_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Refl = Conv.rewr_conv iff_refl_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Symm = Conv.rewr_conv iff_symm_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Dual = Conv.rewrs_conv iff_dual_thms | replay_rewr _ Argo_Proof.Rewr_Imp = Conv.rewr_conv imp_thm | replay_rewr _ Argo_Proof.Rewr_Ite_Prop = Conv.rewr_conv ite_prop_thm | replay_rewr _ Argo_Proof.Rewr_Ite_True = Conv.rewr_conv ite_true_thm | replay_rewr _ Argo_Proof.Rewr_Ite_False = Conv.rewr_conv ite_false_thm | replay_rewr _ Argo_Proof.Rewr_Ite_Eq = Conv.rewr_conv ite_eq_thm | replay_rewr _ Argo_Proof.Rewr_Eq_Refl = Conv.rewr_conv eq_refl_thm | replay_rewr _ Argo_Proof.Rewr_Eq_Symm = Conv.rewr_conv eq_symm_thm | replay_rewr ctxt r = ext_replay_rewr ctxt r fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 fun replay_conv _ Argo_Proof.Keep_Conv ct = Conv.all_conv ct | replay_conv ctxt (Argo_Proof.Then_Conv (c1, c2)) ct = (replay_conv ctxt c1 then_conv replay_conv ctxt c2) ct | replay_conv ctxt (Argo_Proof.Args_Conv (Argo_Expr.App, [c1, c2])) ct = Conv.combination_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct | replay_conv ctxt (Argo_Proof.Args_Conv (_, cs)) ct = replay_args_conv ctxt cs ct | replay_conv ctxt (Argo_Proof.Rewr_Conv r) ct = replay_rewr ctxt r ct and replay_args_conv _ [] ct = Conv.all_conv ct | replay_args_conv ctxt [c] ct = Conv.arg_conv (replay_conv ctxt c) ct | replay_args_conv ctxt [c1, c2] ct = binop_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct | replay_args_conv ctxt (c :: cs) ct = (case Term.head_of (Thm.term_of ct) of Const (\<^const_name>\HOL.If\, _) => let val (cs', c') = split_last cs in Conv.combination_conv (replay_args_conv ctxt (c :: cs')) (replay_conv ctxt c') ct end | _ => binop_conv (replay_conv ctxt c) (replay_args_conv ctxt cs) ct) fun replay_rewrite ctxt c thm = Conv.fconv_rule (HOLogic.Trueprop_conv (replay_conv ctxt c)) thm (* proof replay for clauses *) val prep_clause_rule = @{lemma "P \ \P \ False" by fast} val extract_lit_rule = @{lemma "(\(P \ Q) \ False) \ \P \ \Q \ False" by fast} fun add_lit i thm lits = let val ct = Thm.cprem_of thm 1 in (Thm.implies_elim thm (Thm.assume ct), (i, ct) :: lits) end fun extract_lits [] _ = error "Bad clause" | extract_lits [i] (thm, lits) = add_lit i thm lits | extract_lits (i :: is) (thm, lits) = extract_lits is (add_lit i (discharge thm extract_lit_rule) lits) fun lit_ord ((l1, _), (l2, _)) = int_ord (abs l1, abs l2) fun replay_with_lits [] thm lits = (thm, lits) | replay_with_lits is thm lits = extract_lits is (discharge thm prep_clause_rule, lits) ||> Ord_List.make lit_ord fun replay_clause is thm = replay_with_lits is thm [] (* proof replay for unit resolution *) val unit_rule = @{lemma "(P \ False) \ (\P \ False) \ False" by fast} val unit_rule_var = Thm.dest_arg (Thm.dest_arg1 (Thm.cprem_of unit_rule 1)) val bogus_ct = \<^cterm>\HOL.True\ fun replay_unit_res lit (pthm, plits) (nthm, nlits) = let val plit = the (AList.lookup (op =) plits lit) val nlit = the (AList.lookup (op =) nlits (~lit)) val prune = Ord_List.remove lit_ord (lit, bogus_ct) in unit_rule |> instantiate unit_rule_var (Thm.dest_arg plit) |> Thm.elim_implies (Thm.implies_intr plit pthm) |> Thm.elim_implies (Thm.implies_intr nlit nthm) |> rpair (Ord_List.union lit_ord (prune nlits) (prune plits)) end (* proof replay for hypothesis *) val dneg_rule = @{lemma "\\P \ P" by auto} fun replay_hyp i ct = if i < 0 then (Thm.assume ct, [(~i, ct)]) else let val cu = as_prop (Thm.apply \<^cterm>\HOL.Not\ (Thm.apply \<^cterm>\HOL.Not\ (Thm.dest_arg ct))) in (discharge (Thm.assume cu) dneg_rule, [(~i, cu)]) end (* proof replay for lemma *) fun replay_lemma is (thm, lits) = replay_with_lits is thm lits (* proof replay for reflexivity *) val refl_rule = @{thm refl} val refl_rule_var = Thm.dest_arg1 (Thm.dest_arg (Thm.cprop_of refl_rule)) fun replay_refl ct = Thm.instantiate (Thm.match (refl_rule_var, ct)) refl_rule (* proof replay for symmetry *) val symm_rules = @{lemma "a = b ==> b = a" "\(a = b) \ \(b = a)" by simp_all} fun replay_symm thm = hd (discharges thm symm_rules) (* proof replay for transitivity *) val trans_rules = @{lemma "\(a = b) \ b = c \ \(a = c)" "a = b \ \(b = c) \ \(a = c)" "a = b \ b = c ==> a = c" by simp_all} fun replay_trans thm1 thm2 = hd (discharges thm2 (discharges thm1 trans_rules)) (* proof replay for congruence *) fun replay_cong thm1 thm2 = discharge thm2 (discharge thm1 @{thm cong}) (* proof replay for substitution *) val subst_rule1 = @{lemma "\(p a) \ p = q \ a = b \ \(q b)" by simp} val subst_rule2 = @{lemma "p a \ p = q \ a = b \ q b" by simp} fun replay_subst thm1 thm2 thm3 = subst_rule1 OF [thm1, thm2, thm3] handle THM _ => subst_rule2 OF [thm1, thm2, thm3] (* proof replay *) structure Thm_Cache = Table(type key = Argo_Proof.proof_id val ord = Argo_Proof.proof_id_ord) val unclausify_rule1 = @{lemma "(\P \ False) \ P" by auto} val unclausify_rule2 = @{lemma "(P \ False) \ \P" by auto} fun unclausify (thm, lits) ls = (case (Thm.prop_of thm, lits) of (\<^const>\HOL.Trueprop\ $ \<^const>\HOL.False\, [(_, ct)]) => let val thm = Thm.implies_intr ct thm in (discharge thm unclausify_rule1 handle THM _ => discharge thm unclausify_rule2, ls) end | _ => (thm, Ord_List.union lit_ord lits ls)) fun with_thms f tps = fold_map unclausify tps [] |>> f fun bad_premises () = raise Fail "bad number of premises" fun with_thms1 f = with_thms (fn [thm] => f thm | _ => bad_premises ()) fun with_thms2 f = with_thms (fn [thm1, thm2] => f thm1 thm2 | _ => bad_premises ()) fun with_thms3 f = with_thms (fn [thm1, thm2, thm3] => f thm1 thm2 thm3 | _ => bad_premises ()) fun replay_rule (ctxt, cons, facts) prems rule = (case rule of Argo_Proof.Axiom i => (nth facts i, []) | Argo_Proof.Taut (k, concl) => (replay_taut ctxt k (cprop_of ctxt cons concl), []) | Argo_Proof.Conjunct (i, n) => with_thms1 (replay_conjunct i n) prems | Argo_Proof.Rewrite c => with_thms1 (replay_rewrite ctxt c) prems | Argo_Proof.Clause is => replay_clause is (fst (hd prems)) | Argo_Proof.Unit_Res i => replay_unit_res i (hd prems) (hd (tl prems)) | Argo_Proof.Hyp (i, concl) => replay_hyp i (cprop_of ctxt cons concl) | Argo_Proof.Lemma is => replay_lemma is (hd prems) | Argo_Proof.Refl concl => (replay_refl (cterm_of ctxt cons concl), []) | Argo_Proof.Symm => with_thms1 replay_symm prems | Argo_Proof.Trans => with_thms2 replay_trans prems | Argo_Proof.Cong => with_thms2 replay_cong prems | Argo_Proof.Subst => with_thms3 replay_subst prems | _ => with_thms (ext_replay (cprop_of ctxt cons) ctxt rule) prems) fun with_cache f proof thm_cache = (case Thm_Cache.lookup thm_cache (Argo_Proof.id_of proof) of SOME thm => (thm, thm_cache) | NONE => let val (thm, thm_cache') = f proof thm_cache in (thm, Thm_Cache.update (Argo_Proof.id_of proof, thm) thm_cache') end) fun trace_step ctxt proof_id rule proofs = with_full_tracing ctxt (fn ctxt' => let val id = Argo_Proof.string_of_proof_id proof_id val ids = map (Argo_Proof.string_of_proof_id o Argo_Proof.id_of) proofs val rule_string = Argo_Proof.string_of_rule rule in full_tracing ctxt' (" " ^ id ^ " <- " ^ space_implode " " ids ^ " . " ^ rule_string) end) fun replay_bottom_up (env as (ctxt, _, _)) proof thm_cache = let val (proof_id, rule, proofs) = Argo_Proof.dest proof val (prems, thm_cache) = fold_map (with_cache (replay_bottom_up env)) proofs thm_cache val _ = trace_step ctxt proof_id rule proofs in (replay_rule env prems rule, thm_cache) end fun replay_proof env proof = with_cache (replay_bottom_up env) proof Thm_Cache.empty fun replay ctxt terms facts proof = let val env = (ctxt, Termtab.fold (Contab.update o swap) terms Contab.empty, facts) val _ = tracing ctxt "replaying the proof" val ({elapsed=t, ...}, ((thm, _), _)) = Timing.timing (replay_proof env) proof val _ = tracing ctxt ("replayed the proof in " ^ string_of_int (Time.toMilliseconds t) ^ " ms") in thm end (* normalizing goals *) fun instantiate_elim_rule thm = let val ct = Drule.strip_imp_concl (Thm.cprop_of thm) in (case Thm.term_of ct of \<^const>\HOL.Trueprop\ $ Var (_, \<^typ>\HOL.bool\) => instantiate (Thm.dest_arg ct) \<^cterm>\HOL.False\ thm | Var _ => instantiate ct \<^cprop>\HOL.False\ thm | _ => thm) end fun atomize_conv ctxt ct = (case Thm.term_of ct of \<^const>\HOL.Trueprop\ $ _ => Conv.all_conv | \<^const>\Pure.imp\ $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp} | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq} | Const (\<^const_name>\Pure.all\, _) $ Abs _ => Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all} | _ => Conv.all_conv) ct fun normalize ctxt thm = thm |> instantiate_elim_rule |> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion) - |> Drule.forall_intr_vars + |> Thm.forall_intr_vars |> Conv.fconv_rule (atomize_conv ctxt) (* prover, tactic and method *) datatype result = Satisfiable of term -> bool option | Unsatisfiable fun check props ctxt = (case solve props ctxt of (Proof _, ctxt') => (Unsatisfiable, ctxt') | (Model model, ctxt') => (Satisfiable model, ctxt')) fun prove thms ctxt = let val thms' = map (normalize ctxt) thms in (case solve (map Thm.prop_of thms') ctxt of (Model _, ctxt') => (NONE, ctxt') | (Proof (terms, proof), ctxt') => (SOME (replay ctxt' terms thms' proof), ctxt')) end fun argo_tac ctxt thms = CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 (Conv.try_conv (Conv.rewr_conv @{thm atomize_eq})))) ctxt) THEN' Tactic.resolve_tac ctxt [@{thm ccontr}] THEN' Subgoal.FOCUS (fn {context, prems, ...} => (case with_timeout context (prove (thms @ prems)) context of (SOME thm, _) => Tactic.resolve_tac context [thm] 1 | (NONE, _) => Tactical.no_tac)) ctxt val _ = Theory.setup (Method.setup \<^binding>\argo\ (Scan.optional Attrib.thms [] >> (fn thms => fn ctxt => METHOD (fn facts => HEADGOAL (argo_tac ctxt (thms @ facts))))) "Applies the Argo prover") end diff --git a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML @@ -1,1273 +1,1273 @@ (* Title: HOL/Tools/Ctr_Sugar/ctr_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013 Wrapping existing freely generated type's constructors. *) signature CTR_SUGAR = sig datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar val transfer_ctr_sugar: theory -> ctr_sugar -> ctr_sugar val ctr_sugar_of: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_global: theory -> string -> ctr_sugar option val ctr_sugars_of: Proof.context -> ctr_sugar list val ctr_sugars_of_global: theory -> ctr_sugar list val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_case_global: theory -> string -> ctr_sugar option val ctr_sugar_interpretation: string -> (ctr_sugar -> local_theory -> local_theory) -> theory -> theory val interpret_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val register_ctr_sugar_raw: ctr_sugar -> local_theory -> local_theory val register_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val default_register_ctr_sugar_global: (string -> bool) -> ctr_sugar -> theory -> theory val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list val mk_ctr: typ list -> term -> term val mk_case: typ list -> typ -> term -> term val mk_disc_or_sel: typ list -> term -> term val name_of_ctr: term -> string val name_of_disc: term -> string val dest_ctr: Proof.context -> string -> term -> term * term list val dest_case: Proof.context -> string -> typ list -> term -> (ctr_sugar * term list * term list) option type ('c, 'a) ctr_spec = (binding * 'c) * 'a list val disc_of_ctr_spec: ('c, 'a) ctr_spec -> binding val ctr_of_ctr_spec: ('c, 'a) ctr_spec -> 'c val args_of_ctr_spec: ('c, 'a) ctr_spec -> 'a list val code_plugin: string type ctr_options = (string -> bool) * bool type ctr_options_cmd = (Proof.context -> string -> bool) * bool val fake_local_theory_for_sel_defaults: (binding * typ) list -> Proof.context -> Proof.context val free_constructors: ctr_sugar_kind -> ({prems: thm list, context: Proof.context} -> tactic) list list -> ((ctr_options * binding) * (term, binding) ctr_spec list) * term list -> local_theory -> ctr_sugar * local_theory val free_constructors_cmd: ctr_sugar_kind -> ((((Proof.context -> Plugin_Name.filter) * bool) * binding) * ((binding * string) * binding list) list) * string list -> Proof.context -> Proof.state val default_ctr_options: ctr_options val default_ctr_options_cmd: ctr_options_cmd val parse_bound_term: (binding * string) parser val parse_ctr_options: ctr_options_cmd parser val parse_ctr_spec: 'c parser -> 'a parser -> ('c, 'a) ctr_spec parser val parse_sel_default_eqs: string list parser end; structure Ctr_Sugar : CTR_SUGAR = struct open Ctr_Sugar_Util open Ctr_Sugar_Tactics open Ctr_Sugar_Code datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown; type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; fun morph_ctr_sugar phi ({kind, T, ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts, case_thms, case_cong, case_cong_weak, case_distribs, split, split_asm, disc_defs, disc_thmss, discIs, disc_eq_cases, sel_defs, sel_thmss, distinct_discsss, exhaust_discs, exhaust_sels, collapses, expands, split_sels, split_sel_asms, case_eq_ifs} : ctr_sugar) = {kind = kind, T = Morphism.typ phi T, ctrs = map (Morphism.term phi) ctrs, casex = Morphism.term phi casex, discs = map (Morphism.term phi) discs, selss = map (map (Morphism.term phi)) selss, exhaust = Morphism.thm phi exhaust, nchotomy = Morphism.thm phi nchotomy, injects = map (Morphism.thm phi) injects, distincts = map (Morphism.thm phi) distincts, case_thms = map (Morphism.thm phi) case_thms, case_cong = Morphism.thm phi case_cong, case_cong_weak = Morphism.thm phi case_cong_weak, case_distribs = map (Morphism.thm phi) case_distribs, split = Morphism.thm phi split, split_asm = Morphism.thm phi split_asm, disc_defs = map (Morphism.thm phi) disc_defs, disc_thmss = map (map (Morphism.thm phi)) disc_thmss, discIs = map (Morphism.thm phi) discIs, disc_eq_cases = map (Morphism.thm phi) disc_eq_cases, sel_defs = map (Morphism.thm phi) sel_defs, sel_thmss = map (map (Morphism.thm phi)) sel_thmss, distinct_discsss = map (map (map (Morphism.thm phi))) distinct_discsss, exhaust_discs = map (Morphism.thm phi) exhaust_discs, exhaust_sels = map (Morphism.thm phi) exhaust_sels, collapses = map (Morphism.thm phi) collapses, expands = map (Morphism.thm phi) expands, split_sels = map (Morphism.thm phi) split_sels, split_sel_asms = map (Morphism.thm phi) split_sel_asms, case_eq_ifs = map (Morphism.thm phi) case_eq_ifs}; val transfer_ctr_sugar = morph_ctr_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = (Position.T * ctr_sugar) Symtab.table; val empty = Symtab.empty; val extend = I; fun merge data : T = Symtab.merge (K true) data; ); fun ctr_sugar_of_generic context = Option.map (transfer_ctr_sugar (Context.theory_of context) o #2) o Symtab.lookup (Data.get context); fun ctr_sugars_of_generic context = Symtab.fold (cons o transfer_ctr_sugar (Context.theory_of context) o #2 o #2) (Data.get context) []; fun ctr_sugar_of_case_generic context s = find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of_generic context); val ctr_sugar_of = ctr_sugar_of_generic o Context.Proof; val ctr_sugar_of_global = ctr_sugar_of_generic o Context.Theory; val ctr_sugars_of = ctr_sugars_of_generic o Context.Proof; val ctr_sugars_of_global = ctr_sugars_of_generic o Context.Theory; val ctr_sugar_of_case = ctr_sugar_of_case_generic o Context.Proof; val ctr_sugar_of_case_global = ctr_sugar_of_case_generic o Context.Theory; structure Ctr_Sugar_Plugin = Plugin(type T = ctr_sugar); fun ctr_sugar_interpretation name f = Ctr_Sugar_Plugin.interpretation name (fn ctr_sugar => fn lthy => f (transfer_ctr_sugar (Proof_Context.theory_of lthy) ctr_sugar) lthy); val interpret_ctr_sugar = Ctr_Sugar_Plugin.data; fun register_ctr_sugar_raw (ctr_sugar as {T = Type (name, _), ...}) = Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => fn context => let val pos = Position.thread_data () in Data.map (Symtab.update (name, (pos, morph_ctr_sugar phi ctr_sugar))) context end); fun register_ctr_sugar plugins ctr_sugar = register_ctr_sugar_raw ctr_sugar #> interpret_ctr_sugar plugins ctr_sugar; fun default_register_ctr_sugar_global plugins (ctr_sugar as {T = Type (name, _), ...}) thy = let val tab = Data.get (Context.Theory thy); val pos = Position.thread_data (); in if Symtab.defined tab name then thy else thy |> Context.theory_map (Data.put (Symtab.update_new (name, (pos, ctr_sugar)) tab)) |> Named_Target.theory_map (Ctr_Sugar_Plugin.data plugins ctr_sugar) end; val is_prefix = "is_"; val un_prefix = "un_"; val not_prefix = "not_"; fun mk_unN 1 1 suf = un_prefix ^ suf | mk_unN _ l suf = un_prefix ^ suf ^ string_of_int l; val caseN = "case"; val case_congN = "case_cong"; val case_eq_ifN = "case_eq_if"; val collapseN = "collapse"; val discN = "disc"; val disc_eq_caseN = "disc_eq_case"; val discIN = "discI"; val distinctN = "distinct"; val distinct_discN = "distinct_disc"; val exhaustN = "exhaust"; val exhaust_discN = "exhaust_disc"; val expandN = "expand"; val injectN = "inject"; val nchotomyN = "nchotomy"; val selN = "sel"; val exhaust_selN = "exhaust_sel"; val splitN = "split"; val split_asmN = "split_asm"; val split_selN = "split_sel"; val split_sel_asmN = "split_sel_asm"; val splitsN = "splits"; val split_selsN = "split_sels"; val case_cong_weak_thmsN = "case_cong_weak"; val case_distribN = "case_distrib"; val cong_attrs = @{attributes [cong]}; val dest_attrs = @{attributes [dest]}; val safe_elim_attrs = @{attributes [elim!]}; val iff_attrs = @{attributes [iff]}; val inductsimp_attrs = @{attributes [induct_simp]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun unflat_lookup eq xs ys = map (fn xs' => permute_like_unique eq xs xs' ys); fun mk_half_pairss' _ ([], []) = [] | mk_half_pairss' indent (x :: xs, _ :: ys) = indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys)); fun mk_half_pairss p = mk_half_pairss' [[]] p; fun join_halves n half_xss other_half_xss = (splice (flat half_xss) (flat other_half_xss), map2 (map2 append) (Library.chop_groups n half_xss) (transpose (Library.chop_groups n other_half_xss))); fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_ctr Ts t = let val Type (_, Ts0) = body_type (fastype_of t) in subst_nonatomic_types (Ts0 ~~ Ts) t end; fun mk_case Ts T t = let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t end; fun mk_disc_or_sel Ts t = subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t; val name_of_ctr = name_of_const "constructor" body_type; fun name_of_disc t = (case head_of t of Abs (_, _, \<^const>\Not\ $ (t' $ Bound 0)) => Long_Name.map_base_name (prefix not_prefix) (name_of_disc t') | Abs (_, _, Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t') => Long_Name.map_base_name (prefix is_prefix) (name_of_disc t') | Abs (_, _, \<^const>\Not\ $ (Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t')) => Long_Name.map_base_name (prefix (not_prefix ^ is_prefix)) (name_of_disc t') | t' => name_of_const "discriminator" (perhaps (try domain_type)) t'); val base_name_of_ctr = Long_Name.base_name o name_of_ctr; fun dest_ctr ctxt s t = let val (f, args) = Term.strip_comb t in (case ctr_sugar_of ctxt s of SOME {ctrs, ...} => (case find_first (can (fo_match ctxt f)) ctrs of SOME f' => (f', args) | NONE => raise Fail "dest_ctr") | NONE => raise Fail "dest_ctr") end; fun dest_case ctxt s Ts t = (case Term.strip_comb t of (Const (c, _), args as _ :: _) => (case ctr_sugar_of ctxt s of SOME (ctr_sugar as {casex = Const (case_name, _), discs = discs0, selss = selss0, ...}) => if case_name = c then let val n = length discs0 in if n < length args then let val (branches, obj :: leftovers) = chop n args; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val conds = map (rapp obj) discs; val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss; val branches' = map2 (curry Term.betapplys) branches branch_argss; in SOME (ctr_sugar, conds, branches') end else NONE end else NONE | _ => NONE) | _ => NONE); fun const_or_free_name (Const (s, _)) = Long_Name.base_name s | const_or_free_name (Free (s, _)) = s | const_or_free_name t = raise TERM ("const_or_free_name", [t]) fun extract_sel_default ctxt t = let fun malformed () = error ("Malformed selector default value equation: " ^ Syntax.string_of_term ctxt t); val ((sel, (ctr, vars)), rhs) = fst (Term.replace_dummy_patterns (Syntax.check_term ctxt t) 0) |> HOLogic.dest_eq |>> (Term.dest_comb #>> const_or_free_name ##> (Term.strip_comb #>> (Term.dest_Const #> fst))) handle TERM _ => malformed (); in if forall (is_Free orf is_Var) vars andalso not (has_duplicates (op aconv) vars) then ((ctr, sel), fold_rev Term.lambda vars rhs) else malformed () end; (* Ideally, we would enrich the context with constants rather than free variables. *) fun fake_local_theory_for_sel_defaults sel_bTs = Proof_Context.allow_dummies #> Proof_Context.add_fixes (map (fn (b, T) => (b, SOME T, NoSyn)) sel_bTs) #> snd; type ('c, 'a) ctr_spec = (binding * 'c) * 'a list; fun disc_of_ctr_spec ((disc, _), _) = disc; fun ctr_of_ctr_spec ((_, ctr), _) = ctr; fun args_of_ctr_spec (_, args) = args; val code_plugin = Plugin_Name.declare_setup \<^binding>\code\; fun prepare_free_constructors kind prep_plugins prep_term ((((raw_plugins, discs_sels), raw_case_binding), ctr_specs), sel_default_eqs) no_defs_lthy = let val plugins = prep_plugins no_defs_lthy raw_plugins; (* TODO: sanity checks on arguments *) val raw_ctrs = map ctr_of_ctr_spec ctr_specs; val raw_disc_bindings = map disc_of_ctr_spec ctr_specs; val raw_sel_bindingss = map args_of_ctr_spec ctr_specs; val n = length raw_ctrs; val ks = 1 upto n; val _ = n > 0 orelse error "No constructors specified"; val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs; val (fcT_name, As0) = (case body_type (fastype_of (hd ctrs0)) of Type T' => T' | _ => error "Expected type constructor in body type of constructor"); val _ = forall ((fn Type (T_name, _) => T_name = fcT_name | _ => false) o body_type o fastype_of) (tl ctrs0) orelse error "Constructors not constructing same type"; val fc_b_name = Long_Name.base_name fcT_name; val fc_b = Binding.name fc_b_name; fun qualify mandatory = Binding.qualify mandatory fc_b_name; val (unsorted_As, [B, C]) = no_defs_lthy |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0) ||> fst o mk_TFrees 2; val As = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) As0 unsorted_As; val fcT = Type (fcT_name, As); val ctrs = map (mk_ctr As) ctrs0; val ctr_Tss = map (binder_types o fastype_of) ctrs; val ms = map length ctr_Tss; fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings (k - 1))) orelse nth ms (k - 1) = 0; fun can_rely_on_disc k = can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2)); fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k)); val equal_binding = \<^binding>\=\; fun is_disc_binding_valid b = not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding)); val standard_disc_binding = Binding.name o prefix is_prefix o base_name_of_ctr; val disc_bindings = raw_disc_bindings |> @{map 4} (fn k => fn m => fn ctr => fn disc => qualify false (if Binding.is_empty disc then if m = 0 then equal_binding else if should_omit_disc_binding k then disc else standard_disc_binding ctr else if Binding.eq_name (disc, standard_binding) then standard_disc_binding ctr else disc)) ks ms ctrs0; fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr; val sel_bindingss = @{map 3} (fn ctr => fn m => map2 (fn l => fn sel => qualify false (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then standard_sel_binding m l ctr else sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms raw_sel_bindingss; val add_bindings = Variable.add_fixes (distinct (op =) (filter Symbol_Pos.is_identifier (map Binding.name_of (disc_bindings @ flat sel_bindingss)))) #> snd; val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss; val (((((((((u, exh_y), xss), yss), fs), gs), w), (p, p'))), _) = no_defs_lthy |> add_bindings |> yield_singleton (mk_Frees fc_b_name) fcT ||>> yield_singleton (mk_Frees "y") fcT (* for compatibility with "datatype_realizer.ML" *) ||>> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "z") B ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT; val q = Free (fst p', mk_pred1T B); val xctrs = map2 (curry Term.list_comb) ctrs xss; val yctrs = map2 (curry Term.list_comb) ctrs yss; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides nicer names). Consider removing. *) val eta_fs = map2 (fold_rev Term.lambda) xss xfs; val eta_gs = map2 (fold_rev Term.lambda) xss xgs; val case_binding = qualify false (if Binding.is_empty raw_case_binding orelse Binding.eq_name (raw_case_binding, standard_binding) then Binding.prefix_name (caseN ^ "_") fc_b else raw_case_binding); fun mk_case_disj xctr xf xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf))); val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]] (Const (\<^const_name>\The\, (B --> HOLogic.boolT) --> B) $ Term.lambda w (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_case_disj xctrs xfs xss))); val ((raw_case, (_, raw_case_def)), (lthy, lthy_old)) = no_defs_lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((case_binding, NoSyn), ((Binding.concealed (Thm.def_binding case_binding), []), case_rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val case_def = Morphism.thm phi raw_case_def; val case0 = Morphism.term phi raw_case; val casex = mk_case As B case0; val casexC = mk_case As C case0; val casexBool = mk_case As HOLogic.boolT case0; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val exist_xs_u_eq_ctrs = map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss; val unique_disc_no_def = TrueI; (*arbitrary marker*) val alternate_disc_no_def = FalseE; (*arbitrary marker*) fun alternate_disc_lhs get_udisc k = HOLogic.mk_not (let val b = nth disc_bindings (k - 1) in if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1) end); val no_discs_sels = not discs_sels andalso forall (forall Binding.is_empty) (raw_disc_bindings :: raw_sel_bindingss) andalso null sel_default_eqs; val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy) = if no_discs_sels then (true, [], [], [], [], [], lthy) else let val all_sel_bindings = flat sel_bindingss; val num_all_sel_bindings = length all_sel_bindings; val uniq_sel_bindings = distinct Binding.eq_name all_sel_bindings; val all_sels_distinct = (length uniq_sel_bindings = num_all_sel_bindings); val sel_binding_index = if all_sels_distinct then 1 upto num_all_sel_bindings else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) all_sel_bindings; val all_proto_sels = flat (@{map 3} (fn k => fn xs => map (pair k o pair xs)) ks xss xss); val sel_infos = AList.group (op =) (sel_binding_index ~~ all_proto_sels) |> sort (int_ord o apply2 fst) |> map snd |> curry (op ~~) uniq_sel_bindings; val sel_bindings = map fst sel_infos; val sel_defaults = if null sel_default_eqs then [] else let val sel_Ts = map (curry (op -->) fcT o fastype_of o snd o snd o hd o snd) sel_infos; val fake_lthy = fake_local_theory_for_sel_defaults (sel_bindings ~~ sel_Ts) no_defs_lthy; in map (extract_sel_default fake_lthy o prep_term fake_lthy) sel_default_eqs end; fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT); fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr); fun alternate_disc k = Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k)); fun mk_sel_case_args b proto_sels T = @{map 3} (fn Const (c, _) => fn Ts => fn k => (case AList.lookup (op =) proto_sels k of NONE => (case filter (curry (op =) (c, Binding.name_of b) o fst) sel_defaults of [] => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T) | [(_, t)] => t | _ => error "Multiple default values for selector/constructor pair") | SOME (xs, x) => fold_rev Term.lambda xs x)) ctrs ctr_Tss ks; fun sel_spec b proto_sels = let val _ = (case duplicates (op =) (map fst proto_sels) of k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^ " for constructor " ^ quote (Syntax.string_of_term lthy (nth ctrs (k - 1)))) | [] => ()) val T = (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of [T] => T | T :: T' :: _ => error ("Inconsistent range type for selector " ^ quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. " ^ quote (Syntax.string_of_typ lthy T'))); in mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u, Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u) end; fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss; val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) = lthy |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 3} (fn k => fn exist_xs_u_eq_ctr => fn b => if Binding.is_empty b then if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def) else pair (alternate_disc k, alternate_disc_no_def) else if Binding.eq_name (b, equal_binding) then pair (Term.lambda u exist_xs_u_eq_ctr, refl) else Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr) #>> apsnd snd) ks exist_xs_u_eq_ctrs disc_bindings ||>> apfst split_list o fold_map (fn (b, proto_sels) => Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), sel_spec b proto_sels) #>> apsnd snd) sel_infos ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy lthy'; val disc_defs = map (Morphism.thm phi) raw_disc_defs; val sel_defs = map (Morphism.thm phi) raw_sel_defs; val sel_defss = unflat_selss sel_defs; val discs0 = map (Morphism.term phi) raw_discs; val selss0 = unflat_selss (map (Morphism.term phi) raw_sels); val discs = map (mk_disc_or_sel As) discs0; val selss = map (map (mk_disc_or_sel As)) selss0; in (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') end; fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); val exhaust_goal = let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (exh_y, xctr)]) in fold_rev Logic.all [p, exh_y] (mk_imp_p (map2 mk_prem xctrs xss)) end; val inject_goalss = let fun mk_goal _ _ [] [] = [] | mk_goal xctr yctr xs ys = [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr), Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))]; in @{map 4} mk_goal xctrs yctrs xss yss end; val half_distinct_goalss = let fun mk_goal ((xs, xc), (xs', xc')) = fold_rev Logic.all (xs @ xs') (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc')))); in map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs))) end; val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss; fun after_qed ([exhaust_thm] :: thmss) lthy = let val ((((((((u, u'), (xss, xss')), fs), gs), h), v), p), _) = lthy |> add_bindings |> yield_singleton (apfst (op ~~) oo mk_Frees' fc_b_name) fcT ||>> mk_Freess' "x" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "h") (B --> C) ||>> yield_singleton (mk_Frees (fc_b_name ^ "'")) fcT ||>> yield_singleton (mk_Frees "P") HOLogic.boolT; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; val fcase = Term.list_comb (casex, fs); val ufcase = fcase $ u; val vfcase = fcase $ v; val eta_fcase = Term.list_comb (casex, eta_fs); val eta_gcase = Term.list_comb (casex, eta_gs); val eta_ufcase = eta_fcase $ u; val eta_vgcase = eta_gcase $ v; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val uv_eq = mk_Trueprop_eq (u, v); val ((inject_thms, inject_thmss), half_distinct_thmss) = chop n thmss |>> `flat; val rho_As = map (fn (T, U) => (dest_TVar T, Thm.ctyp_of lthy U)) (map Logic.varifyT_global As ~~ As); fun inst_thm t thm = Thm.instantiate' [] [SOME (Thm.cterm_of lthy t)] (Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm)); val uexhaust_thm = inst_thm u exhaust_thm; val exhaust_cases = map base_name_of_ctr ctrs; val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss; val (distinct_thms, (distinct_thmsss', distinct_thmsss)) = join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose; val nchotomy_thm = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_nchotomy_tac ctxt n exhaust_thm) |> Thm.close_derivation \<^here> end; val case_thms = let val goals = @{map 3} (fn xctr => fn xf => fn xs => fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss; in @{map 4} (fn k => fn goal => fn injects => fn distinctss => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_case_tac ctxt n k case_def injects distinctss) |> Thm.close_derivation \<^here>) ks goals inject_thmss distinct_thmsss end; val (case_cong_thm, case_cong_weak_thm) = let fun mk_prem xctr xs xf xg = fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr), mk_Trueprop_eq (xf, xg))); val goal = Logic.list_implies (uv_eq :: @{map 4} mk_prem xctrs xss xfs xgs, mk_Trueprop_eq (eta_ufcase, eta_vgcase)); val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); val vars = Variable.add_free_names lthy goal []; val weak_vars = Variable.add_free_names lthy weak_goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_case_cong_tac ctxt uexhaust_thm case_thms), Goal.prove_sorry lthy weak_vars [] weak_goal (fn {context = ctxt, prems = _} => etac ctxt arg_cong 1)) |> apply2 (Thm.close_derivation \<^here>) end; val split_lhs = q $ ufcase; fun mk_split_conjunct xctr xs f_xs = list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs)); fun mk_split_disjunct xctr xs f_xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_not (q $ f_xs))); fun mk_split_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj (@{map 3} mk_split_conjunct xctrs xss xfs)); fun mk_split_asm_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_split_disjunct xctrs xss xfs))); fun prove_split selss goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_split_tac ctxt uexhaust_thm case_thms selss inject_thmss distinct_thmsss)) |> Thm.close_derivation \<^here>; fun prove_split_asm asm_goal split_thm = Variable.add_free_names lthy asm_goal [] |> (fn vars => Goal.prove_sorry lthy vars [] asm_goal (fn {context = ctxt, ...} => mk_split_asm_tac ctxt split_thm)) |> Thm.close_derivation \<^here>; val (split_thm, split_asm_thm) = let val goal = mk_split_goal xctrs xss xfs; val asm_goal = mk_split_asm_goal xctrs xss xfs; val thm = prove_split (replicate n []) goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, exhaust_disc_thms, exhaust_sel_thms, all_collapse_thms, safe_collapse_thms, expand_thms, split_sel_thms, split_sel_asm_thms, case_eq_if_thms, disc_eq_case_thms) = if no_discs_sels then ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []) else let val udiscs = map (rapp u) discs; val uselss = map (map (rapp u)) selss; val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss; val usel_fs = map2 (curry Term.list_comb) fs uselss; val vdiscs = map (rapp v) discs; val vselss = map (map (rapp v)) selss; fun make_sel_thm xs' case_thm sel_def = zero_var_indexes (Variable.gen_all lthy (Drule.rename_bvars' (map (SOME o fst) xs') - (Drule.forall_intr_vars (case_thm RS (sel_def RS trans))))); + (Thm.forall_intr_vars (case_thm RS (sel_def RS trans))))); val sel_thmss = @{map 3} (map oo make_sel_thm) xss' case_thms sel_defss; fun has_undefined_rhs thm = (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))) of Const (\<^const_name>\undefined\, _) => true | _ => false); val all_sel_thms = (if all_sels_distinct andalso null sel_default_eqs then flat sel_thmss else map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs (xss' ~~ case_thms)) |> filter_out has_undefined_rhs; fun mk_unique_disc_def () = let val m = the_single ms; val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_unique_disc_def_tac ctxt m uexhaust_thm) |> Thm.close_derivation \<^here> end; fun mk_alternate_disc_def k = let val goal = mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k), nth exist_xs_u_eq_ctrs (k - 1)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) (nth distinct_thms (2 - k)) uexhaust_thm) |> Thm.close_derivation \<^here> end; val has_alternate_disc_def = exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs; val nontriv_disc_defs = disc_defs |> filter_out (member Thm.eq_thm_prop [unique_disc_no_def, alternate_disc_no_def, refl]); val disc_defs' = map2 (fn k => fn def => if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def () else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k else def) ks disc_defs; val discD_thms = map (fn def => def RS iffD1) disc_defs'; val discI_thms = map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms disc_defs'; val not_discI_thms = map2 (fn m => fn def => funpow m (fn thm => allI RS thm) (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]}))) ms disc_defs'; val (disc_thmss', disc_thmss) = let fun mk_thm discI _ [] = refl RS discI | mk_thm _ not_discI [distinct] = distinct RS not_discI; fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss; in @{map 3} mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose end; val nontriv_disc_thmss = map2 (fn b => if is_disc_binding_valid b then I else K []) disc_bindings disc_thmss; fun is_discI_triv b = (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding); val nontriv_discI_thms = flat (map2 (fn b => if is_discI_triv b then K [] else single) disc_bindings discI_thms); val (distinct_disc_thms, (distinct_disc_thmsss', distinct_disc_thmsss)) = let fun mk_goal [] = [] | mk_goal [((_, udisc), (_, udisc'))] = [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc, HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; fun prove tac goal = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => tac ctxt) |> Thm.close_derivation \<^here>; val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); val half_goalss = map mk_goal half_pairss; val half_thmss = @{map 3} (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => fn disc_thm => [prove (fn ctxt => mk_half_distinct_disc_tac ctxt m discD disc_thm) goal]) half_goalss half_pairss (flat disc_thmss'); val other_half_goalss = map (mk_goal o map swap) half_pairss; val other_half_thmss = map2 (map2 (fn thm => prove (fn ctxt => mk_other_half_distinct_disc_tac ctxt thm))) half_thmss other_half_goalss; in join_halves n half_thmss other_half_thmss ||> `transpose |>> has_alternate_disc_def ? K [] end; val exhaust_disc_thm = let fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_disc_tac ctxt n exhaust_thm discI_thms) |> Thm.close_derivation \<^here> end; val (safe_collapse_thms, all_collapse_thms) = let fun mk_goal m udisc usel_ctr = let val prem = HOLogic.mk_Trueprop udisc; val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap); in (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl))) end; val (trivs, goals) = @{map 3} mk_goal ms udiscs usel_ctrs |> split_list; val thms = @{map 5} (fn m => fn discD => fn sel_thms => fn triv => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL (assume_tac ctxt)) |> Thm.close_derivation \<^here> |> not triv ? perhaps (try (fn thm => refl RS thm))) ms discD_thms sel_thmss trivs goals; in (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms), thms) end; val swapped_all_collapse_thms = map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms; val exhaust_sel_thm = let fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_sel_tac ctxt n exhaust_disc_thm swapped_all_collapse_thms) |> Thm.close_derivation \<^here> end; val expand_thm = let fun mk_prems k udisc usels vdisc vsels = (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @ (if null usels then [] else [Logic.list_implies (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc], HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) usels vsels)))]); val goal = Library.foldr Logic.list_implies (@{map 5} mk_prems ks udiscs uselss vdiscs vselss, uv_eq); val uncollapse_thms = map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_expand_tac ctxt n ms (inst_thm u exhaust_disc_thm) (inst_thm v exhaust_disc_thm) uncollapse_thms distinct_disc_thmsss distinct_disc_thmsss') |> Thm.close_derivation \<^here> end; val (split_sel_thm, split_sel_asm_thm) = let val zss = map (K []) xss; val goal = mk_split_goal usel_ctrs zss usel_fs; val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs; val thm = prove_split sel_thmss goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val case_eq_if_thm = let val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss) |> Thm.close_derivation \<^here> end; val disc_eq_case_thms = let fun const_of_bool b = if b then \<^const>\True\ else \<^const>\False\; fun mk_case_args n = map_index (fn (k, argTs) => fold_rev Term.absdummy argTs (const_of_bool (n = k))) ctr_Tss; val goals = map_index (fn (n, udisc) => mk_Trueprop_eq (udisc, list_comb (casexBool, mk_case_args n) $ u)) udiscs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm (flat nontriv_disc_thmss) distinct_thms case_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; in (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, [exhaust_disc_thm], [exhaust_sel_thm], all_collapse_thms, safe_collapse_thms, [expand_thm], [split_sel_thm], [split_sel_asm_thm], [case_eq_if_thm], disc_eq_case_thms) end; val case_distrib_thm = let val args = @{map 2} (fn f => fn argTs => let val (args, _) = mk_Frees "x" argTs lthy in fold_rev Term.lambda args (h $ list_comb (f, args)) end) fs ctr_Tss; val goal = mk_Trueprop_eq (h $ ufcase, list_comb (casexC, args) $ u); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_distrib_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm case_thms) |> Thm.close_derivation \<^here> end; val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases)); val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name)); val nontriv_disc_eq_thmss = map (map (fn th => th RS @{thm eq_False[THEN iffD2]} handle THM _ => th RS @{thm eq_True[THEN iffD2]})) nontriv_disc_thmss; val anonymous_notes = [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs), (flat nontriv_disc_eq_thmss, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = [(caseN, case_thms, nitpicksimp_attrs @ simp_attrs), (case_congN, [case_cong_thm], []), (case_cong_weak_thmsN, [case_cong_weak_thm], cong_attrs), (case_distribN, [case_distrib_thm], []), (case_eq_ifN, case_eq_if_thms, []), (collapseN, safe_collapse_thms, if ms = [0] then [] else simp_attrs), (discN, flat nontriv_disc_thmss, simp_attrs), (disc_eq_caseN, disc_eq_case_thms, []), (discIN, nontriv_discI_thms, []), (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs), (distinct_discN, distinct_disc_thms, dest_attrs), (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]), (exhaust_discN, exhaust_disc_thms, [exhaust_case_names_attr]), (exhaust_selN, exhaust_sel_thms, [exhaust_case_names_attr]), (expandN, expand_thms, []), (injectN, inject_thms, iff_attrs @ inductsimp_attrs), (nchotomyN, [nchotomy_thm], []), (selN, all_sel_thms, nitpicksimp_attrs @ simp_attrs), (splitN, [split_thm], []), (split_asmN, [split_asm_thm], []), (split_selN, split_sel_thms, []), (split_sel_asmN, split_sel_asm_thms, []), (split_selsN, split_sel_thms @ split_sel_asm_thms, []), (splitsN, [split_thm, split_asm_thm], [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify true (Binding.name thmN), attrs), [(thms, [])])); val (noted, lthy') = lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational [casex] case_thms |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (AList.group (eq_list (op aconv)) (map (`(single o lhs_head_of)) all_sel_thms)) |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (filter_out (null o snd) (map single discs ~~ nontriv_disc_eq_thmss)) |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Case_Translation.register (Morphism.term phi casex) (map (Morphism.term phi) ctrs)) |> plugins code_plugin ? (Code.declare_default_eqns (map (rpair true) (flat nontriv_disc_eq_thmss @ case_thms @ all_sel_thms)) #> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => Context.mapping (add_ctr_code fcT_name (map (Morphism.typ phi) As) (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms) (Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms)) I)) |> Local_Theory.notes (anonymous_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms fcT_name exhaustN; val ctr_sugar = {kind = kind, T = fcT, ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm, nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms, case_thms = case_thms, case_cong = case_cong_thm, case_cong_weak = case_cong_weak_thm, case_distribs = [case_distrib_thm], split = split_thm, split_asm = split_asm_thm, disc_defs = nontriv_disc_defs, disc_thmss = disc_thmss, discIs = discI_thms, disc_eq_cases = disc_eq_case_thms, sel_defs = sel_defs, sel_thmss = sel_thmss, distinct_discsss = distinct_disc_thmsss, exhaust_discs = exhaust_disc_thms, exhaust_sels = exhaust_sel_thms, collapses = all_collapse_thms, expands = expand_thms, split_sels = split_sel_thms, split_sel_asms = split_sel_asm_thms, case_eq_ifs = case_eq_if_thms} |> morph_ctr_sugar (substitute_noted_thm noted); in (ctr_sugar, lthy' |> register_ctr_sugar plugins ctr_sugar) end; in (goalss, after_qed, lthy) end; fun free_constructors kind tacss = (fn (goalss, after_qed, lthy) => map2 (map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] [])) goalss tacss |> (fn thms => after_qed thms lthy)) oo prepare_free_constructors kind (K I) (K I); fun free_constructors_cmd kind = (fn (goalss, after_qed, lthy) => Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo prepare_free_constructors kind Plugin_Name.make_filter Syntax.read_term; val parse_bound_term = Parse.binding --| \<^keyword>\:\ -- Parse.term; type ctr_options = Plugin_Name.filter * bool; type ctr_options_cmd = (Proof.context -> Plugin_Name.filter) * bool; val default_ctr_options : ctr_options = (Plugin_Name.default_filter, false); val default_ctr_options_cmd : ctr_options_cmd = (K Plugin_Name.default_filter, false); val parse_ctr_options = Scan.optional (\<^keyword>\(\ |-- Parse.list1 (Plugin_Name.parse_filter >> (apfst o K) || Parse.reserved "discs_sels" >> (apsnd o K o K true)) --| \<^keyword>\)\ >> (fn fs => fold I fs default_ctr_options_cmd)) default_ctr_options_cmd; fun parse_ctr_spec parse_ctr parse_arg = parse_opt_binding_colon -- parse_ctr -- Scan.repeat parse_arg; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.term Parse.binding); val parse_sel_default_eqs = Scan.optional (\<^keyword>\where\ |-- Parse.enum1 "|" Parse.prop) []; val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\free_constructors\ "register an existing freely generated type's constructors" (parse_ctr_options -- Parse.binding --| \<^keyword>\for\ -- parse_ctr_specs -- parse_sel_default_eqs >> free_constructors_cmd Unknown); (** external views **) (* document antiquotations *) local fun antiquote_setup binding co = Document_Output.antiquotation_pretty_source_embedded binding ((Scan.ahead (Scan.lift Parse.not_eof) >> Token.pos_of) -- Args.type_name {proper = true, strict = true}) (fn ctxt => fn (pos, type_name) => let fun err () = error ("Bad " ^ Binding.name_of binding ^ ": " ^ quote type_name ^ Position.here pos); in (case ctr_sugar_of ctxt type_name of NONE => err () | SOME {kind, T = T0, ctrs = ctrs0, ...} => let val _ = if co = (kind = Codatatype) then () else err (); val T = Logic.unvarifyT_global T0; val ctrs = map Logic.unvarify_global ctrs0; val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt); fun pretty_ctr ctr = Pretty.block (Pretty.breaks (Syntax.pretty_term ctxt ctr :: map pretty_typ_bracket (binder_types (fastype_of ctr)))); in Pretty.block (Pretty.keyword1 (Binding.name_of binding) :: Pretty.brk 1 :: Syntax.pretty_typ ctxt T :: Pretty.str " =" :: Pretty.brk 1 :: flat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o pretty_ctr) ctrs))) end) end); in val _ = Theory.setup (antiquote_setup \<^binding>\datatype\ false #> antiquote_setup \<^binding>\codatatype\ true); end; (* theory export *) val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => if Export_Theory.export_enabled context then let val parents = map (Data.get o Context.Theory) (Theory.parents_of thy); val datatypes = (Data.get (Context.Theory thy), []) |-> Symtab.fold (fn (name, (pos, {kind, T, ctrs, ...})) => if kind = Record orelse exists (fn tab => Symtab.defined tab name) parents then I else let val pos_properties = Thy_Info.adjust_pos_properties context pos; val typ = Logic.unvarifyT_global T; val constrs = map Logic.unvarify_global ctrs; val typargs = rev (fold Term.add_tfrees (Logic.mk_type typ :: constrs) []); val constructors = map (fn t => (t, Term.type_of t)) constrs; in cons (pos_properties, (name, (kind = Codatatype, (typargs, (typ, constructors))))) end); in if null datatypes then () else Export_Theory.export_body thy "datatypes" let open XML.Encode Term_XML.Encode in list (pair properties (pair string (pair bool (pair (list (pair string sort)) (pair typ (list (pair (term (Sign.consts_of thy)) typ))))))) datatypes end end else ()); end; diff --git a/src/HOL/Tools/Function/function_core.ML b/src/HOL/Tools/Function/function_core.ML --- a/src/HOL/Tools/Function/function_core.ML +++ b/src/HOL/Tools/Function/function_core.ML @@ -1,925 +1,925 @@ (* Title: HOL/Tools/Function/function_core.ML Author: Alexander Krauss, TU Muenchen Core of the function package. *) signature FUNCTION_CORE = sig val trace: bool Unsynchronized.ref val prepare_function : Function_Common.function_config -> binding (* defname *) -> ((binding * typ) * mixfix) list (* defined symbol *) -> ((string * typ) list * term list * term * term) list (* specification *) -> local_theory -> (term (* f *) * thm (* goalstate *) * (Proof.context -> thm -> Function_Common.function_result) (* continuation *) ) * local_theory end structure Function_Core : FUNCTION_CORE = struct val trace = Unsynchronized.ref false fun trace_msg msg = if ! trace then tracing (msg ()) else () val boolT = HOLogic.boolT val mk_eq = HOLogic.mk_eq open Function_Lib open Function_Common datatype globals = Globals of {fvar: term, domT: typ, ranT: typ, h: term, y: term, x: term, z: term, a: term, P: term, D: term, Pbool:term} datatype rec_call_info = RCInfo of {RIvs: (string * typ) list, (* Call context: fixes and assumes *) CCas: thm list, rcarg: term, (* The recursive argument *) llRI: thm, h_assum: term} datatype clause_context = ClauseContext of {ctxt : Proof.context, qs : term list, gs : term list, lhs: term, rhs: term, cqs: cterm list, ags: thm list, case_hyp : thm} fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) = ClauseContext { ctxt = Proof_Context.transfer thy ctxt, qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp } datatype clause_info = ClauseInfo of {no: int, qglr : ((string * typ) list * term list * term * term), cdata : clause_context, tree: Function_Context_Tree.ctx_tree, lGI: thm, RCs: rec_call_info list} (* Theory dependencies. *) val acc_induct_rule = @{thm accp_induct_rule} val ex1_implies_ex = @{thm Fun_Def.fundef_ex1_existence} val ex1_implies_un = @{thm Fun_Def.fundef_ex1_uniqueness} val ex1_implies_iff = @{thm Fun_Def.fundef_ex1_iff} val acc_downward = @{thm accp_downward} val accI = @{thm accp.accI} fun find_calls tree = let fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs) | add_Ri _ _ _ _ = raise Match in rev (Function_Context_Tree.traverse_tree add_Ri tree []) end (** building proof obligations *) fun mk_compat_proof_obligations domT ranT fvar f glrs = let fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) = let val shift = incr_boundvars (length qs') in Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'), HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs')) |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs') |> fold_rev (fn (n,T) => fn b => Logic.all_const T $ Abs(n,T,b)) (qs @ qs') |> curry abstract_over fvar |> curry subst_bound f end in map mk_impl (unordered_pairs glrs) end fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs = let fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) = HOLogic.mk_Trueprop Pbool |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs))) |> fold_rev (curry Logic.mk_implies) gs |> fold_rev mk_forall_rename (map fst oqs ~~ qs) in HOLogic.mk_Trueprop Pbool |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs) |> mk_forall_rename ("x", x) |> mk_forall_rename ("P", Pbool) end (** making a context with it's own local bindings **) fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) = let val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs fun inst t = subst_bounds (rev qs, t) val gs = map inst pre_gs val lhs = inst pre_lhs val rhs = inst pre_rhs val cqs = map (Thm.cterm_of ctxt') qs val ags = map (Thm.assume o Thm.cterm_of ctxt') gs val case_hyp = Thm.assume (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (mk_eq (x, lhs)))) in ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp } end (* lowlevel term function. FIXME: remove *) fun abstract_over_list vs body = let fun abs lev v tm = if v aconv tm then Bound lev else (case tm of Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t) | t $ u => abs lev v t $ abs lev v u | t => t) in fold_index (fn (i, v) => fn t => abs i v t) vs body end fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms = let val Globals {h, ...} = globals val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata (* Instantiate the GIntro thm with "f" and import into the clause context. *) val lGI = GIntro_thm |> Thm.forall_elim (Thm.cterm_of ctxt f) |> fold Thm.forall_elim cqs |> fold Thm.elim_implies ags fun mk_call_info (rcfix, rcassm, rcarg) RI = let val llRI = RI |> fold Thm.forall_elim cqs |> fold (Thm.forall_elim o Thm.cterm_of ctxt o Free) rcfix |> fold Thm.elim_implies ags |> fold Thm.elim_implies rcassm val h_assum = HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg)) |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm |> fold_rev (Logic.all o Free) rcfix |> Pattern.rewrite_term (Proof_Context.theory_of ctxt) [(f, h)] [] |> abstract_over_list (rev qs) in RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum} end val RC_infos = map2 mk_call_info RCs RIntro_thms in ClauseInfo {no=no, cdata=cdata, qglr=qglr, lGI=lGI, RCs=RC_infos, tree=tree} end fun store_compat_thms 0 thms = [] | store_compat_thms n thms = let val (thms1, thms2) = chop n thms in (thms1 :: store_compat_thms (n - 1) thms2) end (* expects i <= j *) fun lookup_compat_thm i j cts = nth (nth cts (i - 1)) (j - i) (* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *) (* if j < i, then turn around *) fun get_compat_thm ctxt cts i j ctxi ctxj = let val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj val lhsi_eq_lhsj = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj))) in if j < i then let val compat = lookup_compat_thm j i cts in compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *) |> fold Thm.forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *) |> fold Thm.elim_implies agsj |> fold Thm.elim_implies agsi |> Thm.elim_implies ((Thm.assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *) end else let val compat = lookup_compat_thm i j cts in compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *) |> fold Thm.forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *) |> fold Thm.elim_implies agsi |> fold Thm.elim_implies agsj |> Thm.elim_implies (Thm.assume lhsi_eq_lhsj) |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *) end end (* Generates the replacement lemma in fully quantified form. *) fun mk_replacement_lemma ctxt h ih_elim clause = let val ClauseInfo {cdata=ClauseContext {qs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause local open Conv in val ih_conv = arg1_conv o arg_conv o arg_conv end val ih_elim_case = Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs val h_assums = map (fn RCInfo {h_assum, ...} => Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qs, h_assum)))) RCs val (eql, _) = Function_Context_Tree.rewrite_by_tree ctxt h ih_elim_case (Ris ~~ h_assums) tree val replace_lemma = HOLogic.mk_obj_eq eql |> Thm.implies_intr (Thm.cprop_of case_hyp) |> fold_rev (Thm.implies_intr o Thm.cprop_of) h_assums |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags |> fold_rev Thm.forall_intr cqs |> Thm.close_derivation \<^here> in replace_lemma end fun mk_uniqueness_clause ctxt globals compat_store clausei clausej RLj = let val thy = Proof_Context.theory_of ctxt val Globals {h, y, x, fvar, ...} = globals val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} = mk_clause_context x ctxti cdescj val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj' val compat = get_compat_thm ctxt compat_store i j cctxi cctxj val Ghsj' = map (fn RCInfo {h_assum, ...} => Thm.assume (Thm.cterm_of ctxt (subst_bounds (rev qsj', h_assum)))) RCsj val RLj_import = RLj |> fold Thm.forall_elim cqsj' |> fold Thm.elim_implies agsj' |> fold Thm.elim_implies Ghsj' val y_eq_rhsj'h = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h)))) val lhsi_eq_lhsj' = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *) in (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *) |> Thm.implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *) |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *) |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *) |> fold_rev (Thm.implies_intr o Thm.cprop_of) Ghsj' |> fold_rev (Thm.implies_intr o Thm.cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *) |> Thm.implies_intr (Thm.cprop_of y_eq_rhsj'h) |> Thm.implies_intr (Thm.cprop_of lhsi_eq_lhsj') |> fold_rev Thm.forall_intr (Thm.cterm_of ctxt h :: cqsj') end fun mk_uniqueness_case ctxt globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei = let val thy = Proof_Context.theory_of ctxt val Globals {x, y, ranT, fvar, ...} = globals val ClauseInfo {cdata = ClauseContext {lhs, rhs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs val ih_intro_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ih_intro fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case) |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs val existence = fold (curry op COMP o prep_RC) RCs lGI val P = Thm.cterm_of ctxt (mk_eq (y, rhsC)) val G_lhs_y = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (G $ lhs $ y))) val unique_clauses = map2 (mk_uniqueness_clause ctxt globals compat_store clausei) clauses rep_lemmas fun elim_implies_eta A AB = Thm.bicompose (SOME ctxt) {flatten = false, match = true, incremented = false} (false, A, 0) 1 AB |> Seq.list_of |> the_single val uniqueness = G_cases |> Thm.forall_elim (Thm.cterm_of ctxt lhs) |> Thm.forall_elim (Thm.cterm_of ctxt y) |> Thm.forall_elim P |> Thm.elim_implies G_lhs_y |> fold elim_implies_eta unique_clauses |> Thm.implies_intr (Thm.cprop_of G_lhs_y) |> Thm.forall_intr (Thm.cterm_of ctxt y) val P2 = Thm.cterm_of ctxt (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *) val exactly_one = @{thm ex1I} |> Thm.instantiate' [SOME (Thm.ctyp_of ctxt ranT)] [SOME P2, SOME (Thm.cterm_of ctxt rhsC)] |> curry (op COMP) existence |> curry (op COMP) uniqueness |> simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp RS sym]) |> Thm.implies_intr (Thm.cprop_of case_hyp) |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags |> fold_rev Thm.forall_intr cqs val function_value = existence |> Thm.implies_intr ihyp |> Thm.implies_intr (Thm.cprop_of case_hyp) |> Thm.forall_intr (Thm.cterm_of ctxt x) |> Thm.forall_elim (Thm.cterm_of ctxt lhs) |> curry (op RS) refl in (exactly_one, function_value) end fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def = let val Globals {h, domT, ranT, x, ...} = globals (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *) val ihyp = Logic.all_const domT $ Abs ("z", domT, Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x), HOLogic.mk_Trueprop (Const (\<^const_name>\Ex1\, (ranT --> boolT) --> boolT) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0)))) |> Thm.cterm_of ctxt val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0 val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex) val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un) |> Thm.instantiate' [] [NONE, SOME (Thm.cterm_of ctxt h)] val _ = trace_msg (K "Proving Replacement lemmas...") val repLemmas = map (mk_replacement_lemma ctxt h ih_elim) clauses val _ = trace_msg (K "Proving cases for unique existence...") val (ex1s, values) = split_list (map (mk_uniqueness_case ctxt globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses) val _ = trace_msg (K "Proving: Graph is a function") val graph_is_function = complete |> Thm.forall_elim_vars 0 |> fold (curry op COMP) ex1s |> Thm.implies_intr (ihyp) |> Thm.implies_intr (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ x))) |> Thm.forall_intr (Thm.cterm_of ctxt x) |> (fn it => Drule.compose (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *) |> (fn it => fold (Thm.forall_intr o Thm.cterm_of ctxt o Var) (Term.add_vars (Thm.prop_of it) []) it) val goalstate = Conjunction.intr graph_is_function complete |> Thm.close_derivation \<^here> |> Goal.protect 0 |> fold_rev (Thm.implies_intr o Thm.cprop_of) compat |> Thm.implies_intr (Thm.cprop_of complete) in (goalstate, values) end (* wrapper -- restores quantifiers in rule specifications *) fun inductive_def (binding as ((R, T), _)) intrs lthy = let val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) = lthy |> Proof_Context.concealed |> Inductive.add_inductive {quiet_mode = true, verbose = ! trace, alt_name = Binding.empty, coind = false, no_elim = false, no_ind = false, skip_mono = true} [binding] (* relation *) [] (* no parameters *) (map (fn t => (Binding.empty_atts, t)) intrs) (* intro rules *) [] (* no special monos *) ||> Proof_Context.restore_naming lthy fun requantify orig_intro thm = let val (qs, t) = dest_all_all orig_intro val frees = Variable.add_frees lthy t [] |> remove (op =) (Binding.name_of R, T) val vars = Term.add_vars (Thm.prop_of thm) [] val varmap = AList.lookup (op =) (frees ~~ map fst vars) #> the_default ("", 0) in fold_rev (fn Free (n, T) => forall_intr_rename (n, Thm.cterm_of lthy (Var (varmap (n, T), T)))) qs thm end in - ((Rdef, map2 requantify intrs intrs_gen, forall_intr_vars elim_gen, induct), lthy) + ((Rdef, map2 requantify intrs intrs_gen, Thm.forall_intr_vars elim_gen, induct), lthy) end fun define_graph (G_binding, G_type) fvar clauses RCss lthy = let val G = Free (Binding.name_of G_binding, G_type) fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs = let fun mk_h_assm (rcfix, rcassm, rcarg) = HOLogic.mk_Trueprop (G $ rcarg $ (fvar $ rcarg)) |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm |> fold_rev (Logic.all o Free) rcfix in HOLogic.mk_Trueprop (G $ lhs $ rhs) |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs |> fold_rev (curry Logic.mk_implies) gs |> fold_rev Logic.all (fvar :: qs) end val G_intros = map2 mk_GIntro clauses RCss in inductive_def ((G_binding, G_type), NoSyn) G_intros lthy end fun define_function defname (fname, mixfix) domT ranT G default lthy = let val f_def_binding = Thm.make_def_binding (Config.get lthy function_internals) (derived_name_suffix defname "_sumC") val f_def = Abs ("x", domT, Const (\<^const_name>\Fun_Def.THE_default\, ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0)) |> Syntax.check_term lthy in lthy |> Local_Theory.define ((Binding.map_name (suffix "C") fname, mixfix), ((f_def_binding, []), f_def)) end fun define_recursion_relation (R_binding, R_type) qglrs clauses RCss lthy = let val R = Free (Binding.name_of R_binding, R_type) fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) = HOLogic.mk_Trueprop (R $ rcarg $ lhs) |> fold_rev (curry Logic.mk_implies o Thm.prop_of) rcassm |> fold_rev (curry Logic.mk_implies) gs |> fold_rev (Logic.all o Free) rcfix |> fold_rev mk_forall_rename (map fst oqs ~~ qs) (* "!!qs xs. CS ==> G => (r, lhs) : R" *) val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss val ((R, RIntro_thms, R_elim, _), lthy) = inductive_def ((R_binding, R_type), NoSyn) (flat R_intross) lthy in ((R, Library.unflat R_intross RIntro_thms, R_elim), lthy) end fun fix_globals domT ranT fvar ctxt = let val ([h, y, x, z, a, D, P, Pbool], ctxt') = Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt in (Globals {h = Free (h, domT --> ranT), y = Free (y, ranT), x = Free (x, domT), z = Free (z, domT), a = Free (a, domT), D = Free (D, domT --> boolT), P = Free (P, domT --> boolT), Pbool = Free (Pbool, boolT), fvar = fvar, domT = domT, ranT = ranT}, ctxt') end fun inst_RC ctxt fvar f (rcfix, rcassm, rcarg) = let fun inst_term t = subst_bound(f, abstract_over (fvar, t)) in (rcfix, map (Thm.assume o Thm.cterm_of ctxt o inst_term o Thm.prop_of) rcassm, inst_term rcarg) end (********************************************************** * PROVING THE RULES **********************************************************) fun mk_psimps ctxt globals R clauses valthms f_iff graph_is_function = let val Globals {domT, z, ...} = globals fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm = let val lhs_acc = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *) val z_smaller = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *) in ((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward)) |> (fn it => it COMP graph_is_function) |> Thm.implies_intr z_smaller |> Thm.forall_intr (Thm.cterm_of ctxt z) |> (fn it => it COMP valthm) |> Thm.implies_intr lhs_acc |> asm_simplify (put_simpset HOL_basic_ss ctxt addsimps [f_iff]) |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) end in map2 mk_psimp clauses valthms end (** Induction rule **) val acc_subset_induct = @{thm predicate1I} RS @{thm accp_subset_induct} fun mk_partial_induct_rule ctxt globals R complete_thm clauses = let val Globals {domT, x, z, a, P, D, ...} = globals val acc_R = mk_acc domT R val x_D = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (D $ x))) val a_D = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (D $ a)) val D_subset = Thm.cterm_of ctxt (Logic.all x (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x)))) val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *) Logic.all x (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x), HOLogic.mk_Trueprop (D $ z))))) |> Thm.cterm_of ctxt (* Inductive Hypothesis: !!z. (z,x):R ==> P z *) val ihyp = Logic.all_const domT $ Abs ("z", domT, Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x), HOLogic.mk_Trueprop (P $ Bound 0))) |> Thm.cterm_of ctxt val aihyp = Thm.assume ihyp fun prove_case clause = let val ClauseInfo {cdata = ClauseContext {ctxt = ctxt1, qs, cqs, ags, gs, lhs, case_hyp, ...}, RCs, qglr = (oqs, _, _, _), ...} = clause val case_hyp_conv = K (case_hyp RS eq_reflection) local open Conv in val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D val sih = fconv_rule (Conv.binder_conv (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt1) aihyp end fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih |> Thm.forall_elim (Thm.cterm_of ctxt rcarg) |> Thm.elim_implies llRI |> fold_rev (Thm.implies_intr o Thm.cprop_of) CCas |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) RIvs val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *) val step = HOLogic.mk_Trueprop (P $ lhs) |> fold_rev (curry Logic.mk_implies o Thm.prop_of) P_recs |> fold_rev (curry Logic.mk_implies) gs |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs)) |> fold_rev mk_forall_rename (map fst oqs ~~ qs) |> Thm.cterm_of ctxt val P_lhs = Thm.assume step |> fold Thm.forall_elim cqs |> Thm.elim_implies lhs_D |> fold Thm.elim_implies ags |> fold Thm.elim_implies P_recs val res = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (P $ x)) |> Conv.arg_conv (Conv.arg_conv case_hyp_conv) |> Thm.symmetric (* P lhs == P x *) |> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *) |> Thm.implies_intr (Thm.cprop_of case_hyp) |> fold_rev (Thm.implies_intr o Thm.cprop_of) ags |> fold_rev Thm.forall_intr cqs in (res, step) end val (cases, steps) = split_list (map prove_case clauses) val istep = complete_thm |> Thm.forall_elim_vars 0 |> fold (curry op COMP) cases (* P x *) |> Thm.implies_intr ihyp |> Thm.implies_intr (Thm.cprop_of x_D) |> Thm.forall_intr (Thm.cterm_of ctxt x) val subset_induct_rule = acc_subset_induct |> (curry op COMP) (Thm.assume D_subset) |> (curry op COMP) (Thm.assume D_dcl) |> (curry op COMP) (Thm.assume a_D) |> (curry op COMP) istep |> fold_rev Thm.implies_intr steps |> Thm.implies_intr a_D |> Thm.implies_intr D_dcl |> Thm.implies_intr D_subset val simple_induct_rule = subset_induct_rule |> Thm.forall_intr (Thm.cterm_of ctxt D) |> Thm.forall_elim (Thm.cterm_of ctxt acc_R) |> assume_tac ctxt 1 |> Seq.hd |> (curry op COMP) (acc_downward |> (Thm.instantiate' [SOME (Thm.ctyp_of ctxt domT)] (map (SOME o Thm.cterm_of ctxt) [R, x, z])) |> Thm.forall_intr (Thm.cterm_of ctxt z) |> Thm.forall_intr (Thm.cterm_of ctxt x)) |> Thm.forall_intr (Thm.cterm_of ctxt a) |> Thm.forall_intr (Thm.cterm_of ctxt P) in simple_induct_rule end (* FIXME: broken by design *) fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause = let val ClauseInfo {cdata = ClauseContext {gs, lhs, cqs, ...}, qglr = (oqs, _, _, _), ...} = clause val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs) |> fold_rev (curry Logic.mk_implies) gs |> Thm.cterm_of ctxt in Goal.init goal |> (SINGLE (resolve_tac ctxt [accI] 1)) |> the |> (SINGLE (eresolve_tac ctxt [Thm.forall_elim_vars 0 R_cases] 1)) |> the |> (SINGLE (auto_tac ctxt)) |> the |> Goal.conclude |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) end (** Termination rule **) val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule} val wf_in_rel = @{thm Fun_Def.wf_in_rel} val in_rel_def = @{thm Fun_Def.in_rel_def} fun mk_nest_term_case ctxt globals R' ihyp clause = let val Globals {z, ...} = globals val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, lhs, case_hyp, ...}, tree, qglr=(oqs, _, _, _), ...} = clause val ih_case = full_simplify (put_simpset HOL_basic_ss ctxt addsimps [case_hyp]) ihyp fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) = let val used = (u @ sub) |> map (fn (ctx, thm) => Function_Context_Tree.export_thm ctxt ctx thm) val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs) |> fold_rev (curry Logic.mk_implies o Thm.prop_of) used (* additional hyps *) |> Function_Context_Tree.export_term (fixes, assumes) |> fold_rev (curry Logic.mk_implies o Thm.prop_of) ags |> fold_rev mk_forall_rename (map fst oqs ~~ qs) |> Thm.cterm_of ctxt val thm = Thm.assume hyp |> fold Thm.forall_elim cqs |> fold Thm.elim_implies ags |> Function_Context_Tree.import_thm ctxt (fixes, assumes) |> fold Thm.elim_implies used (* "(arg, lhs) : R'" *) val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg)) |> Thm.cterm_of ctxt |> Thm.assume val acc = thm COMP ih_case val z_acc_local = acc |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (Thm.symmetric (z_eq_arg RS eq_reflection))))) val ethm = z_acc_local |> Function_Context_Tree.export_thm ctxt (fixes, z_eq_arg :: case_hyp :: ags @ assumes) |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) val sub' = sub @ [(([],[]), acc)] in (sub', (hyp :: hyps, ethm :: thms)) end | step _ _ _ _ = raise Match in Function_Context_Tree.traverse_tree step tree end fun mk_nest_term_rule ctxt globals R R_cases clauses = let val Globals { domT, x, z, ... } = globals val acc_R = mk_acc domT R val ([Rn], ctxt') = Variable.variant_fixes ["R"] ctxt val R' = Free (Rn, fastype_of R) val Rrel = Free (Rn, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT))) val inrel_R = Const (\<^const_name>\Fun_Def.in_rel\, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel val wfR' = HOLogic.mk_Trueprop (Const (\<^const_name>\Wellfounded.wfP\, (domT --> domT --> boolT) --> boolT) $ R') |> Thm.cterm_of ctxt' (* "wf R'" *) (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *) val ihyp = Logic.all_const domT $ Abs ("z", domT, Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x), HOLogic.mk_Trueprop (acc_R $ Bound 0))) |> Thm.cterm_of ctxt' val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0 val R_z_x = Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (R $ z $ x)) val (hyps, cases) = fold (mk_nest_term_case ctxt' globals R' ihyp_a) clauses ([], []) in R_cases |> Thm.forall_elim (Thm.cterm_of ctxt' z) |> Thm.forall_elim (Thm.cterm_of ctxt' x) |> Thm.forall_elim (Thm.cterm_of ctxt' (acc_R $ z)) |> curry op COMP (Thm.assume R_z_x) |> fold_rev (curry op COMP) cases |> Thm.implies_intr R_z_x |> Thm.forall_intr (Thm.cterm_of ctxt' z) |> (fn it => it COMP accI) |> Thm.implies_intr ihyp |> Thm.forall_intr (Thm.cterm_of ctxt' x) |> (fn it => Drule.compose (it, 2, wf_induct_rule)) |> curry op RS (Thm.assume wfR') - |> forall_intr_vars + |> Thm.forall_intr_vars |> (fn it => it COMP allI) |> fold Thm.implies_intr hyps |> Thm.implies_intr wfR' |> Thm.forall_intr (Thm.cterm_of ctxt' R') |> Thm.forall_elim (Thm.cterm_of ctxt' inrel_R) |> curry op RS wf_in_rel |> full_simplify (put_simpset HOL_basic_ss ctxt' addsimps [in_rel_def]) |> Thm.forall_intr_name ("R", Thm.cterm_of ctxt' Rrel) end fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy = let val FunctionConfig {domintros, default=default_opt, ...} = config val default_str = the_default "%x. HOL.undefined" (* FIXME proper term!? *) default_opt val fvar = (Binding.name_of fname, fT) val domT = domain_type fT val ranT = range_type fT val default = Syntax.parse_term lthy default_str |> Type.constraint fT |> Syntax.check_term lthy val (globals, ctxt') = fix_globals domT ranT (Free fvar) lthy val Globals { x, h, ... } = globals val clauses = map (mk_clause_context x ctxt') abstract_qglrs val n = length abstract_qglrs fun build_tree (ClauseContext { ctxt, rhs, ...}) = Function_Context_Tree.mk_tree (Free fvar) h ctxt rhs val trees = map build_tree clauses val RCss = map find_calls trees val ((G, GIntro_thms, G_elim, G_induct), lthy) = PROFILE "def_graph" (define_graph (derived_name_suffix defname "_graph", domT --> ranT --> boolT) (Free fvar) clauses RCss) lthy val ((f, (_, f_defthm)), lthy) = PROFILE "def_fun" (define_function defname (fname, mixfix) domT ranT G default) lthy val RCss = map (map (inst_RC lthy (Free fvar) f)) RCss val trees = map (Function_Context_Tree.inst_tree lthy (Free fvar) f) trees val ((R, RIntro_thmss, R_elim), lthy) = PROFILE "def_rel" (define_recursion_relation (derived_name_suffix defname "_rel", domT --> domT --> boolT) abstract_qglrs clauses RCss) lthy val dom = mk_acc domT R val (_, lthy) = Local_Theory.abbrev Syntax.mode_default ((derived_name_suffix defname "_dom", NoSyn), dom) lthy val newthy = Proof_Context.theory_of lthy val clauses = map (transfer_clause_ctx newthy) clauses val xclauses = PROFILE "xclauses" (@{map 7} (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms) RIntro_thmss val complete = mk_completeness globals clauses abstract_qglrs |> Thm.cterm_of lthy |> Thm.assume val compat = mk_compat_proof_obligations domT ranT (Free fvar) f abstract_qglrs |> map (Thm.cterm_of lthy #> Thm.assume) val compat_store = store_compat_thms n compat val (goalstate, values) = PROFILE "prove_stuff" (prove_stuff lthy globals G f R xclauses complete compat compat_store G_elim) f_defthm fun mk_partial_rules newctxt provedgoal = let val (graph_is_function, complete_thm) = provedgoal |> Conjunction.elim |> apfst (Thm.forall_elim_vars 0) val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff) val psimps = PROFILE "Proving simplification rules" (mk_psimps newctxt globals R xclauses values f_iff) graph_is_function val simple_pinduct = PROFILE "Proving partial induction rule" (mk_partial_induct_rule newctxt globals R complete_thm) xclauses val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newctxt globals R R_elim) xclauses val dom_intros = if domintros then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses) else NONE in FunctionResult {fs=[f], G=G, R=R, dom=dom, cases=[complete_thm], psimps=psimps, pelims=[], simple_pinducts=[simple_pinduct], termination=total_intro, domintros=dom_intros} end in ((f, goalstate, mk_partial_rules), lthy) end end diff --git a/src/HOL/Tools/Meson/meson_clausify.ML b/src/HOL/Tools/Meson/meson_clausify.ML --- a/src/HOL/Tools/Meson/meson_clausify.ML +++ b/src/HOL/Tools/Meson/meson_clausify.ML @@ -1,388 +1,388 @@ (* Title: HOL/Tools/Meson/meson_clausify.ML Author: Jia Meng, Cambridge University Computer Laboratory and NICTA Author: Jasmin Blanchette, TU Muenchen Transformation of HOL theorems into CNF forms. *) signature MESON_CLAUSIFY = sig val new_skolem_var_prefix : string val new_nonskolem_var_prefix : string val is_zapped_var_name : string -> bool val is_quasi_lambda_free : term -> bool val introduce_combinators_in_cterm : Proof.context -> cterm -> thm val introduce_combinators_in_theorem : Proof.context -> thm -> thm val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool val ss_only : thm list -> Proof.context -> Proof.context val cnf_axiom : Meson.simp_options -> Proof.context -> bool -> bool -> int -> thm -> (thm * term) option * thm list end; structure Meson_Clausify : MESON_CLAUSIFY = struct (* the extra "Meson" helps prevent clashes (FIXME) *) val new_skolem_var_prefix = "MesonSK" val new_nonskolem_var_prefix = "MesonV" fun is_zapped_var_name s = exists (fn prefix => String.isPrefix prefix s) [new_skolem_var_prefix, new_nonskolem_var_prefix] (**** Transformation of Elimination Rules into First-Order Formulas****) val cfalse = Thm.cterm_of \<^theory_context>\HOL\ \<^term>\False\; val ctp_false = Thm.cterm_of \<^theory_context>\HOL\ (HOLogic.mk_Trueprop \<^term>\False\); (* Converts an elim-rule into an equivalent theorem that does not have the predicate variable. Leaves other theorems unchanged. We simply instantiate the conclusion variable to False. (Cf. "transform_elim_prop" in "Sledgehammer_Util".) *) fun transform_elim_theorem th = (case Thm.concl_of th of (*conclusion variable*) \<^const>\Trueprop\ $ (Var (v as (_, \<^typ>\bool\))) => Thm.instantiate ([], [(v, cfalse)]) th | Var (v as (_, \<^typ>\prop\)) => Thm.instantiate ([], [(v, ctp_false)]) th | _ => th) (**** SKOLEMIZATION BY INFERENCE (lcp) ****) fun mk_old_skolem_term_wrapper t = let val T = fastype_of t in Const (\<^const_name>\Meson.skolem\, T --> T) $ t end fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') | beta_eta_in_abs_body t = Envir.beta_eta_contract t (*Traverse a theorem, accumulating Skolem function definitions.*) fun old_skolem_defs th = let fun dec_sko (Const (\<^const_name>\Ex\, _) $ (body as Abs (_, T, p))) rhss = (*Existential: declare a Skolem function, then insert into body and continue*) let val args = Misc_Legacy.term_frees body (* Forms a lambda-abstraction over the formal parameters *) val rhs = fold_rev (absfree o dest_Free) args (HOLogic.choice_const T $ beta_eta_in_abs_body body) |> mk_old_skolem_term_wrapper val comb = list_comb (rhs, args) in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end | dec_sko (Const (\<^const_name>\All\,_) $ Abs (a, T, p)) rhss = (*Universal quant: insert a free variable into body and continue*) let val fname = singleton (Name.variant_list (Misc_Legacy.add_term_names (p, []))) a in dec_sko (subst_bound (Free(fname,T), p)) rhss end | dec_sko (\<^const>\conj\ $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q | dec_sko (\<^const>\disj\ $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q | dec_sko (\<^const>\Trueprop\ $ p) rhss = dec_sko p rhss | dec_sko _ rhss = rhss in dec_sko (Thm.prop_of th) [] end; (**** REPLACING ABSTRACTIONS BY COMBINATORS ****) fun is_quasi_lambda_free (Const (\<^const_name>\Meson.skolem\, _) $ _) = true | is_quasi_lambda_free (t1 $ t2) = is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2 | is_quasi_lambda_free (Abs _) = false | is_quasi_lambda_free _ = true fun abstract ctxt ct = let val Abs (_, _, body) = Thm.term_of ct val (x, cbody) = Thm.dest_abs NONE ct val (A, cbodyT) = Thm.dest_funT (Thm.ctyp_of_cterm ct) fun makeK () = Thm.instantiate' [SOME A, SOME cbodyT] [SOME cbody] @{thm abs_K} in case body of Const _ => makeK() | Free _ => makeK() | Var _ => makeK() (*though Var isn't expected*) | Bound 0 => Thm.instantiate' [SOME A] [] @{thm abs_I} (*identity: I*) | rator$rand => if Term.is_dependent rator then (*C or S*) if Term.is_dependent rand then (*S*) let val crator = Thm.lambda x (Thm.dest_fun cbody) val crand = Thm.lambda x (Thm.dest_arg cbody) val (C, B) = Thm.dest_funT (Thm.dest_ctyp1 (Thm.ctyp_of_cterm crator)) val abs_S' = @{thm abs_S} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_S') in Thm.transitive abs_S' (Conv.binop_conv (abstract ctxt) rhs) end else (*C*) let val crator = Thm.lambda x (Thm.dest_fun cbody) val crand = Thm.dest_arg cbody val (C, B) = Thm.dest_funT (Thm.dest_ctyp1 (Thm.ctyp_of_cterm crator)) val abs_C' = @{thm abs_C} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_C') in Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv (abstract ctxt)) rhs) end else if Term.is_dependent rand then (*B or eta*) if rand = Bound 0 then Thm.eta_conversion ct else (*B*) let val crator = Thm.dest_fun cbody val crand = Thm.lambda x (Thm.dest_arg cbody) val (C, B) = Thm.dest_funT (Thm.ctyp_of_cterm crator) val abs_B' = @{thm abs_B} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_B') in Thm.transitive abs_B' (Conv.arg_conv (abstract ctxt) rhs) end else makeK () | _ => raise Fail "abstract: Bad term" end; (* Traverse a theorem, remplacing lambda-abstractions with combinators. *) fun introduce_combinators_in_cterm ctxt ct = if is_quasi_lambda_free (Thm.term_of ct) then Thm.reflexive ct else case Thm.term_of ct of Abs _ => let val (cv, cta) = Thm.dest_abs NONE ct val (v, _) = dest_Free (Thm.term_of cv) val u_th = introduce_combinators_in_cterm ctxt cta val cu = Thm.rhs_of u_th val comb_eq = abstract ctxt (Thm.lambda cv cu) in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end | _ $ _ => let val (ct1, ct2) = Thm.dest_comb ct in Thm.combination (introduce_combinators_in_cterm ctxt ct1) (introduce_combinators_in_cterm ctxt ct2) end fun introduce_combinators_in_theorem ctxt th = if is_quasi_lambda_free (Thm.prop_of th) then th else let val th = Drule.eta_contraction_rule th val eqth = introduce_combinators_in_cterm ctxt (Thm.cprop_of th) in Thm.equal_elim eqth th end handle THM (msg, _, _) => (warning ("Error in the combinator translation of " ^ Thm.string_of_thm ctxt th ^ "\nException message: " ^ msg); (* A type variable of sort "{}" will make "abstraction" fail. *) TrueI) (*cterms are used throughout for efficiency*) val cTrueprop = Thm.cterm_of \<^theory_context>\HOL\ HOLogic.Trueprop; (*Given an abstraction over n variables, replace the bound variables by free ones. Return the body, along with the list of free variables.*) fun c_variant_abs_multi (ct0, vars) = let val (cv,ct) = Thm.dest_abs NONE ct0 in c_variant_abs_multi (ct, cv::vars) end handle CTERM _ => (ct0, rev vars); (* Given the definition of a Skolem function, return a theorem to replace an existential formula by a use of that function. Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *) fun old_skolem_theorem_of_def ctxt rhs0 = let val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> Thm.cterm_of ctxt val rhs' = rhs |> Thm.dest_comb |> snd val (ch, frees) = c_variant_abs_multi (rhs', []) val (hilbert, cabs) = ch |> Thm.dest_comb |>> Thm.term_of val T = case hilbert of Const (_, Type (\<^type_name>\fun\, [_, T])) => T | _ => raise TERM ("old_skolem_theorem_of_def: expected \"Eps\"", [hilbert]) val cex = Thm.cterm_of ctxt (HOLogic.exists_const T) val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs) val conc = Drule.list_comb (rhs, frees) |> Drule.beta_conv cabs |> Thm.apply cTrueprop fun tacf [prem] = rewrite_goals_tac ctxt @{thms skolem_def [abs_def]} THEN resolve_tac ctxt [(prem |> rewrite_rule ctxt @{thms skolem_def [abs_def]}) RS Global_Theory.get_thm (Proof_Context.theory_of ctxt) "Hilbert_Choice.someI_ex"] 1 in Goal.prove_internal ctxt [ex_tm] conc tacf |> forall_intr_list frees |> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) |> Thm.varifyT_global end fun to_definitional_cnf_with_quantifiers ctxt th = let val eqth = CNF.make_cnfx_thm ctxt (HOLogic.dest_Trueprop (Thm.prop_of th)) val eqth = eqth RS @{thm eq_reflection} val eqth = eqth RS @{thm TruepropI} in Thm.equal_elim eqth th end fun zapped_var_name ((ax_no, cluster_no), skolem) index_no s = (if skolem then new_skolem_var_prefix else new_nonskolem_var_prefix) ^ "_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^ string_of_int index_no ^ "_" ^ Name.desymbolize (SOME false) s fun cluster_of_zapped_var_name s = let val get_int = the o Int.fromString o nth (space_explode "_" s) in ((get_int 1, (get_int 2, get_int 3)), String.isPrefix new_skolem_var_prefix s) end fun rename_bound_vars_to_be_zapped ax_no = let fun aux (cluster as (cluster_no, cluster_skolem)) index_no pos t = case t of (t1 as Const (s, _)) $ Abs (s', T, t') => if s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse s = \<^const_name>\Ex\ then let val skolem = (pos = (s = \<^const_name>\Ex\)) val (cluster, index_no) = if skolem = cluster_skolem then (cluster, index_no) else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0) val s' = zapped_var_name cluster index_no s' in t1 $ Abs (s', T, aux cluster (index_no + 1) pos t') end else t | (t1 as Const (s, _)) $ t2 $ t3 => if s = \<^const_name>\Pure.imp\ orelse s = \<^const_name>\HOL.implies\ then t1 $ aux cluster index_no (not pos) t2 $ aux cluster index_no pos t3 else if s = \<^const_name>\HOL.conj\ orelse s = \<^const_name>\HOL.disj\ then t1 $ aux cluster index_no pos t2 $ aux cluster index_no pos t3 else t | (t1 as Const (s, _)) $ t2 => if s = \<^const_name>\Trueprop\ then t1 $ aux cluster index_no pos t2 else if s = \<^const_name>\Not\ then t1 $ aux cluster index_no (not pos) t2 else t | _ => t in aux ((ax_no, 0), true) 0 true end fun zap pos ct = ct |> (case Thm.term_of ct of Const (s, _) $ Abs (s', _, _) => if s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse s = \<^const_name>\Ex\ then Thm.dest_comb #> snd #> Thm.dest_abs (SOME s') #> snd #> zap pos else Conv.all_conv | Const (s, _) $ _ $ _ => if s = \<^const_name>\Pure.imp\ orelse s = \<^const_name>\implies\ then Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos) else if s = \<^const_name>\conj\ orelse s = \<^const_name>\disj\ then Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos) else Conv.all_conv | Const (s, _) $ _ => if s = \<^const_name>\Trueprop\ then Conv.arg_conv (zap pos) else if s = \<^const_name>\Not\ then Conv.arg_conv (zap (not pos)) else Conv.all_conv | _ => Conv.all_conv) fun ss_only ths ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths val cheat_choice = \<^prop>\\x. \y. Q x y \ \f. \x. Q x (f x)\ |> Logic.varify_global |> Skip_Proof.make_thm \<^theory> (* Converts an Isabelle theorem into NNF. *) fun nnf_axiom simp_options choice_ths new_skolem ax_no th ctxt = let val thy = Proof_Context.theory_of ctxt val th = th |> transform_elim_theorem |> zero_var_indexes - |> new_skolem ? forall_intr_vars + |> new_skolem ? Thm.forall_intr_vars val (th, ctxt) = Variable.import true [th] ctxt |>> snd |>> the_single val th = th |> Conv.fconv_rule (Object_Logic.atomize ctxt) |> Meson.cong_extensionalize_thm ctxt |> Meson.abs_extensionalize_thm ctxt |> Meson.make_nnf simp_options ctxt in if new_skolem then let fun skolemize choice_ths = Meson.skolemize_with_choice_theorems simp_options ctxt choice_ths #> simplify (ss_only @{thms all_simps[symmetric]} ctxt) val no_choice = null choice_ths val pull_out = if no_choice then simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]} ctxt) else skolemize choice_ths val discharger_th = th |> pull_out val discharger_th = discharger_th |> Meson.has_too_many_clauses ctxt (Thm.concl_of discharger_th) ? (to_definitional_cnf_with_quantifiers ctxt #> pull_out) val zapped_th = discharger_th |> Thm.prop_of |> rename_bound_vars_to_be_zapped ax_no |> (if no_choice then Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> Thm.cprop_of else Thm.cterm_of ctxt) |> zap true val fixes = [] |> Term.add_free_names (Thm.prop_of zapped_th) |> filter is_zapped_var_name val ctxt' = ctxt |> Variable.add_fixes_direct fixes val fully_skolemized_t = zapped_th |> singleton (Variable.export ctxt' ctxt) |> Thm.cprop_of |> Thm.dest_equals |> snd |> Thm.term_of in if exists_subterm (fn Var ((s, _), _) => String.isPrefix new_skolem_var_prefix s | _ => false) fully_skolemized_t then let val (fully_skolemized_ct, ctxt) = yield_singleton (Variable.import_terms true) fully_skolemized_t ctxt |>> Thm.cterm_of ctxt in (SOME (discharger_th, fully_skolemized_ct), (Thm.assume fully_skolemized_ct, ctxt)) end else (NONE, (th, ctxt)) end else (NONE, (th |> Meson.has_too_many_clauses ctxt (Thm.concl_of th) ? to_definitional_cnf_with_quantifiers ctxt, ctxt)) end (* Convert a theorem to CNF, with additional premises due to skolemization. *) fun cnf_axiom simp_options ctxt0 new_skolem combinators ax_no th = let val choice_ths = Meson.choice_theorems (Proof_Context.theory_of ctxt0) val (opt, (nnf_th, ctxt1)) = nnf_axiom simp_options choice_ths new_skolem ax_no th ctxt0 fun clausify th = Meson.make_cnf (if new_skolem orelse null choice_ths then [] else map (old_skolem_theorem_of_def ctxt1) (old_skolem_defs th)) th ctxt1 val (cnf_ths, ctxt2) = clausify nnf_th fun intr_imp ct th = Thm.instantiate ([], [((("i", 0), \<^typ>\nat\), Thm.cterm_of ctxt2 (HOLogic.mk_nat ax_no))]) (zero_var_indexes @{thm skolem_COMBK_D}) RS Thm.implies_intr ct th in (opt |> Option.map (I #>> singleton (Variable.export ctxt2 ctxt0) ##> (Thm.term_of #> HOLogic.dest_Trueprop #> singleton (Variable.export_terms ctxt2 ctxt0))), cnf_ths |> map (combinators ? introduce_combinators_in_theorem ctxt2 #> (case opt of SOME (_, ct) => intr_imp ct | NONE => I)) |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf |> map (Thm.close_derivation \<^here>)) end handle THM _ => (NONE, []) end; diff --git a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML @@ -1,1225 +1,1227 @@ (* Title: HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Author: Lukas Bulwahn, TU Muenchen Auxilary functions for predicate compiler. *) signature PREDICATE_COMPILE_AUX = sig val find_indices : ('a -> bool) -> 'a list -> int list (* mode *) datatype mode = Bool | Input | Output | Pair of mode * mode | Fun of mode * mode val eq_mode : mode * mode -> bool val mode_ord: mode ord val list_fun_mode : mode list -> mode val strip_fun_mode : mode -> mode list val dest_fun_mode : mode -> mode list val dest_tuple_mode : mode -> mode list val all_modes_of_typ : typ -> mode list val all_smodes_of_typ : typ -> mode list val fold_map_aterms_prodT : ('a -> 'a -> 'a) -> (typ -> 'b -> 'a * 'b) -> typ -> 'b -> 'a * 'b val map_filter_prod : (term -> term option) -> term -> term option val replace_ho_args : mode -> term list -> term list -> term list val ho_arg_modes_of : mode -> mode list val ho_argsT_of : mode -> typ list -> typ list val ho_args_of : mode -> term list -> term list val ho_args_of_typ : typ -> term list -> term list val ho_argsT_of_typ : typ list -> typ list val split_map_mode : (mode -> term -> term option * term option) -> mode -> term list -> term list * term list val split_map_modeT : (mode -> typ -> typ option * typ option) -> mode -> typ list -> typ list * typ list val split_mode : mode -> term list -> term list * term list val split_modeT : mode -> typ list -> typ list * typ list val string_of_mode : mode -> string val ascii_string_of_mode : mode -> string (* premises *) datatype indprem = Prem of term | Negprem of term | Sidecond of term | Generator of (string * typ) val dest_indprem : indprem -> term val map_indprem : (term -> term) -> indprem -> indprem (* general syntactic functions *) val is_equationlike : thm -> bool val is_pred_equation : thm -> bool val is_intro : string -> thm -> bool val is_predT : typ -> bool val lookup_constr : Proof.context -> (string * typ) -> int option val is_constrt : Proof.context -> term -> bool val strip_ex : term -> (string * typ) list * term val focus_ex : term -> Name.context -> ((string * typ) list * term) * Name.context val strip_all : term -> (string * typ) list * term val strip_intro_concl : thm -> term * term list (* introduction rule combinators *) val map_atoms : (term -> term) -> term -> term val fold_atoms : (term -> 'a -> 'a) -> term -> 'a -> 'a val fold_map_atoms : (term -> 'a -> term * 'a) -> term -> 'a -> term * 'a val maps_premises : (term -> term list) -> term -> term val map_concl : (term -> term) -> term -> term val map_term : theory -> (term -> term) -> thm -> thm (* split theorems of case expressions *) val prepare_split_thm : Proof.context -> thm -> thm val find_split_thm : theory -> term -> thm option (* datastructures and setup for generic compilation *) datatype compilation_funs = CompilationFuns of { mk_monadT : typ -> typ, dest_monadT : typ -> typ, mk_empty : typ -> term, mk_single : term -> term, mk_bind : term * term -> term, mk_plus : term * term -> term, mk_if : term -> term, mk_iterate_upto : typ -> term * term * term -> term, mk_not : term -> term, mk_map : typ -> typ -> term -> term -> term }; val mk_monadT : compilation_funs -> typ -> typ val dest_monadT : compilation_funs -> typ -> typ val mk_empty : compilation_funs -> typ -> term val mk_single : compilation_funs -> term -> term val mk_bind : compilation_funs -> term * term -> term val mk_plus : compilation_funs -> term * term -> term val mk_if : compilation_funs -> term -> term val mk_iterate_upto : compilation_funs -> typ -> term * term * term -> term val mk_not : compilation_funs -> term -> term val mk_map : compilation_funs -> typ -> typ -> term -> term -> term val funT_of : compilation_funs -> mode -> typ -> typ (* Different compilations *) datatype compilation = Pred | Depth_Limited | Random | Depth_Limited_Random | DSeq | Annotated | Pos_Random_DSeq | Neg_Random_DSeq | New_Pos_Random_DSeq | New_Neg_Random_DSeq | Pos_Generator_DSeq | Neg_Generator_DSeq | Pos_Generator_CPS | Neg_Generator_CPS val negative_compilation_of : compilation -> compilation val compilation_for_polarity : bool -> compilation -> compilation val is_depth_limited_compilation : compilation -> bool val string_of_compilation : compilation -> string val compilation_names : (string * compilation) list val non_random_compilations : compilation list val random_compilations : compilation list (* Different options for compiler *) datatype options = Options of { expected_modes : (string * mode list) option, proposed_modes : (string * mode list) list, proposed_names : ((string * mode) * string) list, show_steps : bool, show_proof_trace : bool, show_intermediate_results : bool, show_mode_inference : bool, show_modes : bool, show_compilation : bool, show_caught_failures : bool, show_invalid_clauses : bool, skip_proof : bool, no_topmost_reordering : bool, function_flattening : bool, fail_safe_function_flattening : bool, specialise : bool, no_higher_order_predicate : string list, inductify : bool, detect_switches : bool, smart_depth_limiting : bool, compilation : compilation }; val expected_modes : options -> (string * mode list) option val proposed_modes : options -> string -> mode list option val proposed_names : options -> string -> mode -> string option val show_steps : options -> bool val show_proof_trace : options -> bool val show_intermediate_results : options -> bool val show_mode_inference : options -> bool val show_modes : options -> bool val show_compilation : options -> bool val show_caught_failures : options -> bool val show_invalid_clauses : options -> bool val skip_proof : options -> bool val no_topmost_reordering : options -> bool val function_flattening : options -> bool val fail_safe_function_flattening : options -> bool val specialise : options -> bool val no_higher_order_predicate : options -> string list val is_inductify : options -> bool val detect_switches : options -> bool val smart_depth_limiting : options -> bool val compilation : options -> compilation val default_options : options val bool_options : string list val print_step : options -> string -> unit (* conversions *) val imp_prems_conv : conv -> conv (* simple transformations *) val split_conjuncts_in_assms : Proof.context -> thm -> thm val dest_conjunct_prem : thm -> thm list val expand_tuples : theory -> thm -> thm val case_betapply : theory -> term -> term val eta_contract_ho_arguments : theory -> thm -> thm val remove_equalities : theory -> thm -> thm val remove_pointless_clauses : thm -> thm list val peephole_optimisation : theory -> thm -> thm option (* auxillary *) val unify_consts : theory -> term list -> term list -> (term list * term list) val mk_casesrule : Proof.context -> term -> thm list -> term val preprocess_intro : theory -> thm -> thm val define_quickcheck_predicate : term -> theory -> (((string * typ) * (string * typ) list) * thm) * theory end structure Predicate_Compile_Aux : PREDICATE_COMPILE_AUX = struct (* general functions *) fun comb_option f (SOME x1, SOME x2) = SOME (f (x1, x2)) | comb_option f (NONE, SOME x2) = SOME x2 | comb_option f (SOME x1, NONE) = SOME x1 | comb_option f (NONE, NONE) = NONE fun map2_optional f (x :: xs) (y :: ys) = f x (SOME y) :: (map2_optional f xs ys) | map2_optional f (x :: xs) [] = (f x NONE) :: (map2_optional f xs []) | map2_optional f [] [] = [] fun find_indices f xs = map_filter (fn (i, true) => SOME i | (_, false) => NONE) (map_index (apsnd f) xs) (* mode *) datatype mode = Bool | Input | Output | Pair of mode * mode | Fun of mode * mode (* equality of instantiatedness with respect to equivalences: Pair Input Input == Input and Pair Output Output == Output *) fun eq_mode (Fun (m1, m2), Fun (m3, m4)) = eq_mode (m1, m3) andalso eq_mode (m2, m4) | eq_mode (Pair (m1, m2), Pair (m3, m4)) = eq_mode (m1, m3) andalso eq_mode (m2, m4) | eq_mode (Pair (m1, m2), Input) = eq_mode (m1, Input) andalso eq_mode (m2, Input) | eq_mode (Pair (m1, m2), Output) = eq_mode (m1, Output) andalso eq_mode (m2, Output) | eq_mode (Input, Pair (m1, m2)) = eq_mode (Input, m1) andalso eq_mode (Input, m2) | eq_mode (Output, Pair (m1, m2)) = eq_mode (Output, m1) andalso eq_mode (Output, m2) | eq_mode (Input, Input) = true | eq_mode (Output, Output) = true | eq_mode (Bool, Bool) = true | eq_mode _ = false fun mode_ord (Input, Output) = LESS | mode_ord (Output, Input) = GREATER | mode_ord (Input, Input) = EQUAL | mode_ord (Output, Output) = EQUAL | mode_ord (Bool, Bool) = EQUAL | mode_ord (Pair (m1, m2), Pair (m3, m4)) = prod_ord mode_ord mode_ord ((m1, m2), (m3, m4)) | mode_ord (Fun (m1, m2), Fun (m3, m4)) = prod_ord mode_ord mode_ord ((m1, m2), (m3, m4)) fun list_fun_mode [] = Bool | list_fun_mode (m :: ms) = Fun (m, list_fun_mode ms) (* name: binder_modes? *) fun strip_fun_mode (Fun (mode, mode')) = mode :: strip_fun_mode mode' | strip_fun_mode Bool = [] | strip_fun_mode _ = raise Fail "Bad mode for strip_fun_mode" (* name: strip_fun_mode? *) fun dest_fun_mode (Fun (mode, mode')) = mode :: dest_fun_mode mode' | dest_fun_mode mode = [mode] fun dest_tuple_mode (Pair (mode, mode')) = mode :: dest_tuple_mode mode' | dest_tuple_mode _ = [] fun all_modes_of_typ' (T as Type ("fun", _)) = let val (S, U) = strip_type T in if U = HOLogic.boolT then fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2) (map all_modes_of_typ' S) [Bool] else [Input, Output] end | all_modes_of_typ' (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = map_product (curry Pair) (all_modes_of_typ' T1) (all_modes_of_typ' T2) | all_modes_of_typ' _ = [Input, Output] fun all_modes_of_typ (T as Type ("fun", _)) = let val (S, U) = strip_type T in if U = \<^typ>\bool\ then fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2) (map all_modes_of_typ' S) [Bool] else raise Fail "Invocation of all_modes_of_typ with a non-predicate type" end | all_modes_of_typ \<^typ>\bool\ = [Bool] | all_modes_of_typ _ = raise Fail "Invocation of all_modes_of_typ with a non-predicate type" fun all_smodes_of_typ (T as Type ("fun", _)) = let val (S, U) = strip_type T fun all_smodes (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = map_product (curry Pair) (all_smodes T1) (all_smodes T2) | all_smodes _ = [Input, Output] in if U = HOLogic.boolT then fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2) (map all_smodes S) [Bool] else raise Fail "invalid type for predicate" end fun ho_arg_modes_of mode = let fun ho_arg_mode (m as Fun _) = [m] | ho_arg_mode (Pair (m1, m2)) = ho_arg_mode m1 @ ho_arg_mode m2 | ho_arg_mode _ = [] in maps ho_arg_mode (strip_fun_mode mode) end fun ho_args_of mode ts = let fun ho_arg (Fun _) (SOME t) = [t] | ho_arg (Fun _) NONE = raise Fail "mode and term do not match" | ho_arg (Pair (m1, m2)) (SOME (Const (\<^const_name>\Pair\, _) $ t1 $ t2)) = ho_arg m1 (SOME t1) @ ho_arg m2 (SOME t2) | ho_arg (Pair (m1, m2)) NONE = ho_arg m1 NONE @ ho_arg m2 NONE | ho_arg _ _ = [] in flat (map2_optional ho_arg (strip_fun_mode mode) ts) end fun ho_args_of_typ T ts = let fun ho_arg (T as Type ("fun", [_, _])) (SOME t) = if body_type T = \<^typ>\bool\ then [t] else [] | ho_arg (Type ("fun", [_, _])) NONE = raise Fail "mode and term do not match" | ho_arg (Type(\<^type_name>\Product_Type.prod\, [T1, T2])) (SOME (Const (\<^const_name>\Pair\, _) $ t1 $ t2)) = ho_arg T1 (SOME t1) @ ho_arg T2 (SOME t2) | ho_arg (Type(\<^type_name>\Product_Type.prod\, [T1, T2])) NONE = ho_arg T1 NONE @ ho_arg T2 NONE | ho_arg _ _ = [] in flat (map2_optional ho_arg (binder_types T) ts) end fun ho_argsT_of_typ Ts = let fun ho_arg (T as Type("fun", [_,_])) = if body_type T = \<^typ>\bool\ then [T] else [] | ho_arg (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = ho_arg T1 @ ho_arg T2 | ho_arg _ = [] in maps ho_arg Ts end (* temporary function should be replaced by unsplit_input or so? *) fun replace_ho_args mode hoargs ts = let fun replace (Fun _, _) (arg' :: hoargs') = (arg', hoargs') | replace (Pair (m1, m2), Const (\<^const_name>\Pair\, T) $ t1 $ t2) hoargs = let val (t1', hoargs') = replace (m1, t1) hoargs val (t2', hoargs'') = replace (m2, t2) hoargs' in (Const (\<^const_name>\Pair\, T) $ t1' $ t2', hoargs'') end | replace (_, t) hoargs = (t, hoargs) in fst (fold_map replace (strip_fun_mode mode ~~ ts) hoargs) end fun ho_argsT_of mode Ts = let fun ho_arg (Fun _) T = [T] | ho_arg (Pair (m1, m2)) (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = ho_arg m1 T1 @ ho_arg m2 T2 | ho_arg _ _ = [] in flat (map2 ho_arg (strip_fun_mode mode) Ts) end (* splits mode and maps function to higher-order argument types *) fun split_map_mode f mode ts = let fun split_arg_mode' (m as Fun _) t = f m t | split_arg_mode' (Pair (m1, m2)) (Const (\<^const_name>\Pair\, _) $ t1 $ t2) = let val (i1, o1) = split_arg_mode' m1 t1 val (i2, o2) = split_arg_mode' m2 t2 in (comb_option HOLogic.mk_prod (i1, i2), comb_option HOLogic.mk_prod (o1, o2)) end | split_arg_mode' m t = if eq_mode (m, Input) then (SOME t, NONE) else if eq_mode (m, Output) then (NONE, SOME t) else raise Fail "split_map_mode: mode and term do not match" in (apply2 (map_filter I) o split_list) (map2 split_arg_mode' (strip_fun_mode mode) ts) end (* splits mode and maps function to higher-order argument types *) fun split_map_modeT f mode Ts = let fun split_arg_mode' (m as Fun _) T = f m T | split_arg_mode' (Pair (m1, m2)) (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = let val (i1, o1) = split_arg_mode' m1 T1 val (i2, o2) = split_arg_mode' m2 T2 in (comb_option HOLogic.mk_prodT (i1, i2), comb_option HOLogic.mk_prodT (o1, o2)) end | split_arg_mode' Input T = (SOME T, NONE) | split_arg_mode' Output T = (NONE, SOME T) | split_arg_mode' _ _ = raise Fail "split_modeT': mode and type do not match" in (apply2 (map_filter I) o split_list) (map2 split_arg_mode' (strip_fun_mode mode) Ts) end fun split_mode mode ts = split_map_mode (fn _ => fn _ => (NONE, NONE)) mode ts fun fold_map_aterms_prodT comb f (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) s = let val (x1, s') = fold_map_aterms_prodT comb f T1 s val (x2, s'') = fold_map_aterms_prodT comb f T2 s' in (comb x1 x2, s'') end | fold_map_aterms_prodT _ f T s = f T s fun map_filter_prod f (Const (\<^const_name>\Pair\, _) $ t1 $ t2) = comb_option HOLogic.mk_prod (map_filter_prod f t1, map_filter_prod f t2) | map_filter_prod f t = f t fun split_modeT mode Ts = let fun split_arg_mode (Fun _) _ = ([], []) | split_arg_mode (Pair (m1, m2)) (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = let val (i1, o1) = split_arg_mode m1 T1 val (i2, o2) = split_arg_mode m2 T2 in (i1 @ i2, o1 @ o2) end | split_arg_mode Input T = ([T], []) | split_arg_mode Output T = ([], [T]) | split_arg_mode _ _ = raise Fail "split_modeT: mode and type do not match" in (apply2 flat o split_list) (map2 split_arg_mode (strip_fun_mode mode) Ts) end fun string_of_mode mode = let fun string_of_mode1 Input = "i" | string_of_mode1 Output = "o" | string_of_mode1 Bool = "bool" | string_of_mode1 mode = "(" ^ (string_of_mode3 mode) ^ ")" and string_of_mode2 (Pair (m1, m2)) = string_of_mode3 m1 ^ " * " ^ string_of_mode2 m2 | string_of_mode2 mode = string_of_mode1 mode and string_of_mode3 (Fun (m1, m2)) = string_of_mode2 m1 ^ " => " ^ string_of_mode3 m2 | string_of_mode3 mode = string_of_mode2 mode in string_of_mode3 mode end fun ascii_string_of_mode mode' = let fun ascii_string_of_mode' Input = "i" | ascii_string_of_mode' Output = "o" | ascii_string_of_mode' Bool = "b" | ascii_string_of_mode' (Pair (m1, m2)) = "P" ^ ascii_string_of_mode' m1 ^ ascii_string_of_mode'_Pair m2 | ascii_string_of_mode' (Fun (m1, m2)) = "F" ^ ascii_string_of_mode' m1 ^ ascii_string_of_mode'_Fun m2 ^ "B" and ascii_string_of_mode'_Fun (Fun (m1, m2)) = ascii_string_of_mode' m1 ^ (if m2 = Bool then "" else "_" ^ ascii_string_of_mode'_Fun m2) | ascii_string_of_mode'_Fun Bool = "B" | ascii_string_of_mode'_Fun m = ascii_string_of_mode' m and ascii_string_of_mode'_Pair (Pair (m1, m2)) = ascii_string_of_mode' m1 ^ ascii_string_of_mode'_Pair m2 | ascii_string_of_mode'_Pair m = ascii_string_of_mode' m in ascii_string_of_mode'_Fun mode' end (* premises *) datatype indprem = Prem of term | Negprem of term | Sidecond of term | Generator of (string * typ) fun dest_indprem (Prem t) = t | dest_indprem (Negprem t) = t | dest_indprem (Sidecond t) = t | dest_indprem (Generator _) = raise Fail "cannot destruct generator" fun map_indprem f (Prem t) = Prem (f t) | map_indprem f (Negprem t) = Negprem (f t) | map_indprem f (Sidecond t) = Sidecond (f t) | map_indprem f (Generator (v, T)) = Generator (dest_Free (f (Free (v, T)))) (* general syntactic functions *) fun is_equationlike_term (Const (\<^const_name>\Pure.eq\, _) $ _ $ _) = true | is_equationlike_term (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _)) = true | is_equationlike_term _ = false val is_equationlike = is_equationlike_term o Thm.prop_of fun is_pred_equation_term (Const (\<^const_name>\Pure.eq\, _) $ u $ v) = (fastype_of u = \<^typ>\bool\) andalso (fastype_of v = \<^typ>\bool\) | is_pred_equation_term _ = false val is_pred_equation = is_pred_equation_term o Thm.prop_of fun is_intro_term constname t = the_default false (try (fn t => case fst (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl t))) of Const (c, _) => c = constname | _ => false) t) fun is_intro constname t = is_intro_term constname (Thm.prop_of t) fun is_predT (T as Type("fun", [_, _])) = (body_type T = \<^typ>\bool\) | is_predT _ = false fun lookup_constr ctxt = let val tab = Ctr_Sugar.ctr_sugars_of ctxt |> maps (map_filter (try dest_Const) o #ctrs) |> map (fn (c, T) => ((c, (fst o dest_Type o body_type) T), BNF_Util.num_binder_types T)) in fn (c, T) => case body_type T of Type (Tname, _) => AList.lookup (op =) tab (c, Tname) | _ => NONE end; fun is_constrt ctxt = let val lookup_constr = lookup_constr ctxt fun check t = (case strip_comb t of (Var _, []) => true | (Free _, []) => true | (Const cT, ts) => (case lookup_constr cT of SOME i => length ts = i andalso forall check ts | _ => false) | _ => false) in check end fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t) fun strip_ex (Const (\<^const_name>\Ex\, _) $ Abs (x, T, t)) = let val (xTs, t') = strip_ex t in ((x, T) :: xTs, t') end | strip_ex t = ([], t) fun focus_ex t nctxt = let val ((xs, Ts), t') = apfst split_list (strip_ex t) val (xs', nctxt') = fold_map Name.variant xs nctxt; val ps' = xs' ~~ Ts; val vs = map Free ps'; val t'' = Term.subst_bounds (rev vs, t'); in ((ps', t''), nctxt') end val strip_intro_concl = strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o Thm.prop_of (* introduction rule combinators *) fun map_atoms f intro = let val (literals, head) = Logic.strip_horn intro fun appl t = (case t of (\<^term>\Not\ $ t') => HOLogic.mk_not (f t') | _ => f t) in Logic.list_implies (map (HOLogic.mk_Trueprop o appl o HOLogic.dest_Trueprop) literals, head) end fun fold_atoms f intro s = let val (literals, _) = Logic.strip_horn intro fun appl t s = (case t of (\<^term>\Not\ $ t') => f t' s | _ => f t s) in fold appl (map HOLogic.dest_Trueprop literals) s end fun fold_map_atoms f intro s = let val (literals, head) = Logic.strip_horn intro fun appl t s = (case t of (\<^term>\Not\ $ t') => apfst HOLogic.mk_not (f t' s) | _ => f t s) val (literals', s') = fold_map appl (map HOLogic.dest_Trueprop literals) s in (Logic.list_implies (map HOLogic.mk_Trueprop literals', head), s') end; fun map_filter_premises f intro = let val (premises, head) = Logic.strip_horn intro in Logic.list_implies (map_filter f premises, head) end fun maps_premises f intro = let val (premises, head) = Logic.strip_horn intro in Logic.list_implies (maps f premises, head) end fun map_concl f intro = let val (premises, head) = Logic.strip_horn intro in Logic.list_implies (premises, f head) end (* combinators to apply a function to all basic parts of nested products *) fun map_products f (Const (\<^const_name>\Pair\, T) $ t1 $ t2) = Const (\<^const_name>\Pair\, T) $ map_products f t1 $ map_products f t2 | map_products f t = f t (* split theorems of case expressions *) fun prepare_split_thm ctxt split_thm = (split_thm RS @{thm iffD2}) |> Local_Defs.unfold0 ctxt [@{thm atomize_conjL[symmetric]}, @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}] fun find_split_thm thy (Const (name, _)) = Option.map #split (Ctr_Sugar.ctr_sugar_of_case (Proof_Context.init_global thy) name) | find_split_thm _ _ = NONE (* lifting term operations to theorems *) fun map_term thy f th = Skip_Proof.make_thm thy (f (Thm.prop_of th)) (* fun equals_conv lhs_cv rhs_cv ct = case Thm.term_of ct of Const (@{const_name Pure.eq}, _) $ _ $ _ => Conv.arg_conv cv ct | _ => error "equals_conv" *) (* Different compilations *) datatype compilation = Pred | Depth_Limited | Random | Depth_Limited_Random | DSeq | Annotated | Pos_Random_DSeq | Neg_Random_DSeq | New_Pos_Random_DSeq | New_Neg_Random_DSeq | Pos_Generator_DSeq | Neg_Generator_DSeq | Pos_Generator_CPS | Neg_Generator_CPS fun negative_compilation_of Pos_Random_DSeq = Neg_Random_DSeq | negative_compilation_of Neg_Random_DSeq = Pos_Random_DSeq | negative_compilation_of New_Pos_Random_DSeq = New_Neg_Random_DSeq | negative_compilation_of New_Neg_Random_DSeq = New_Pos_Random_DSeq | negative_compilation_of Pos_Generator_DSeq = Neg_Generator_DSeq | negative_compilation_of Neg_Generator_DSeq = Pos_Generator_DSeq | negative_compilation_of Pos_Generator_CPS = Neg_Generator_CPS | negative_compilation_of Neg_Generator_CPS = Pos_Generator_CPS | negative_compilation_of c = c fun compilation_for_polarity false Pos_Random_DSeq = Neg_Random_DSeq | compilation_for_polarity false New_Pos_Random_DSeq = New_Neg_Random_DSeq | compilation_for_polarity _ c = c fun is_depth_limited_compilation c = (c = New_Pos_Random_DSeq) orelse (c = New_Neg_Random_DSeq) orelse (c = Pos_Generator_DSeq) orelse (c = Pos_Generator_DSeq) fun string_of_compilation c = (case c of Pred => "" | Random => "random" | Depth_Limited => "depth limited" | Depth_Limited_Random => "depth limited random" | DSeq => "dseq" | Annotated => "annotated" | Pos_Random_DSeq => "pos_random dseq" | Neg_Random_DSeq => "neg_random_dseq" | New_Pos_Random_DSeq => "new_pos_random dseq" | New_Neg_Random_DSeq => "new_neg_random_dseq" | Pos_Generator_DSeq => "pos_generator_dseq" | Neg_Generator_DSeq => "neg_generator_dseq" | Pos_Generator_CPS => "pos_generator_cps" | Neg_Generator_CPS => "neg_generator_cps") val compilation_names = [("pred", Pred), ("random", Random), ("depth_limited", Depth_Limited), ("depth_limited_random", Depth_Limited_Random), (*("annotated", Annotated),*) ("dseq", DSeq), ("random_dseq", Pos_Random_DSeq), ("new_random_dseq", New_Pos_Random_DSeq), ("generator_dseq", Pos_Generator_DSeq), ("generator_cps", Pos_Generator_CPS)] val non_random_compilations = [Pred, Depth_Limited, DSeq, Annotated] val random_compilations = [Random, Depth_Limited_Random, Pos_Random_DSeq, Neg_Random_DSeq, New_Pos_Random_DSeq, New_Neg_Random_DSeq, Pos_Generator_CPS, Neg_Generator_CPS] (* datastructures and setup for generic compilation *) datatype compilation_funs = CompilationFuns of { mk_monadT : typ -> typ, dest_monadT : typ -> typ, mk_empty : typ -> term, mk_single : term -> term, mk_bind : term * term -> term, mk_plus : term * term -> term, mk_if : term -> term, mk_iterate_upto : typ -> term * term * term -> term, mk_not : term -> term, mk_map : typ -> typ -> term -> term -> term } fun mk_monadT (CompilationFuns funs) = #mk_monadT funs fun dest_monadT (CompilationFuns funs) = #dest_monadT funs fun mk_empty (CompilationFuns funs) = #mk_empty funs fun mk_single (CompilationFuns funs) = #mk_single funs fun mk_bind (CompilationFuns funs) = #mk_bind funs fun mk_plus (CompilationFuns funs) = #mk_plus funs fun mk_if (CompilationFuns funs) = #mk_if funs fun mk_iterate_upto (CompilationFuns funs) = #mk_iterate_upto funs fun mk_not (CompilationFuns funs) = #mk_not funs fun mk_map (CompilationFuns funs) = #mk_map funs (** function types and names of different compilations **) fun funT_of compfuns mode T = let val Ts = binder_types T val (inTs, outTs) = split_map_modeT (fn m => fn T => (SOME (funT_of compfuns m T), NONE)) mode Ts in inTs ---> (mk_monadT compfuns (HOLogic.mk_tupleT outTs)) end (* Different options for compiler *) datatype options = Options of { expected_modes : (string * mode list) option, proposed_modes : (string * mode list) list, proposed_names : ((string * mode) * string) list, show_steps : bool, show_proof_trace : bool, show_intermediate_results : bool, show_mode_inference : bool, show_modes : bool, show_compilation : bool, show_caught_failures : bool, show_invalid_clauses : bool, skip_proof : bool, no_topmost_reordering : bool, function_flattening : bool, specialise : bool, fail_safe_function_flattening : bool, no_higher_order_predicate : string list, inductify : bool, detect_switches : bool, smart_depth_limiting : bool, compilation : compilation } fun expected_modes (Options opt) = #expected_modes opt fun proposed_modes (Options opt) = AList.lookup (op =) (#proposed_modes opt) fun proposed_names (Options opt) name mode = AList.lookup (eq_pair (op =) eq_mode) (#proposed_names opt) (name, mode) fun show_steps (Options opt) = #show_steps opt fun show_intermediate_results (Options opt) = #show_intermediate_results opt fun show_proof_trace (Options opt) = #show_proof_trace opt fun show_modes (Options opt) = #show_modes opt fun show_mode_inference (Options opt) = #show_mode_inference opt fun show_compilation (Options opt) = #show_compilation opt fun show_caught_failures (Options opt) = #show_caught_failures opt fun show_invalid_clauses (Options opt) = #show_invalid_clauses opt fun skip_proof (Options opt) = #skip_proof opt fun function_flattening (Options opt) = #function_flattening opt fun fail_safe_function_flattening (Options opt) = #fail_safe_function_flattening opt fun specialise (Options opt) = #specialise opt fun no_topmost_reordering (Options opt) = #no_topmost_reordering opt fun no_higher_order_predicate (Options opt) = #no_higher_order_predicate opt fun is_inductify (Options opt) = #inductify opt fun compilation (Options opt) = #compilation opt fun detect_switches (Options opt) = #detect_switches opt fun smart_depth_limiting (Options opt) = #smart_depth_limiting opt val default_options = Options { expected_modes = NONE, proposed_modes = [], proposed_names = [], show_steps = false, show_intermediate_results = false, show_proof_trace = false, show_modes = false, show_mode_inference = false, show_compilation = false, show_caught_failures = false, show_invalid_clauses = false, skip_proof = true, no_topmost_reordering = false, function_flattening = false, specialise = false, fail_safe_function_flattening = false, no_higher_order_predicate = [], inductify = false, detect_switches = true, smart_depth_limiting = false, compilation = Pred } val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace", "show_modes", "show_mode_inference", "show_compilation", "show_invalid_clauses", "skip_proof", "inductify", "no_function_flattening", "detect_switches", "specialise", "no_topmost_reordering", "smart_depth_limiting"] fun print_step options s = if show_steps options then tracing s else () (* simple transformations *) (** tuple processing **) fun rewrite_args [] (pats, intro_t, ctxt) = (pats, intro_t, ctxt) | rewrite_args (arg::args) (pats, intro_t, ctxt) = (case HOLogic.strip_tupleT (fastype_of arg) of (_ :: _ :: _) => let fun rewrite_arg' (Const (\<^const_name>\Pair\, _) $ _ $ t2, Type (\<^type_name>\Product_Type.prod\, [_, T2])) (args, (pats, intro_t, ctxt)) = rewrite_arg' (t2, T2) (args, (pats, intro_t, ctxt)) | rewrite_arg' (t, Type (\<^type_name>\Product_Type.prod\, [T1, T2])) (args, (pats, intro_t, ctxt)) = let val thy = Proof_Context.theory_of ctxt val ([x, y], ctxt') = Variable.variant_fixes ["x", "y"] ctxt val pat = (t, HOLogic.mk_prod (Free (x, T1), Free (y, T2))) val intro_t' = Pattern.rewrite_term thy [pat] [] intro_t val args' = map (Pattern.rewrite_term thy [pat] []) args in rewrite_arg' (Free (y, T2), T2) (args', (pat::pats, intro_t', ctxt')) end | rewrite_arg' _ (args, (pats, intro_t, ctxt)) = (args, (pats, intro_t, ctxt)) val (args', (pats, intro_t', ctxt')) = rewrite_arg' (arg, fastype_of arg) (args, (pats, intro_t, ctxt)) in rewrite_args args' (pats, intro_t', ctxt') end | _ => rewrite_args args (pats, intro_t, ctxt)) fun rewrite_prem atom = let val (_, args) = strip_comb atom in rewrite_args args end fun split_conjuncts_in_assms ctxt th = let val ((_, [fixed_th]), ctxt') = Variable.import false [th] ctxt fun split_conjs i nprems th = if i > nprems then th else (case try (op RSN) (@{thm conjI}, (i, th)) of SOME th' => split_conjs i (nprems + 1) th' | NONE => split_conjs (i + 1) nprems th) in singleton (Variable.export ctxt' ctxt) (split_conjs 1 (Thm.nprems_of fixed_th) fixed_th) end fun dest_conjunct_prem th = (case HOLogic.dest_Trueprop (Thm.prop_of th) of (Const (\<^const_name>\HOL.conj\, _) $ _ $ _) => dest_conjunct_prem (th RS @{thm conjunct1}) @ dest_conjunct_prem (th RS @{thm conjunct2}) | _ => [th]) fun expand_tuples thy intro = let val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *) val (((T_insts, t_insts), [intro']), ctxt1) = Variable.import false [intro] ctxt val intro_t = Thm.prop_of intro' val concl = Logic.strip_imp_concl intro_t val (_, args) = strip_comb (HOLogic.dest_Trueprop concl) val (pats', intro_t', ctxt2) = rewrite_args args ([], intro_t, ctxt1) val (pats', _, ctxt3) = fold_atoms rewrite_prem intro_t' (pats', intro_t', ctxt2) fun rewrite_pat (ct1, ct2) = (ct1, Thm.cterm_of ctxt3 (Pattern.rewrite_term thy pats' [] (Thm.term_of ct2))) - val t_insts' = map rewrite_pat t_insts - val intro'' = Thm.instantiate (T_insts, t_insts') intro + val t_insts' = map rewrite_pat (Term_Subst.Vars.dest t_insts) + val intro'' = Thm.instantiate (Term_Subst.TVars.dest T_insts, t_insts') intro val [intro'''] = Variable.export ctxt3 ctxt [intro''] val intro'''' = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm prod.inject}]) intro''' (* splitting conjunctions introduced by prod.inject*) val intro''''' = split_conjuncts_in_assms ctxt intro'''' in intro''''' end (** making case distributivity rules **) (*** this should be part of the datatype package ***) fun datatype_name_of_case_name thy = Ctr_Sugar.ctr_sugar_of_case (Proof_Context.init_global thy) #> the #> #ctrs #> hd #> fastype_of #> body_type #> dest_Type #> fst fun make_case_comb thy Tcon = let val ctxt = Proof_Context.init_global thy val SOME {casex, ...} = Ctr_Sugar.ctr_sugar_of ctxt Tcon val casex' = Type.legacy_freeze casex val Ts = BNF_Util.binder_fun_types (fastype_of casex') in list_comb (casex', map_index (fn (j, T) => Free ("f" ^ string_of_int j, T)) Ts) end fun make_case_distrib thy Tcon = let val comb = make_case_comb thy Tcon; val Type ("fun", [T, T']) = fastype_of comb; val (Const (case_name, _), fs) = strip_comb comb val used = Term.add_tfree_names comb [] val U = TFree (singleton (Name.variant_list used) "'t", \<^sort>\type\) val x = Free ("x", T) val f = Free ("f", T' --> U) fun apply_f f' = let val Ts = binder_types (fastype_of f') val bs = map Bound ((length Ts - 1) downto 0) in fold_rev absdummy Ts (f $ (list_comb (f', bs))) end val fs' = map apply_f fs val case_c' = Const (case_name, (map fastype_of fs') @ [T] ---> U) in HOLogic.mk_eq (f $ (comb $ x), list_comb (case_c', fs') $ x) end fun case_rewrite thy Tcon = (Drule.export_without_context o Skip_Proof.make_thm thy o HOLogic.mk_Trueprop) (make_case_distrib thy Tcon) fun instantiated_case_rewrite thy Tcon = let val th = case_rewrite thy Tcon val ctxt = Proof_Context.init_global thy val f = fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th))))) val Type ("fun", [uninst_T, uninst_T']) = fastype_of f val ([yname], ctxt') = Variable.add_fixes ["y"] ctxt val T' = TFree ("'t'", \<^sort>\type\) val U = TFree ("'u", \<^sort>\type\) val y = Free (yname, U) val f' = absdummy (U --> T') (Bound 0 $ y) val th' = Thm.instantiate ([(dest_TVar uninst_T, Thm.ctyp_of ctxt' (U --> T')), (dest_TVar uninst_T', Thm.ctyp_of ctxt' T')], [((fst (dest_Var f), (U --> T') --> T'), Thm.cterm_of ctxt' f')]) th val [th'] = Variable.export (Variable.declare_thm th' ctxt') ctxt [th'] in th' end fun case_betapply thy t = let val case_name = fst (dest_Const (fst (strip_comb t))) val Tcon = datatype_name_of_case_name thy case_name val th = instantiated_case_rewrite thy Tcon in Raw_Simplifier.rewrite_term thy [th RS @{thm eq_reflection}] [] t end (*** conversions ***) fun imp_prems_conv cv ct = (case Thm.term_of ct of Const (\<^const_name>\Pure.imp\, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct | _ => Conv.all_conv ct) (** eta contract higher-order arguments **) fun eta_contract_ho_arguments thy intro = let fun f atom = list_comb (apsnd ((map o map_products) Envir.eta_contract) (strip_comb atom)) in map_term thy (map_concl f o map_atoms f) intro end (** remove equalities **) fun remove_equalities thy intro = let fun remove_eqs intro_t = let val (prems, concl) = Logic.strip_horn intro_t fun remove_eq (prems, concl) = let fun removable_eq prem = (case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) prem of SOME (lhs, rhs) => (case lhs of Var _ => true | _ => (case rhs of Var _ => true | _ => false)) | NONE => false) in (case find_first removable_eq prems of NONE => (prems, concl) | SOME eq => let val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq) val prems' = remove (op =) eq prems val subst = (case lhs of (v as Var _) => (fn t => if t = v then rhs else t) | _ => (case rhs of (v as Var _) => (fn t => if t = v then lhs else t))) in remove_eq (map (map_aterms subst) prems', map_aterms subst concl) end) end in Logic.list_implies (remove_eq (prems, concl)) end in map_term thy remove_eqs intro end (* Some last processing *) fun remove_pointless_clauses intro = if Logic.strip_imp_prems (Thm.prop_of intro) = [\<^prop>\False\] then [] else [intro] (* some peephole optimisations *) fun peephole_optimisation thy intro = let val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *) val process = rewrite_rule ctxt (Named_Theorems.get ctxt \<^named_theorems>\code_pred_simp\) fun process_False intro_t = if member (op =) (Logic.strip_imp_prems intro_t) \<^prop>\False\ then NONE else SOME intro_t fun process_True intro_t = map_filter_premises (fn p => if p = \<^prop>\True\ then NONE else SOME p) intro_t in Option.map (Skip_Proof.make_thm thy) (process_False (process_True (Thm.prop_of (process intro)))) end (* importing introduction rules *) fun import_intros inp_pred [] ctxt = let val (outp_pred, ctxt') = yield_singleton (Variable.import_terms true) inp_pred ctxt val T = fastype_of outp_pred val paramTs = ho_argsT_of_typ (binder_types T) val (param_names, _) = Variable.variant_fixes (map (fn i => "p" ^ (string_of_int i)) (1 upto (length paramTs))) ctxt' val params = map2 (curry Free) param_names paramTs in (((outp_pred, params), []), ctxt') end | import_intros inp_pred (th :: ths) ctxt = let val ((_, [th']), ctxt') = Variable.import true [th] ctxt val thy = Proof_Context.theory_of ctxt' val (pred, args) = strip_intro_concl th' val T = fastype_of pred val ho_args = ho_args_of_typ T args fun subst_of (pred', pred) = let val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty handle Type.TYPE_MATCH => error ("Type mismatch of predicate " ^ fst (dest_Const pred) ^ " (trying to match " ^ Syntax.string_of_typ ctxt' (fastype_of pred') ^ " and " ^ Syntax.string_of_typ ctxt' (fastype_of pred) ^ ")" ^ " in " ^ Thm.string_of_thm ctxt' th) in Vartab.fold (fn (xi, (S, T)) => Term_Subst.TVars.add ((xi, S), T)) subst Term_Subst.TVars.empty end fun instantiate_typ th = let val (pred', _) = strip_intro_concl th val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then raise Fail "Trying to instantiate another predicate" else () val instT = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt' T)) (subst_of (pred', pred)) []; in Thm.instantiate (instT, []) th end fun instantiate_ho_args th = let val (_, args') = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o Thm.prop_of) th val ho_args' = map dest_Var (ho_args_of_typ T args') in Thm.instantiate ([], ho_args' ~~ map (Thm.cterm_of ctxt') ho_args) th end val outp_pred = Term_Subst.instantiate (subst_of (inp_pred, pred), Term_Subst.Vars.empty) inp_pred val ((_, ths'), ctxt1) = Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt' in (((outp_pred, ho_args), th' :: ths'), ctxt1) end (* generation of case rules from user-given introduction rules *) fun mk_args2 (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) st = let val (t1, st') = mk_args2 T1 st val (t2, st'') = mk_args2 T2 st' in (HOLogic.mk_prod (t1, t2), st'') end (*| mk_args2 (T as Type ("fun", _)) (params, ctxt) = let val (S, U) = strip_type T in if U = HOLogic.boolT then (hd params, (tl params, ctxt)) else let val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt in (Free (x, T), (params, ctxt')) end end*) | mk_args2 T (params, ctxt) = let val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt in (Free (x, T), (params, ctxt')) end fun mk_casesrule ctxt pred introrules = let (* TODO: can be simplified if parameters are not treated specially ? *) val (((pred, params), intros_th), ctxt1) = import_intros pred introrules ctxt (* TODO: distinct required ? -- test case with more than one parameter! *) val params = distinct (op aconv) params val intros = map Thm.prop_of intros_th val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1 val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT)) val argsT = binder_types (fastype_of pred) (* TODO: can be simplified if parameters are not treated specially ? <-- see uncommented code! *) val (argvs, _) = fold_map mk_args2 argsT (params, ctxt2) fun mk_case intro = let val (_, args) = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl) intro val prems = Logic.strip_imp_prems intro val eqprems = map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) argvs args val frees = map Free (fold Term.add_frees (args @ prems) []) in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end val assm = HOLogic.mk_Trueprop (list_comb (pred, argvs)) val cases = map mk_case intros in Logic.list_implies (assm :: cases, prop) end; (* unifying constants to have the same type variables *) fun unify_consts thy cs intr_ts = let val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I); fun varify (t, (i, ts)) = - let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify_global [] t)) + let + val t' = map_types (Logic.incr_tvar (i + 1)) + (#2 (Type.varify_global Term_Subst.TFrees.empty t)) in (maxidx_of_term t', t' :: ts) end val (i, cs') = List.foldr varify (~1, []) cs val (i', intr_ts') = List.foldr varify (i, []) intr_ts val rec_consts = fold add_term_consts_2 cs' [] val intr_consts = fold add_term_consts_2 intr_ts' [] fun unify (cname, cT) = let val consts = map snd (filter (fn c => fst c = cname) intr_consts) in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end val (env, _) = fold unify rec_consts (Vartab.empty, i') val subst = map_types (Envir.norm_type env) in (map subst cs', map subst intr_ts') end handle Type.TUNIFY => (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts)) (* preprocessing rules *) fun preprocess_equality thy rule = Conv.fconv_rule (imp_prems_conv (HOLogic.Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq}))))) (Thm.transfer thy rule) fun preprocess_intro thy = expand_tuples thy #> preprocess_equality thy (* defining a quickcheck predicate *) fun strip_imp_prems (Const(\<^const_name>\HOL.implies\, _) $ A $ B) = A :: strip_imp_prems B | strip_imp_prems _ = []; fun strip_imp_concl (Const(\<^const_name>\HOL.implies\, _) $ _ $ B) = strip_imp_concl B | strip_imp_concl A = A; fun strip_horn A = (strip_imp_prems A, strip_imp_concl A) fun define_quickcheck_predicate t thy = let val (vs, t') = strip_abs t val vs' = Variable.variant_frees (Proof_Context.init_global thy) [] vs (* FIXME proper context!? *) val t'' = subst_bounds (map Free (rev vs'), t') val (prems, concl) = strip_horn t'' val constname = "quickcheck" val full_constname = Sign.full_bname thy constname val constT = map snd vs' ---> \<^typ>\bool\ val thy1 = Sign.add_consts [(Binding.name constname, constT, NoSyn)] thy val const = Const (full_constname, constT) val t = Logic.list_implies (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]), HOLogic.mk_Trueprop (list_comb (const, map Free vs'))) val intro = Goal.prove (Proof_Context.init_global thy1) (map fst vs') [] t (fn {context = ctxt, ...} => ALLGOALS (Skip_Proof.cheat_tac ctxt)) in ((((full_constname, constT), vs'), intro), thy1) end end diff --git a/src/HOL/Tools/Qelim/qelim.ML b/src/HOL/Tools/Qelim/qelim.ML --- a/src/HOL/Tools/Qelim/qelim.ML +++ b/src/HOL/Tools/Qelim/qelim.ML @@ -1,73 +1,73 @@ (* Title: HOL/Tools/Qelim/qelim.ML Author: Amine Chaieb, TU Muenchen Generic quantifier elimination conversions for HOL formulae. *) signature QELIM = sig val gen_qelim_conv: Proof.context -> conv -> conv -> conv -> (cterm -> 'a -> 'a) -> 'a -> ('a -> conv) -> ('a -> conv) -> ('a -> conv) -> conv val standard_qelim_conv: Proof.context -> (cterm list -> conv) -> (cterm list -> conv) -> (cterm list -> conv) -> conv end; structure Qelim: QELIM = struct val all_not_ex = mk_meta_eq @{thm "all_not_ex"}; fun gen_qelim_conv ctxt precv postcv simpex_conv ins env atcv ncv qcv = let fun conv env p = case Thm.term_of p of Const(s,T)$_$_ => if domain_type T = HOLogic.boolT andalso member (op =) [\<^const_name>\HOL.conj\, \<^const_name>\HOL.disj\, \<^const_name>\HOL.implies\, \<^const_name>\HOL.eq\] s then Conv.binop_conv (conv env) p else atcv env p | Const(\<^const_name>\Not\,_)$_ => Conv.arg_conv (conv env) p | Const(\<^const_name>\Ex\,_)$Abs(s,_,_) => let val (e,p0) = Thm.dest_comb p val (x,p') = Thm.dest_abs (SOME s) p0 val env' = ins x env val th = Thm.abstract_rule s x ((conv env' then_conv ncv env') p') |> Drule.arg_cong_rule e val th' = simpex_conv (Thm.rhs_of th) val (_, r) = Thm.dest_equals (Thm.cprop_of th') in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th)) else Thm.transitive (Thm.transitive th th') (conv env r) end | Const(\<^const_name>\Ex\,_)$ _ => (Thm.eta_long_conversion then_conv conv env) p | Const(\<^const_name>\All\, _)$_ => let val allT = Thm.ctyp_of_cterm (Thm.dest_fun p) val T = Thm.dest_ctyp0 (Thm.dest_ctyp0 allT) val p = Thm.dest_arg p val th = Thm.instantiate' [SOME T] [SOME p] all_not_ex in Thm.transitive th (conv env (Thm.rhs_of th)) end | _ => atcv env p in precv then_conv (conv env) then_conv postcv end (* Instantiation of some parameter for most common cases *) local val ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib}); fun pcv ctxt = Simplifier.rewrite (put_simpset ss ctxt) in fun standard_qelim_conv ctxt atcv ncv qcv p = let val pcv = pcv ctxt - in gen_qelim_conv ctxt pcv pcv pcv cons (Drule.cterm_add_frees p []) atcv ncv qcv p end + in gen_qelim_conv ctxt pcv pcv pcv cons (Misc_Legacy.cterm_frees p) atcv ncv qcv p end end; end; diff --git a/src/HOL/Tools/Quotient/quotient_tacs.ML b/src/HOL/Tools/Quotient/quotient_tacs.ML --- a/src/HOL/Tools/Quotient/quotient_tacs.ML +++ b/src/HOL/Tools/Quotient/quotient_tacs.ML @@ -1,764 +1,764 @@ (* Title: HOL/Tools/Quotient/quotient_tacs.ML Author: Cezary Kaliszyk and Christian Urban Tactics for solving goal arising from lifting theorems to quotient types. *) signature QUOTIENT_TACS = sig val regularize_tac: Proof.context -> int -> tactic val injection_tac: Proof.context -> int -> tactic val all_injection_tac: Proof.context -> int -> tactic val clean_tac: Proof.context -> int -> tactic val descend_procedure_tac: Proof.context -> thm list -> int -> tactic val descend_tac: Proof.context -> thm list -> int -> tactic val partiality_descend_procedure_tac: Proof.context -> thm list -> int -> tactic val partiality_descend_tac: Proof.context -> thm list -> int -> tactic val lift_procedure_tac: Proof.context -> thm list -> thm -> int -> tactic val lift_tac: Proof.context -> thm list -> thm list -> int -> tactic val lifted: Proof.context -> typ list -> thm list -> thm -> thm val lifted_attrib: attribute end; structure Quotient_Tacs: QUOTIENT_TACS = struct (** various helper fuctions **) (* Since HOL_basic_ss is too "big" for us, we *) (* need to set up our own minimal simpset. *) fun mk_minimal_simpset ctxt = empty_simpset ctxt |> Simplifier.set_subgoaler asm_simp_tac |> Simplifier.set_mksimps (mksimps []) (* composition of two theorems, used in maps *) fun OF1 thm1 thm2 = thm2 RS thm1 fun atomize_thm ctxt thm = let - val thm' = Thm.legacy_freezeT (forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? no! *) + val thm' = Thm.legacy_freezeT (Thm.forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? no! *) val thm'' = Object_Logic.atomize ctxt (Thm.cprop_of thm') in @{thm equal_elim_rule1} OF [thm'', thm'] end (*** Regularize Tactic ***) (** solvers for equivp and quotient assumptions **) fun equiv_tac ctxt = REPEAT_ALL_NEW (resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_equiv\))) val equiv_solver = mk_solver "Equivalence goal solver" equiv_tac fun quotient_tac ctxt = (REPEAT_ALL_NEW (FIRST' [resolve_tac ctxt @{thms identity_quotient3}, resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_thm\))])) val quotient_solver = mk_solver "Quotient goal solver" quotient_tac fun solve_quotient_assm ctxt thm = case Seq.pull (quotient_tac ctxt 1 thm) of SOME (t, _) => t | _ => error "Solve_quotient_assm failed. Possibly a quotient theorem is missing." fun get_match_inst thy pat trm = let val univ = Unify.matchers (Context.Theory thy) [(pat, trm)] val SOME (env, _) = Seq.pull univ (* raises Bind, if no unifier *) (* FIXME fragile *) val tenv = Vartab.dest (Envir.term_env env) val tyenv = Vartab.dest (Envir.type_env env) fun prep_ty (x, (S, ty)) = ((x, S), Thm.global_ctyp_of thy ty) fun prep_trm (x, (T, t)) = ((x, T), Thm.global_cterm_of thy t) in (map prep_ty tyenv, map prep_trm tenv) end (* Calculates the instantiations for the lemmas: ball_reg_eqv_range and bex_reg_eqv_range Since the left-hand-side contains a non-pattern '?P (f ?x)' we rely on unification/instantiation to check whether the theorem applies and return NONE if it doesn't. *) fun calculate_inst ctxt ball_bex_thm redex R1 R2 = let val thy = Proof_Context.theory_of ctxt fun get_lhs thm = fst (Logic.dest_equals (Thm.concl_of thm)) val ty_inst = map (SOME o Thm.ctyp_of ctxt) [domain_type (fastype_of R2)] val trm_inst = map (SOME o Thm.cterm_of ctxt) [R2, R1] in (case try (Thm.instantiate' ty_inst trm_inst) ball_bex_thm of NONE => NONE | SOME thm' => (case try (get_match_inst thy (get_lhs thm')) (Thm.term_of redex) of NONE => NONE | SOME inst2 => try (Drule.instantiate_normalize inst2) thm')) end fun ball_bex_range_simproc ctxt redex = (case Thm.term_of redex of (Const (\<^const_name>\Ball\, _) $ (Const (\<^const_name>\Respects\, _) $ (Const (\<^const_name>\rel_fun\, _) $ R1 $ R2)) $ _) => calculate_inst ctxt @{thm ball_reg_eqv_range[THEN eq_reflection]} redex R1 R2 | (Const (\<^const_name>\Bex\, _) $ (Const (\<^const_name>\Respects\, _) $ (Const (\<^const_name>\rel_fun\, _) $ R1 $ R2)) $ _) => calculate_inst ctxt @{thm bex_reg_eqv_range[THEN eq_reflection]} redex R1 R2 | _ => NONE) (* Regularize works as follows: 0. preliminary simplification step according to ball_reg_eqv bex_reg_eqv babs_reg_eqv ball_reg_eqv_range bex_reg_eqv_range 1. eliminating simple Ball/Bex instances (ball_reg_right bex_reg_left) 2. monos 3. commutation rules for ball and bex (ball_all_comm bex_ex_comm) 4. then rel-equalities, which need to be instantiated with 'eq_imp_rel' to avoid loops 5. then simplification like 0 finally jump back to 1 *) fun reflp_get ctxt = map_filter (fn th => if Thm.no_prems th then SOME (OF1 @{thm equivp_reflp} th) else NONE handle THM _ => NONE) (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_equiv\)) val eq_imp_rel = @{lemma "equivp R \ a = b \ R a b" by (simp add: equivp_reflp)} fun eq_imp_rel_get ctxt = map (OF1 eq_imp_rel) (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_equiv\)) val regularize_simproc = Simplifier.make_simproc \<^context> "regularize" {lhss = [\<^term>\Ball (Respects (R1 ===> R2)) P\, \<^term>\Bex (Respects (R1 ===> R2)) P\], proc = K ball_bex_range_simproc}; fun regularize_tac ctxt = let val simpset = mk_minimal_simpset ctxt addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp} addsimprocs [regularize_simproc] addSolver equiv_solver addSolver quotient_solver val eq_eqvs = eq_imp_rel_get ctxt in simp_tac simpset THEN' TRY o REPEAT_ALL_NEW (CHANGED o FIRST' [resolve_tac ctxt @{thms ball_reg_right bex_reg_left bex1_bexeq_reg}, resolve_tac ctxt (Inductive.get_monos ctxt), resolve_tac ctxt @{thms ball_all_comm bex_ex_comm}, resolve_tac ctxt eq_eqvs, simp_tac simpset]) end (*** Injection Tactic ***) (* Looks for Quot_True assumptions, and in case its parameter is an application, it returns the function and the argument. *) fun find_qt_asm asms = let fun find_fun trm = (case trm of (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\Quot_True\, _) $ _)) => true | _ => false) in (case find_first find_fun asms of SOME (_ $ (_ $ (f $ a))) => SOME (f, a) | _ => NONE) end fun quot_true_simple_conv ctxt fnctn ctrm = (case Thm.term_of ctrm of (Const (\<^const_name>\Quot_True\, _) $ x) => let val fx = fnctn x; val cx = Thm.cterm_of ctxt x; val cfx = Thm.cterm_of ctxt fx; val cxt = Thm.ctyp_of ctxt (fastype_of x); val cfxt = Thm.ctyp_of ctxt (fastype_of fx); val thm = Thm.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QT_imp} in Conv.rewr_conv thm ctrm end) fun quot_true_conv ctxt fnctn ctrm = (case Thm.term_of ctrm of (Const (\<^const_name>\Quot_True\, _) $ _) => quot_true_simple_conv ctxt fnctn ctrm | _ $ _ => Conv.comb_conv (quot_true_conv ctxt fnctn) ctrm | Abs _ => Conv.abs_conv (fn (_, ctxt) => quot_true_conv ctxt fnctn) ctxt ctrm | _ => Conv.all_conv ctrm) fun quot_true_tac ctxt fnctn = CONVERSION ((Conv.params_conv ~1 (fn ctxt => (Conv.prems_conv ~1 (quot_true_conv ctxt fnctn)))) ctxt) fun dest_comb (f $ a) = (f, a) fun dest_bcomb ((_ $ l) $ r) = (l, r) fun unlam t = (case t of Abs a => snd (Term.dest_abs a) | _ => unlam (Abs("", domain_type (fastype_of t), (incr_boundvars 1 t) $ (Bound 0)))) val bare_concl = HOLogic.dest_Trueprop o Logic.strip_assums_concl (* We apply apply_rsp only in case if the type needs lifting. This is the case if the type of the data in the Quot_True assumption is different from the corresponding type in the goal. *) val apply_rsp_tac = Subgoal.FOCUS (fn {concl, asms, context = ctxt,...} => let val bare_concl = HOLogic.dest_Trueprop (Thm.term_of concl) val qt_asm = find_qt_asm (map Thm.term_of asms) in case (bare_concl, qt_asm) of (R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) => if fastype_of qt_fun = fastype_of f then no_tac else let val ty_x = fastype_of x val ty_b = fastype_of qt_arg val ty_f = range_type (fastype_of f) val ty_inst = map (SOME o Thm.ctyp_of ctxt) [ty_x, ty_b, ty_f] val t_inst = map (SOME o Thm.cterm_of ctxt) [R2, f, g, x, y]; val inst_thm = Thm.instantiate' ty_inst ([NONE, NONE, NONE] @ t_inst) @{thm apply_rspQ3} in (resolve_tac ctxt [inst_thm] THEN' SOLVED' (quotient_tac ctxt)) 1 end | _ => no_tac end) (* Instantiates and applies 'equals_rsp'. Since the theorem is complex we rely on instantiation to tell us if it applies *) fun equals_rsp_tac R ctxt = case try (Thm.cterm_of ctxt) R of (* There can be loose bounds in R *) (* FIXME fragile *) SOME ctm => let val ty = domain_type (fastype_of R) in case try (Thm.instantiate' [SOME (Thm.ctyp_of ctxt ty)] [SOME (Thm.cterm_of ctxt R)]) @{thm equals_rsp} of SOME thm => resolve_tac ctxt [thm] THEN' quotient_tac ctxt | NONE => K no_tac end | _ => K no_tac fun rep_abs_rsp_tac ctxt = SUBGOAL (fn (goal, i) => (case try bare_concl goal of SOME (rel $ _ $ (rep $ (abs $ _))) => (let val (ty_a, ty_b) = dest_funT (fastype_of abs); val ty_inst = map (SOME o Thm.ctyp_of ctxt) [ty_a, ty_b]; in case try (map (SOME o Thm.cterm_of ctxt)) [rel, abs, rep] of SOME t_inst => (case try (Thm.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of SOME inst_thm => (resolve_tac ctxt [inst_thm] THEN' quotient_tac ctxt) i | NONE => no_tac) | NONE => no_tac end handle TERM _ => no_tac) | _ => no_tac)) (* Injection means to prove that the regularized theorem implies the abs/rep injected one. The deterministic part: - remove lambdas from both sides - prove Ball/Bex/Babs equalities using ball_rsp, bex_rsp, babs_rsp - prove Ball/Bex relations using rel_funI - reflexivity of equality - prove equality of relations using equals_rsp - use user-supplied RSP theorems - solve 'relation of relations' goals using quot_rel_rsp - remove rep_abs from the right side (Lambdas under respects may have left us some assumptions) Then in order: - split applications of lifted type (apply_rsp) - split applications of non-lifted type (cong_tac) - apply extentionality - assumption - reflexivity of the relation *) fun injection_match_tac ctxt = SUBGOAL (fn (goal, i) => (case bare_concl goal of (* (R1 ===> R2) (%x...) (%x...) ----> [|R1 x y|] ==> R2 (...x) (...y) *) (Const (\<^const_name>\rel_fun\, _) $ _ $ _) $ (Abs _) $ (Abs _) => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *) | (Const (\<^const_name>\HOL.eq\,_) $ (Const(\<^const_name>\Ball\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Ball\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _)) => resolve_tac ctxt @{thms ball_rsp} THEN' dresolve_tac ctxt @{thms QT_all} (* (R1 ===> op =) (Ball...) (Ball...) ----> [|R1 x y|] ==> (Ball...x) = (Ball...y) *) | (Const (\<^const_name>\rel_fun\, _) $ _ $ _) $ (Const(\<^const_name>\Ball\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Ball\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *) | Const (\<^const_name>\HOL.eq\,_) $ (Const(\<^const_name>\Bex\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Bex\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) => resolve_tac ctxt @{thms bex_rsp} THEN' dresolve_tac ctxt @{thms QT_ex} (* (R1 ===> op =) (Bex...) (Bex...) ----> [|R1 x y|] ==> (Bex...x) = (Bex...y) *) | (Const (\<^const_name>\rel_fun\, _) $ _ $ _) $ (Const(\<^const_name>\Bex\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Bex\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam | (Const (\<^const_name>\rel_fun\, _) $ _ $ _) $ (Const(\<^const_name>\Bex1_rel\,_) $ _) $ (Const(\<^const_name>\Bex1_rel\,_) $ _) => resolve_tac ctxt @{thms bex1_rel_rsp} THEN' quotient_tac ctxt | (_ $ (Const(\<^const_name>\Babs\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Babs\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _)) => resolve_tac ctxt @{thms babs_rsp} THEN' quotient_tac ctxt | Const (\<^const_name>\HOL.eq\,_) $ (R $ _ $ _) $ (_ $ _ $ _) => (resolve_tac ctxt @{thms refl} ORELSE' (equals_rsp_tac R ctxt THEN' RANGE [ quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)])) (* reflexivity of operators arising from Cong_tac *) | Const (\<^const_name>\HOL.eq\,_) $ _ $ _ => resolve_tac ctxt @{thms refl} (* respectfulness of constants; in particular of a simple relation *) | _ $ (Const _) $ (Const _) (* rel_fun, list_rel, etc but not equality *) => resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_respect\)) THEN_ALL_NEW quotient_tac ctxt (* R (...) (Rep (Abs ...)) ----> R (...) (...) *) (* observe map_fun *) | _ $ _ $ _ => (resolve_tac ctxt @{thms quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt) ORELSE' rep_abs_rsp_tac ctxt | _ => K no_tac) i) fun injection_step_tac ctxt rel_refl = FIRST' [ injection_match_tac ctxt, (* R (t $ ...) (t' $ ...) ----> apply_rsp provided type of t needs lifting *) apply_rsp_tac ctxt THEN' RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)], (* (op =) (t $ ...) (t' $ ...) ----> Cong provided type of t does not need lifting *) (* merge with previous tactic *) Cong_Tac.cong_tac ctxt @{thm cong} THEN' RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)], (* resolving with R x y assumptions *) assume_tac ctxt, (* (op =) (%x...) (%y...) ----> (op =) (...) (...) *) resolve_tac ctxt @{thms ext} THEN' quot_true_tac ctxt unlam, (* reflexivity of the basic relations *) (* R ... ... *) resolve_tac ctxt rel_refl] fun injection_tac ctxt = let val rel_refl = reflp_get ctxt in injection_step_tac ctxt rel_refl end fun all_injection_tac ctxt = REPEAT_ALL_NEW (injection_tac ctxt) (*** Cleaning of the Theorem ***) (* expands all map_funs, except in front of the (bound) variables listed in xs *) fun map_fun_simple_conv xs ctrm = (case Thm.term_of ctrm of ((Const (\<^const_name>\map_fun\, _) $ _ $ _) $ h $ _) => if member (op=) xs h then Conv.all_conv ctrm else Conv.rewr_conv @{thm map_fun_apply [THEN eq_reflection]} ctrm | _ => Conv.all_conv ctrm) fun map_fun_conv xs ctxt ctrm = (case Thm.term_of ctrm of _ $ _ => (Conv.comb_conv (map_fun_conv xs ctxt) then_conv map_fun_simple_conv xs) ctrm | Abs _ => Conv.abs_conv (fn (x, ctxt) => map_fun_conv (Thm.term_of x :: xs) ctxt) ctxt ctrm | _ => Conv.all_conv ctrm) fun map_fun_tac ctxt = CONVERSION (map_fun_conv [] ctxt) (* custom matching functions *) fun mk_abs u i t = if incr_boundvars i u aconv t then Bound i else case t of t1 $ t2 => mk_abs u i t1 $ mk_abs u i t2 | Abs (s, T, t') => Abs (s, T, mk_abs u (i + 1) t') | Bound j => if i = j then error "make_inst" else t | _ => t fun make_inst lhs t = let val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs; val _ $ (Abs (_, _, (_ $ g))) = t; in (f, Abs ("x", T, mk_abs u 0 g)) end fun make_inst_id lhs t = let val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs; val _ $ (Abs (_, _, g)) = t; in (f, Abs ("x", T, mk_abs u 0 g)) end (* Simplifies a redex using the 'lambda_prs' theorem. First instantiates the types and known subterms. Then solves the quotient assumptions to get Rep2 and Abs1 Finally instantiates the function f using make_inst If Rep2 is an identity then the pattern is simpler and make_inst_id is used *) fun lambda_prs_simple_conv ctxt ctrm = (case Thm.term_of ctrm of (Const (\<^const_name>\map_fun\, _) $ r1 $ a2) $ (Abs _) => let val (ty_b, ty_a) = dest_funT (fastype_of r1) val (ty_c, ty_d) = dest_funT (fastype_of a2) val tyinst = map (SOME o Thm.ctyp_of ctxt) [ty_a, ty_b, ty_c, ty_d] val tinst = [NONE, NONE, SOME (Thm.cterm_of ctxt r1), NONE, SOME (Thm.cterm_of ctxt a2)] val thm1 = Thm.instantiate' tyinst tinst @{thm lambda_prs[THEN eq_reflection]} val thm2 = solve_quotient_assm ctxt (solve_quotient_assm ctxt thm1) val thm3 = rewrite_rule ctxt @{thms id_apply[THEN eq_reflection]} thm2 val (insp, inst) = if ty_c = ty_d then make_inst_id (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm) else make_inst (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm) val thm4 = Drule.instantiate_normalize ([], [(dest_Var insp, Thm.cterm_of ctxt inst)]) thm3 in Conv.rewr_conv thm4 ctrm end | _ => Conv.all_conv ctrm) fun lambda_prs_conv ctxt = Conv.top_conv lambda_prs_simple_conv ctxt fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt) (* Cleaning consists of: 1. unfolding of ---> in front of everything, except bound variables (this prevents lambda_prs from becoming stuck) 2. simplification with lambda_prs 3. simplification with: - Quotient_abs_rep Quotient_rel_rep babs_prs all_prs ex_prs ex1_prs - id_simps and preservation lemmas and - symmetric versions of the definitions (that is definitions of quotient constants are folded) 4. test for refl *) fun clean_tac ctxt = let val thy = Proof_Context.theory_of ctxt val defs = map (Thm.symmetric o #def) (Quotient_Info.dest_quotconsts_global thy) val prs = rev (Named_Theorems.get ctxt \<^named_theorems>\quot_preserve\) val ids = rev (Named_Theorems.get ctxt \<^named_theorems>\id_simps\) val thms = @{thms Quotient3_abs_rep Quotient3_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs val simpset = (mk_minimal_simpset ctxt) addsimps thms addSolver quotient_solver in EVERY' [ map_fun_tac ctxt, lambda_prs_tac ctxt, simp_tac simpset, TRY o resolve_tac ctxt [refl]] end (* Tactic for Generalising Free Variables in a Goal *) fun inst_spec ctrm = Thm.instantiate' [SOME (Thm.ctyp_of_cterm ctrm)] [NONE, SOME ctrm] @{thm spec} fun inst_spec_tac ctxt ctrms = EVERY' (map (dresolve_tac ctxt o single o inst_spec) ctrms) fun all_list xs trm = fold (fn (x, T) => fn t' => HOLogic.mk_all (x, T, t')) xs trm fun apply_under_Trueprop f = HOLogic.dest_Trueprop #> f #> HOLogic.mk_Trueprop fun gen_frees_tac ctxt = SUBGOAL (fn (concl, i) => let val vrs = Term.add_frees concl [] val cvrs = map (Thm.cterm_of ctxt o Free) vrs val concl' = apply_under_Trueprop (all_list vrs) concl val goal = Logic.mk_implies (concl', concl) val rule = Goal.prove ctxt [] [] goal (K (EVERY1 [inst_spec_tac ctxt (rev cvrs), assume_tac ctxt])) in resolve_tac ctxt [rule] i end) (** The General Shape of the Lifting Procedure **) (* - A is the original raw theorem - B is the regularized theorem - C is the rep/abs injected version of B - D is the lifted theorem - 1st prem is the regularization step - 2nd prem is the rep/abs injection step - 3rd prem is the cleaning part the Quot_True premise in 2nd records the lifted theorem *) val procedure_thm = @{lemma "\A; A \ B; Quot_True D \ B = C; C = D\ \ D" by (simp add: Quot_True_def)} (* in case of partial equivalence relations, this form of the procedure theorem results in solvable proof obligations *) val partiality_procedure_thm = @{lemma "[|B; Quot_True D ==> B = C; C = D|] ==> D" by (simp add: Quot_True_def)} fun lift_match_error ctxt msg rtrm qtrm = let val rtrm_str = Syntax.string_of_term ctxt rtrm val qtrm_str = Syntax.string_of_term ctxt qtrm val msg = cat_lines [enclose "[" "]" msg, "The quotient theorem", qtrm_str, "", "does not match with original theorem", rtrm_str] in error msg end fun procedure_inst ctxt rtrm qtrm = let val rtrm' = HOLogic.dest_Trueprop rtrm val qtrm' = HOLogic.dest_Trueprop qtrm val reg_goal = Quotient_Term.regularize_trm_chk ctxt (rtrm', qtrm') handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm val inj_goal = Quotient_Term.inj_repabs_trm_chk ctxt (reg_goal, qtrm') handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm in Thm.instantiate' [] [SOME (Thm.cterm_of ctxt rtrm'), SOME (Thm.cterm_of ctxt reg_goal), NONE, SOME (Thm.cterm_of ctxt inj_goal)] procedure_thm end (* Since we use Ball and Bex during the lifting and descending, we cannot deal with lemmas containing them, unless we unfold them by default. *) val default_unfolds = @{thms Ball_def Bex_def} (** descending as tactic **) fun descend_procedure_tac ctxt simps = let val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in full_simp_tac simpset THEN' Object_Logic.full_atomize_tac ctxt THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => let val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal val rule = procedure_inst ctxt rtrm goal in resolve_tac ctxt [rule] i end) end fun descend_tac ctxt simps = let val mk_tac_raw = descend_procedure_tac ctxt simps THEN' RANGE [Object_Logic.rulify_tac ctxt THEN' (K all_tac), regularize_tac ctxt, all_injection_tac ctxt, clean_tac ctxt] in Goal.conjunction_tac THEN_ALL_NEW mk_tac_raw end (** descending for partial equivalence relations **) fun partiality_procedure_inst ctxt rtrm qtrm = let val rtrm' = HOLogic.dest_Trueprop rtrm val qtrm' = HOLogic.dest_Trueprop qtrm val reg_goal = Quotient_Term.regularize_trm_chk ctxt (rtrm', qtrm') handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm val inj_goal = Quotient_Term.inj_repabs_trm_chk ctxt (reg_goal, qtrm') handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm in Thm.instantiate' [] [SOME (Thm.cterm_of ctxt reg_goal), NONE, SOME (Thm.cterm_of ctxt inj_goal)] partiality_procedure_thm end fun partiality_descend_procedure_tac ctxt simps = let val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in full_simp_tac simpset THEN' Object_Logic.full_atomize_tac ctxt THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => let val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal val rule = partiality_procedure_inst ctxt rtrm goal in resolve_tac ctxt [rule] i end) end fun partiality_descend_tac ctxt simps = let val mk_tac_raw = partiality_descend_procedure_tac ctxt simps THEN' RANGE [Object_Logic.rulify_tac ctxt THEN' (K all_tac), all_injection_tac ctxt, clean_tac ctxt] in Goal.conjunction_tac THEN_ALL_NEW mk_tac_raw end (** lifting as a tactic **) (* the tactic leaves three subgoals to be proved *) fun lift_procedure_tac ctxt simps rthm = let val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in full_simp_tac simpset THEN' Object_Logic.full_atomize_tac ctxt THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => let (* full_atomize_tac contracts eta redexes, so we do it also in the original theorem *) val rthm' = rthm |> full_simplify simpset |> Drule.eta_contraction_rule |> Thm.forall_intr_frees |> atomize_thm ctxt val rule = procedure_inst ctxt (Thm.prop_of rthm') goal in (resolve_tac ctxt [rule] THEN' resolve_tac ctxt [rthm']) i end) end fun lift_single_tac ctxt simps rthm = lift_procedure_tac ctxt simps rthm THEN' RANGE [ regularize_tac ctxt, all_injection_tac ctxt, clean_tac ctxt ] fun lift_tac ctxt simps rthms = Goal.conjunction_tac THEN' RANGE (map (lift_single_tac ctxt simps) rthms) (* automated lifting with pre-simplification of the theorems; for internal usage *) fun lifted ctxt qtys simps rthm = let val ((_, [rthm']), ctxt') = Variable.import true [rthm] ctxt val goal = Quotient_Term.derive_qtrm ctxt' qtys (Thm.prop_of rthm') in Goal.prove ctxt' [] [] goal (K (HEADGOAL (lift_single_tac ctxt' simps rthm'))) |> singleton (Proof_Context.export ctxt' ctxt) end (* lifting as an attribute *) val lifted_attrib = Thm.rule_attribute [] (fn context => let val ctxt = Context.proof_of context val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) in lifted ctxt qtys [] end) end; (* structure *) diff --git a/src/HOL/Tools/SMT/smt_normalize.ML b/src/HOL/Tools/SMT/smt_normalize.ML --- a/src/HOL/Tools/SMT/smt_normalize.ML +++ b/src/HOL/Tools/SMT/smt_normalize.ML @@ -1,535 +1,535 @@ (* Title: HOL/Tools/SMT/smt_normalize.ML Author: Sascha Boehme, TU Muenchen Normalization steps on theorems required by SMT solvers. *) signature SMT_NORMALIZE = sig val drop_fact_warning: Proof.context -> thm -> unit val atomize_conv: Proof.context -> conv val special_quant_table: (string * thm) list val case_bool_entry: string * thm val abs_min_max_table: (string * thm) list type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list val add_extra_norm: SMT_Util.class * extra_norm -> Context.generic -> Context.generic val normalize: Proof.context -> thm list -> (int * thm) list end; structure SMT_Normalize: SMT_NORMALIZE = struct fun drop_fact_warning ctxt = SMT_Config.verbose_msg ctxt (prefix "Warning: dropping assumption: " o Thm.string_of_thm ctxt) (* general theorem normalizations *) (** instantiate elimination rules **) local val (cpfalse, cfalse) = `SMT_Util.mk_cprop (Thm.cterm_of \<^context> \<^const>\False\) fun inst f ct thm = let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm)) in Thm.instantiate ([], [(dest_Var (Thm.term_of cv), ct)]) thm end in fun instantiate_elim thm = (case Thm.concl_of thm of \<^const>\Trueprop\ $ Var (_, \<^typ>\bool\) => inst Thm.dest_arg cfalse thm | Var _ => inst I cpfalse thm | _ => thm) end (** normalize definitions **) fun norm_def thm = (case Thm.prop_of thm of \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ Abs _) => norm_def (thm RS @{thm fun_cong}) | Const (\<^const_name>\Pure.eq\, _) $ _ $ Abs _ => norm_def (HOLogic.mk_obj_eq thm) | _ => thm) (** atomization **) fun atomize_conv ctxt ct = (case Thm.term_of ct of \<^const>\Pure.imp\ $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp} | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq} | Const (\<^const_name>\Pure.all\, _) $ Abs _ => Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all} | _ => Conv.all_conv) ct handle CTERM _ => Conv.all_conv ct val setup_atomize = fold SMT_Builtin.add_builtin_fun_ext'' [\<^const_name>\Pure.imp\, \<^const_name>\Pure.eq\, \<^const_name>\Pure.all\, \<^const_name>\Trueprop\] (** unfold special quantifiers **) val special_quant_table = [ (\<^const_name>\Ex1\, @{thm Ex1_def_raw}), (\<^const_name>\Ball\, @{thm Ball_def_raw}), (\<^const_name>\Bex\, @{thm Bex_def_raw})] local fun special_quant (Const (n, _)) = AList.lookup (op =) special_quant_table n | special_quant _ = NONE fun special_quant_conv _ ct = (case special_quant (Thm.term_of ct) of SOME thm => Conv.rewr_conv thm | NONE => Conv.all_conv) ct in fun unfold_special_quants_conv ctxt = SMT_Util.if_exists_conv (is_some o special_quant) (Conv.top_conv special_quant_conv ctxt) val setup_unfolded_quants = fold (SMT_Builtin.add_builtin_fun_ext'' o fst) special_quant_table end (** trigger inference **) local (*** check trigger syntax ***) fun dest_trigger (Const (\<^const_name>\pat\, _) $ _) = SOME true | dest_trigger (Const (\<^const_name>\nopat\, _) $ _) = SOME false | dest_trigger _ = NONE fun eq_list [] = false | eq_list (b :: bs) = forall (equal b) bs fun proper_trigger t = t |> these o try SMT_Util.dest_symb_list |> map (map_filter dest_trigger o these o try SMT_Util.dest_symb_list) |> (fn [] => false | bss => forall eq_list bss) fun proper_quant inside f t = (case t of Const (\<^const_name>\All\, _) $ Abs (_, _, u) => proper_quant true f u | Const (\<^const_name>\Ex\, _) $ Abs (_, _, u) => proper_quant true f u | \<^const>\trigger\ $ p $ u => (if inside then f p else false) andalso proper_quant false f u | Abs (_, _, u) => proper_quant false f u | u1 $ u2 => proper_quant false f u1 andalso proper_quant false f u2 | _ => true) fun check_trigger_error ctxt t = error ("SMT triggers must only occur under quantifier and multipatterns " ^ "must have the same kind: " ^ Syntax.string_of_term ctxt t) fun check_trigger_conv ctxt ct = if proper_quant false proper_trigger (SMT_Util.term_of ct) then Conv.all_conv ct else check_trigger_error ctxt (Thm.term_of ct) (*** infer simple triggers ***) fun dest_cond_eq ct = (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, _) $ _ $ _ => Thm.dest_binop ct | \<^const>\HOL.implies\ $ _ $ _ => dest_cond_eq (Thm.dest_arg ct) | _ => raise CTERM ("no equation", [ct])) fun get_constrs thy (Type (n, _)) = these (BNF_LFP_Compat.get_constrs thy n) | get_constrs _ _ = [] fun is_constr thy (n, T) = let fun match (m, U) = m = n andalso Sign.typ_instance thy (T, U) in can (the o find_first match o get_constrs thy o Term.body_type) T end fun is_constr_pat thy t = (case Term.strip_comb t of (Free _, []) => true | (Const c, ts) => is_constr thy c andalso forall (is_constr_pat thy) ts | _ => false) fun is_simp_lhs ctxt t = (case Term.strip_comb t of (Const c, ts as _ :: _) => not (SMT_Builtin.is_builtin_fun_ext ctxt c ts) andalso forall (is_constr_pat (Proof_Context.theory_of ctxt)) ts | _ => false) fun has_all_vars vs t = subset (op aconv) (vs, map Free (Term.add_frees t [])) fun minimal_pats vs ct = if has_all_vars vs (Thm.term_of ct) then (case Thm.term_of ct of _ $ _ => (case apply2 (minimal_pats vs) (Thm.dest_comb ct) of ([], []) => [[ct]] | (ctss, ctss') => union (eq_set (op aconvc)) ctss ctss') | _ => []) else [] fun proper_mpat _ _ _ [] = false | proper_mpat thy gen u cts = let val tps = (op ~~) (`gen (map Thm.term_of cts)) fun some_match u = tps |> exists (fn (t', t) => Pattern.matches thy (t', u) andalso not (t aconv u)) in not (Term.exists_subterm some_match u) end val pat = SMT_Util.mk_const_pat \<^theory> \<^const_name>\pat\ Thm.dest_ctyp0 fun mk_pat ct = Thm.apply (SMT_Util.instT' ct pat) ct fun mk_clist T = apply2 (Thm.cterm_of \<^context>) (SMT_Util.symb_cons_const T, SMT_Util.symb_nil_const T) fun mk_list (ccons, cnil) f cts = fold_rev (Thm.mk_binop ccons o f) cts cnil val mk_pat_list = mk_list (mk_clist \<^typ>\pattern\) val mk_mpat_list = mk_list (mk_clist \<^typ>\pattern symb_list\) fun mk_trigger ctss = mk_mpat_list (mk_pat_list mk_pat) ctss val trigger_eq = mk_meta_eq @{lemma "p = trigger t p" by (simp add: trigger_def)} fun insert_trigger_conv [] ct = Conv.all_conv ct | insert_trigger_conv ctss ct = let val (ctr, cp) = Thm.dest_binop (Thm.rhs_of trigger_eq) ||> rpair ct val inst = map (apfst (dest_Var o Thm.term_of)) [cp, (ctr, mk_trigger ctss)] in Thm.instantiate ([], inst) trigger_eq end fun infer_trigger_eq_conv outer_ctxt (ctxt, cvs) ct = let val (lhs, rhs) = dest_cond_eq ct val vs = map Thm.term_of cvs val thy = Proof_Context.theory_of ctxt fun get_mpats ct = if is_simp_lhs ctxt (Thm.term_of ct) then minimal_pats vs ct else [] val gen = Variable.export_terms ctxt outer_ctxt val filter_mpats = filter (proper_mpat thy gen (Thm.term_of rhs)) in insert_trigger_conv (filter_mpats (get_mpats lhs)) ct end fun has_trigger (\<^const>\trigger\ $ _ $ _) = true | has_trigger _ = false fun try_trigger_conv cv ct = if SMT_Util.under_quant has_trigger (SMT_Util.term_of ct) then Conv.all_conv ct else Conv.try_conv cv ct fun infer_trigger_conv ctxt = if Config.get ctxt SMT_Config.infer_triggers then try_trigger_conv (SMT_Util.under_quant_conv (infer_trigger_eq_conv ctxt) ctxt) else Conv.all_conv in fun trigger_conv ctxt = SMT_Util.prop_conv (check_trigger_conv ctxt then_conv infer_trigger_conv ctxt) val setup_trigger = fold SMT_Builtin.add_builtin_fun_ext'' [\<^const_name>\pat\, \<^const_name>\nopat\, \<^const_name>\trigger\] end (** combined general normalizations **) fun gen_normalize1_conv ctxt = atomize_conv ctxt then_conv unfold_special_quants_conv ctxt then_conv Thm.beta_conversion true then_conv trigger_conv ctxt fun gen_normalize1 ctxt = instantiate_elim #> norm_def #> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion) #> - Drule.forall_intr_vars #> + Thm.forall_intr_vars #> Conv.fconv_rule (gen_normalize1_conv ctxt) #> (* Z3 4.3.1 silently normalizes "P --> Q --> R" to "P & Q --> R" *) Raw_Simplifier.rewrite_rule ctxt @{thms HOL.imp_conjL[symmetric, THEN eq_reflection]} fun gen_norm1_safe ctxt (i, thm) = (case try (gen_normalize1 ctxt) thm of SOME thm' => SOME (i, thm') | NONE => (drop_fact_warning ctxt thm; NONE)) fun gen_normalize ctxt iwthms = map_filter (gen_norm1_safe ctxt) iwthms (* unfolding of definitions and theory-specific rewritings *) fun expand_head_conv cv ct = (case Thm.term_of ct of _ $ _ => Conv.fun_conv (expand_head_conv cv) then_conv Conv.try_conv (Thm.beta_conversion false) | _ => cv) ct (** rewrite bool case expressions as if expressions **) val case_bool_entry = (\<^const_name>\bool.case_bool\, @{thm case_bool_if}) local fun is_case_bool (Const (\<^const_name>\bool.case_bool\, _)) = true | is_case_bool _ = false fun unfold_conv _ = SMT_Util.if_true_conv (is_case_bool o Term.head_of) (expand_head_conv (Conv.rewr_conv @{thm case_bool_if})) in fun rewrite_case_bool_conv ctxt = SMT_Util.if_exists_conv is_case_bool (Conv.top_conv unfold_conv ctxt) val setup_case_bool = SMT_Builtin.add_builtin_fun_ext'' \<^const_name>\bool.case_bool\ end (** unfold abs, min and max **) val abs_min_max_table = [ (\<^const_name>\min\, @{thm min_def_raw}), (\<^const_name>\max\, @{thm max_def_raw}), (\<^const_name>\abs\, @{thm abs_if_raw})] local fun abs_min_max ctxt (Const (n, Type (\<^type_name>\fun\, [T, _]))) = (case AList.lookup (op =) abs_min_max_table n of NONE => NONE | SOME thm => if SMT_Builtin.is_builtin_typ_ext ctxt T then SOME thm else NONE) | abs_min_max _ _ = NONE fun unfold_amm_conv ctxt ct = (case abs_min_max ctxt (Term.head_of (Thm.term_of ct)) of SOME thm => expand_head_conv (Conv.rewr_conv thm) | NONE => Conv.all_conv) ct in fun unfold_abs_min_max_conv ctxt = SMT_Util.if_exists_conv (is_some o abs_min_max ctxt) (Conv.top_conv unfold_amm_conv ctxt) val setup_abs_min_max = fold (SMT_Builtin.add_builtin_fun_ext'' o fst) abs_min_max_table end (** embedding of standard natural number operations into integer operations **) local val simple_nat_ops = [ @{const HOL.eq (nat)}, @{const less (nat)}, @{const less_eq (nat)}, \<^const>\Suc\, @{const plus (nat)}, @{const minus (nat)}] val nat_consts = simple_nat_ops @ [@{const numeral (nat)}, @{const zero_class.zero (nat)}, @{const one_class.one (nat)}] @ [@{const times (nat)}, @{const divide (nat)}, @{const modulo (nat)}] val is_nat_const = member (op aconv) nat_consts val nat_int_thm = Thm.symmetric (mk_meta_eq @{thm nat_int}) val nat_int_comp_thms = map mk_meta_eq @{thms nat_int_comparison} val int_ops_thms = map mk_meta_eq @{thms int_ops} val int_if_thm = mk_meta_eq @{thm int_if} fun if_conv cv1 cv2 = Conv.combination_conv (Conv.combination_conv (Conv.arg_conv cv1) cv2) cv2 fun int_ops_conv cv ctxt ct = (case Thm.term_of ct of @{const of_nat (int)} $ (Const (\<^const_name>\If\, _) $ _ $ _ $ _) => Conv.rewr_conv int_if_thm then_conv if_conv (cv ctxt) (int_ops_conv cv ctxt) | @{const of_nat (int)} $ _ => (Conv.rewrs_conv int_ops_thms then_conv Conv.top_sweep_conv (int_ops_conv cv) ctxt) else_conv Conv.arg_conv (Conv.sub_conv cv ctxt) | _ => Conv.no_conv) ct val unfold_nat_let_conv = Conv.rewr_conv @{lemma "Let (n::nat) f \ f n" by (rule Let_def)} val drop_nat_int_conv = Conv.rewr_conv (Thm.symmetric nat_int_thm) fun nat_to_int_conv ctxt ct = ( Conv.top_conv (K (Conv.try_conv unfold_nat_let_conv)) ctxt then_conv Conv.top_sweep_conv nat_conv ctxt then_conv Conv.top_conv (K (Conv.try_conv drop_nat_int_conv)) ctxt) ct and nat_conv ctxt ct = ( Conv.rewrs_conv (nat_int_thm :: nat_int_comp_thms) then_conv Conv.top_sweep_conv (int_ops_conv nat_to_int_conv) ctxt) ct fun add_int_of_nat vs ct cu (q, cts) = (case Thm.term_of ct of @{const of_nat(int)} => if Term.exists_subterm (member (op aconv) vs) (Thm.term_of cu) then (true, cts) else (q, insert (op aconvc) cu cts) | _ => (q, cts)) fun add_apps f vs ct = (case Thm.term_of ct of _ $ _ => let val (cu1, cu2) = Thm.dest_comb ct in f vs cu1 cu2 #> add_apps f vs cu1 #> add_apps f vs cu2 end | Abs _ => let val (cv, cu) = Thm.dest_abs NONE ct in add_apps f (Thm.term_of cv :: vs) cu end | _ => I) val int_thm = @{lemma "(0::int) <= int (n::nat)" by simp} val nat_int_thms = @{lemma "\n::nat. (0::int) <= int n" "\n::nat. nat (int n) = n" "\i::int. int (nat i) = (if 0 <= i then i else 0)" by simp_all} val var = Term.dest_Var (Thm.term_of (funpow 3 Thm.dest_arg (Thm.cprop_of int_thm))) in fun nat_as_int_conv ctxt = SMT_Util.if_exists_conv is_nat_const (nat_to_int_conv ctxt) fun add_int_of_nat_constraints thms = let val (q, cts) = fold (add_apps add_int_of_nat [] o Thm.cprop_of) thms (false, []) in if q then (thms, nat_int_thms) else (thms, map (fn ct => Thm.instantiate ([], [(var, ct)]) int_thm) cts) end val setup_nat_as_int = SMT_Builtin.add_builtin_typ_ext (\<^typ>\nat\, fn ctxt => K (Config.get ctxt SMT_Config.nat_as_int)) #> fold (SMT_Builtin.add_builtin_fun_ext' o Term.dest_Const) simple_nat_ops end (** normalize numerals **) local (* rewrite Numeral1 into 1 rewrite - 0 into 0 *) fun is_irregular_number (Const (\<^const_name>\numeral\, _) $ Const (\<^const_name>\num.One\, _)) = true | is_irregular_number (Const (\<^const_name>\uminus\, _) $ Const (\<^const_name>\Groups.zero\, _)) = true | is_irregular_number _ = false fun is_strange_number ctxt t = is_irregular_number t andalso SMT_Builtin.is_builtin_num ctxt t val proper_num_ss = simpset_of (put_simpset HOL_ss \<^context> addsimps @{thms Num.numeral_One minus_zero}) fun norm_num_conv ctxt = SMT_Util.if_conv (is_strange_number ctxt) (Simplifier.rewrite (put_simpset proper_num_ss ctxt)) Conv.no_conv in fun normalize_numerals_conv ctxt = SMT_Util.if_exists_conv (is_strange_number ctxt) (Conv.top_sweep_conv norm_num_conv ctxt) end (** combined unfoldings and rewritings **) fun burrow_ids f ithms = let val (is, thms) = split_list ithms val (thms', extra_thms) = f thms in (is ~~ thms') @ map (pair ~1) extra_thms end fun unfold_conv ctxt = rewrite_case_bool_conv ctxt then_conv unfold_abs_min_max_conv ctxt then_conv (if Config.get ctxt SMT_Config.nat_as_int then nat_as_int_conv ctxt else Conv.all_conv) then_conv Thm.beta_conversion true fun unfold_polymorph ctxt = map (apsnd (Conv.fconv_rule (unfold_conv ctxt))) fun unfold_monomorph ctxt = map (apsnd (Conv.fconv_rule (normalize_numerals_conv ctxt))) #> Config.get ctxt SMT_Config.nat_as_int ? burrow_ids add_int_of_nat_constraints (* overall normalization *) type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list structure Extra_Norms = Generic_Data ( type T = extra_norm SMT_Util.dict val empty = [] val extend = I fun merge data = SMT_Util.dict_merge fst data ) fun add_extra_norm (cs, norm) = Extra_Norms.map (SMT_Util.dict_update (cs, norm)) fun apply_extra_norms ctxt ithms = let val cs = SMT_Config.solver_class_of ctxt val es = SMT_Util.dict_lookup (Extra_Norms.get (Context.Proof ctxt)) cs in burrow_ids (fold (fn e => e ctxt) es o rpair []) ithms end local val ignored = member (op =) [\<^const_name>\All\, \<^const_name>\Ex\, \<^const_name>\Let\, \<^const_name>\If\, \<^const_name>\HOL.eq\] val schematic_consts_of = let fun collect (\<^const>\trigger\ $ p $ t) = collect_trigger p #> collect t | collect (t $ u) = collect t #> collect u | collect (Abs (_, _, t)) = collect t | collect (t as Const (n, _)) = if not (ignored n) then Monomorph.add_schematic_consts_of t else I | collect _ = I and collect_trigger t = let val dest = these o try SMT_Util.dest_symb_list in fold (fold collect_pat o dest) (dest t) end and collect_pat (Const (\<^const_name>\pat\, _) $ t) = collect t | collect_pat (Const (\<^const_name>\nopat\, _) $ t) = collect t | collect_pat _ = I in (fn t => collect t Symtab.empty) end in fun monomorph ctxt xthms = let val (xs, thms) = split_list xthms in map (pair 1) thms |> Monomorph.monomorph schematic_consts_of ctxt |> maps (uncurry (map o pair)) o map2 pair xs o map (map snd) end end fun normalize ctxt wthms = wthms |> map_index I |> gen_normalize ctxt |> unfold_polymorph ctxt |> monomorph ctxt |> unfold_monomorph ctxt |> apply_extra_norms ctxt val _ = Theory.setup (Context.theory_map ( setup_atomize #> setup_unfolded_quants #> setup_trigger #> setup_case_bool #> setup_abs_min_max #> setup_nat_as_int)) end; diff --git a/src/HOL/Tools/Transfer/transfer.ML b/src/HOL/Tools/Transfer/transfer.ML --- a/src/HOL/Tools/Transfer/transfer.ML +++ b/src/HOL/Tools/Transfer/transfer.ML @@ -1,932 +1,932 @@ (* Title: HOL/Tools/Transfer/transfer.ML Author: Brian Huffman, TU Muenchen Author: Ondrej Kuncar, TU Muenchen Generic theorem transfer method. *) signature TRANSFER = sig type pred_data val mk_pred_data: thm -> thm -> thm list -> pred_data val rel_eq_onp: pred_data -> thm val pred_def: pred_data -> thm val pred_simps: pred_data -> thm list val update_pred_simps: thm list -> pred_data -> pred_data val bottom_rewr_conv: thm list -> conv val top_rewr_conv: thm list -> conv val top_sweep_rewr_conv: thm list -> conv val prep_conv: conv val fold_relator_eqs_conv: Proof.context -> conv val unfold_relator_eqs_conv: Proof.context -> conv val get_transfer_raw: Proof.context -> thm list val get_relator_eq: Proof.context -> thm list val retrieve_relator_eq: Proof.context -> term -> thm list val get_sym_relator_eq: Proof.context -> thm list val get_relator_eq_raw: Proof.context -> thm list val get_relator_domain: Proof.context -> thm list val morph_pred_data: morphism -> pred_data -> pred_data val lookup_pred_data: Proof.context -> string -> pred_data option val update_pred_data: string -> pred_data -> Context.generic -> Context.generic val is_compound_lhs: Proof.context -> term -> bool val is_compound_rhs: Proof.context -> term -> bool val transfer_add: attribute val transfer_del: attribute val transfer_raw_add: thm -> Context.generic -> Context.generic val transfer_raw_del: thm -> Context.generic -> Context.generic val transferred_attribute: thm list -> attribute val untransferred_attribute: thm list -> attribute val prep_transfer_domain_thm: Proof.context -> thm -> thm val transfer_domain_add: attribute val transfer_domain_del: attribute val transfer_rule_of_term: Proof.context -> bool -> term -> thm val transfer_rule_of_lhs: Proof.context -> term -> thm val eq_tac: Proof.context -> int -> tactic val transfer_start_tac: bool -> Proof.context -> int -> tactic val transfer_prover_start_tac: Proof.context -> int -> tactic val transfer_step_tac: Proof.context -> int -> tactic val transfer_end_tac: Proof.context -> int -> tactic val transfer_prover_end_tac: Proof.context -> int -> tactic val transfer_tac: bool -> Proof.context -> int -> tactic val transfer_prover_tac: Proof.context -> int -> tactic val gen_frees_tac: (string * typ) list -> Proof.context -> int -> tactic end structure Transfer : TRANSFER = struct fun bottom_rewr_conv rewrs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv rewrs))) \<^context> fun top_rewr_conv rewrs = Conv.top_conv (K (Conv.try_conv (Conv.rewrs_conv rewrs))) \<^context> fun top_sweep_rewr_conv rewrs = Conv.top_sweep_conv (K (Conv.rewrs_conv rewrs)) \<^context> (** Theory Data **) val compound_xhs_empty_net = Item_Net.init (Thm.eq_thm_prop o apply2 snd) (single o fst); val rewr_rules = Item_Net.init Thm.eq_thm_prop (single o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.concl_of); datatype pred_data = PRED_DATA of {pred_def: thm, rel_eq_onp: thm, pred_simps: thm list} fun mk_pred_data pred_def rel_eq_onp pred_simps = PRED_DATA {pred_def = pred_def, rel_eq_onp = rel_eq_onp, pred_simps = pred_simps} fun map_pred_data' f1 f2 f3 (PRED_DATA {pred_def, rel_eq_onp, pred_simps}) = PRED_DATA {pred_def = f1 pred_def, rel_eq_onp = f2 rel_eq_onp, pred_simps = f3 pred_simps} fun rep_pred_data (PRED_DATA p) = p val rel_eq_onp = #rel_eq_onp o rep_pred_data val pred_def = #pred_def o rep_pred_data val pred_simps = #pred_simps o rep_pred_data fun update_pred_simps new_pred_data = map_pred_data' I I (K new_pred_data) structure Data = Generic_Data ( type T = { transfer_raw : thm Item_Net.T, known_frees : (string * typ) list, compound_lhs : (term * thm) Item_Net.T, compound_rhs : (term * thm) Item_Net.T, relator_eq : thm Item_Net.T, relator_eq_raw : thm Item_Net.T, relator_domain : thm Item_Net.T, pred_data : pred_data Symtab.table } val empty = { transfer_raw = Thm.item_net_intro, known_frees = [], compound_lhs = compound_xhs_empty_net, compound_rhs = compound_xhs_empty_net, relator_eq = rewr_rules, relator_eq_raw = Thm.item_net, relator_domain = Thm.item_net, pred_data = Symtab.empty } val extend = I fun merge ( { transfer_raw = t1, known_frees = k1, compound_lhs = l1, compound_rhs = c1, relator_eq = r1, relator_eq_raw = rw1, relator_domain = rd1, pred_data = pd1 }, { transfer_raw = t2, known_frees = k2, compound_lhs = l2, compound_rhs = c2, relator_eq = r2, relator_eq_raw = rw2, relator_domain = rd2, pred_data = pd2 } ) = { transfer_raw = Item_Net.merge (t1, t2), known_frees = Library.merge (op =) (k1, k2), compound_lhs = Item_Net.merge (l1, l2), compound_rhs = Item_Net.merge (c1, c2), relator_eq = Item_Net.merge (r1, r2), relator_eq_raw = Item_Net.merge (rw1, rw2), relator_domain = Item_Net.merge (rd1, rd2), pred_data = Symtab.merge (K true) (pd1, pd2) } ) fun get_net_content f ctxt = Item_Net.content (f (Data.get (Context.Proof ctxt))) |> map (Thm.transfer' ctxt) val get_transfer_raw = get_net_content #transfer_raw val get_known_frees = #known_frees o Data.get o Context.Proof fun is_compound f ctxt t = Item_Net.retrieve (f (Data.get (Context.Proof ctxt))) t |> exists (fn (pat, _) => Pattern.matches (Proof_Context.theory_of ctxt) (pat, t)); val is_compound_lhs = is_compound #compound_lhs val is_compound_rhs = is_compound #compound_rhs val get_relator_eq = get_net_content #relator_eq #> map safe_mk_meta_eq fun retrieve_relator_eq ctxt t = Item_Net.retrieve (#relator_eq (Data.get (Context.Proof ctxt))) t |> map (Thm.transfer' ctxt) val get_sym_relator_eq = get_net_content #relator_eq #> map (safe_mk_meta_eq #> Thm.symmetric) val get_relator_eq_raw = get_net_content #relator_eq_raw val get_relator_domain = get_net_content #relator_domain val get_pred_data = #pred_data o Data.get o Context.Proof fun map_data f1 f2 f3 f4 f5 f6 f7 f8 { transfer_raw, known_frees, compound_lhs, compound_rhs, relator_eq, relator_eq_raw, relator_domain, pred_data } = { transfer_raw = f1 transfer_raw, known_frees = f2 known_frees, compound_lhs = f3 compound_lhs, compound_rhs = f4 compound_rhs, relator_eq = f5 relator_eq, relator_eq_raw = f6 relator_eq_raw, relator_domain = f7 relator_domain, pred_data = f8 pred_data } fun map_transfer_raw f = map_data f I I I I I I I fun map_known_frees f = map_data I f I I I I I I fun map_compound_lhs f = map_data I I f I I I I I fun map_compound_rhs f = map_data I I I f I I I I fun map_relator_eq f = map_data I I I I f I I I fun map_relator_eq_raw f = map_data I I I I I f I I fun map_relator_domain f = map_data I I I I I I f I fun map_pred_data f = map_data I I I I I I I f val add_transfer_thm = Thm.trim_context #> (fn thm => Data.map (map_transfer_raw (Item_Net.update thm) o map_compound_lhs (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Rel\, _) $ _ $ (lhs as (_ $ _)) $ _ => Item_Net.update (lhs, thm) | _ => I) o map_compound_rhs (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Rel\, _) $ _ $ _ $ (rhs as (_ $ _)) => Item_Net.update (rhs, thm) | _ => I) o map_known_frees (Term.add_frees (Thm.concl_of thm)))) fun del_transfer_thm thm = Data.map (map_transfer_raw (Item_Net.remove thm) o map_compound_lhs (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Rel\, _) $ _ $ (lhs as (_ $ _)) $ _ => Item_Net.remove (lhs, thm) | _ => I) o map_compound_rhs (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Rel\, _) $ _ $ _ $ (rhs as (_ $ _)) => Item_Net.remove (rhs, thm) | _ => I)) fun transfer_raw_add thm ctxt = add_transfer_thm thm ctxt fun transfer_raw_del thm ctxt = del_transfer_thm thm ctxt (** Conversions **) fun transfer_rel_conv conv = Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.fun2_conv (Conv.arg_conv conv))) val Rel_rule = Thm.symmetric @{thm Rel_def} fun Rel_conv ct = let val (cT, cT') = Thm.dest_funT (Thm.ctyp_of_cterm ct) val (cU, _) = Thm.dest_funT cT' in Thm.instantiate' [SOME cT, SOME cU] [SOME ct] Rel_rule end (* Conversion to preprocess a transfer rule *) fun safe_Rel_conv ct = Conv.try_conv (HOLogic.Trueprop_conv (Conv.fun_conv (Conv.fun_conv Rel_conv))) ct fun prep_conv ct = ( Conv.implies_conv safe_Rel_conv prep_conv else_conv safe_Rel_conv else_conv Conv.all_conv) ct fun fold_relator_eqs_conv ctxt ct = (bottom_rewr_conv (get_relator_eq ctxt)) ct; fun unfold_relator_eqs_conv ctxt ct = (top_rewr_conv (get_sym_relator_eq ctxt)) ct; (** Replacing explicit equalities with is_equality premises **) fun mk_is_equality t = Const (\<^const_name>\is_equality\, Term.fastype_of t --> HOLogic.boolT) $ t fun gen_abstract_equalities ctxt (dest : term -> term * (term -> term)) thm = let val prop = Thm.prop_of thm val (t, mk_prop') = dest prop (* Only consider "(=)" at non-base types *) fun is_eq (Const (\<^const_name>\HOL.eq\, Type ("fun", [T, _]))) = (case T of Type (_, []) => false | _ => true) | is_eq _ = false val add_eqs = Term.fold_aterms (fn t => if is_eq t then insert (op =) t else I) val eq_consts = rev (add_eqs t []) val eqTs = map (snd o dest_Const) eq_consts val used = Term.add_free_names prop [] val names = map (K "") eqTs |> Name.variant_list used val frees = map Free (names ~~ eqTs) val prems = map (HOLogic.mk_Trueprop o mk_is_equality) frees val prop1 = mk_prop' (Term.subst_atomic (eq_consts ~~ frees) t) val prop2 = fold Logic.all frees (Logic.list_implies (prems, prop1)) val cprop = Thm.cterm_of ctxt prop2 val equal_thm = Raw_Simplifier.rewrite ctxt false @{thms is_equality_lemma} cprop fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm in forall_elim (thm COMP (equal_thm COMP @{thm equal_elim_rule2})) end handle TERM _ => thm fun abstract_equalities_transfer ctxt thm = let fun dest prop = let val prems = Logic.strip_imp_prems prop val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl) in (rel, fn rel' => Logic.list_implies (prems, HOLogic.mk_Trueprop (rel' $ x $ y))) end val contracted_eq_thm = Conv.fconv_rule (transfer_rel_conv (fold_relator_eqs_conv ctxt)) thm handle CTERM _ => thm in gen_abstract_equalities ctxt dest contracted_eq_thm end fun abstract_equalities_relator_eq ctxt rel_eq_thm = gen_abstract_equalities ctxt (fn x => (x, I)) (rel_eq_thm RS @{thm is_equality_def [THEN iffD2]}) fun abstract_equalities_domain ctxt thm = let fun dest prop = let val prems = Logic.strip_imp_prems prop val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) val ((eq, dom), y) = apfst Term.dest_comb (Term.dest_comb concl) in (dom, fn dom' => Logic.list_implies (prems, HOLogic.mk_Trueprop (eq $ dom' $ y))) end fun transfer_rel_conv conv = Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.arg1_conv (Conv.arg_conv conv))) val contracted_eq_thm = Conv.fconv_rule (transfer_rel_conv (fold_relator_eqs_conv ctxt)) thm in gen_abstract_equalities ctxt dest contracted_eq_thm end (** Replacing explicit Domainp predicates with Domainp assumptions **) fun mk_Domainp_assm (T, R) = HOLogic.mk_eq ((Const (\<^const_name>\Domainp\, Term.fastype_of T --> Term.fastype_of R) $ T), R) fun fold_Domainp f (t as Const (\<^const_name>\Domainp\,_) $ (Var (_,_))) = f t | fold_Domainp f (t $ u) = fold_Domainp f t #> fold_Domainp f u | fold_Domainp f (Abs (_, _, t)) = fold_Domainp f t | fold_Domainp _ _ = I fun subst_terms tab t = let val t' = Termtab.lookup tab t in case t' of SOME t' => t' | NONE => (case t of u $ v => (subst_terms tab u) $ (subst_terms tab v) | Abs (a, T, t) => Abs (a, T, subst_terms tab t) | t => t) end fun gen_abstract_domains ctxt (dest : term -> term * (term -> term)) thm = let val prop = Thm.prop_of thm val (t, mk_prop') = dest prop val Domainp_tms = rev (fold_Domainp (fn t => insert op= t) t []) val Domainp_Ts = map (snd o dest_funT o snd o dest_Const o fst o dest_comb) Domainp_tms val used = Term.add_free_names t [] val rels = map (snd o dest_comb) Domainp_tms val rel_names = map (fst o fst o dest_Var) rels val names = map (fn name => ("D" ^ name)) rel_names |> Name.variant_list used val frees = map Free (names ~~ Domainp_Ts) val prems = map (HOLogic.mk_Trueprop o mk_Domainp_assm) (rels ~~ frees); val t' = subst_terms (fold Termtab.update (Domainp_tms ~~ frees) Termtab.empty) t val prop1 = fold Logic.all frees (Logic.list_implies (prems, mk_prop' t')) val prop2 = Logic.list_rename_params (rev names) prop1 val cprop = Thm.cterm_of ctxt prop2 val equal_thm = Raw_Simplifier.rewrite ctxt false @{thms Domainp_lemma} cprop fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm; in forall_elim (thm COMP (equal_thm COMP @{thm equal_elim_rule2})) end handle TERM _ => thm fun abstract_domains_transfer ctxt thm = let fun dest prop = let val prems = Logic.strip_imp_prems prop val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl) in (x, fn x' => Logic.list_implies (prems, HOLogic.mk_Trueprop (rel $ x' $ y))) end in gen_abstract_domains ctxt dest thm end fun abstract_domains_relator_domain ctxt thm = let fun dest prop = let val prems = Logic.strip_imp_prems prop val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl) in (y, fn y' => Logic.list_implies (prems, HOLogic.mk_Trueprop (rel $ x $ y'))) end in gen_abstract_domains ctxt dest thm end fun detect_transfer_rules thm = let fun is_transfer_rule tm = case (HOLogic.dest_Trueprop tm) of (Const (\<^const_name>\HOL.eq\, _)) $ ((Const (\<^const_name>\Domainp\, _)) $ _) $ _ => false | _ $ _ $ _ => true | _ => false fun safe_transfer_rule_conv ctm = if is_transfer_rule (Thm.term_of ctm) then safe_Rel_conv ctm else Conv.all_conv ctm in Conv.fconv_rule (Conv.prems_conv ~1 safe_transfer_rule_conv) thm end (** Adding transfer domain rules **) fun prep_transfer_domain_thm ctxt = abstract_equalities_domain ctxt o detect_transfer_rules fun add_transfer_domain_thm thm ctxt = (add_transfer_thm o prep_transfer_domain_thm (Context.proof_of ctxt)) thm ctxt fun del_transfer_domain_thm thm ctxt = (del_transfer_thm o prep_transfer_domain_thm (Context.proof_of ctxt)) thm ctxt (** Transfer proof method **) val post_simps = @{thms transfer_forall_eq [symmetric] transfer_implies_eq [symmetric] transfer_bforall_unfold} fun gen_frees_tac keepers ctxt = SUBGOAL (fn (t, i) => let val keepers = keepers @ get_known_frees ctxt val vs = rev (Term.add_frees t []) val vs' = filter_out (member (op =) keepers) vs in Induct.arbitrary_tac ctxt 0 vs' i end) fun mk_relT (T, U) = T --> U --> HOLogic.boolT fun mk_Rel t = let val T = fastype_of t in Const (\<^const_name>\Transfer.Rel\, T --> T) $ t end fun transfer_rule_of_terms (prj : typ * typ -> typ) ctxt tab t u = let (* precondition: prj(T,U) must consist of only TFrees and type "fun" *) fun rel (T as Type ("fun", [T1, T2])) (U as Type ("fun", [U1, U2])) = let val r1 = rel T1 U1 val r2 = rel T2 U2 val rT = fastype_of r1 --> fastype_of r2 --> mk_relT (T, U) in Const (\<^const_name>\rel_fun\, rT) $ r1 $ r2 end | rel T U = let val (a, _) = dest_TFree (prj (T, U)) in Free (the (AList.lookup (op =) tab a), mk_relT (T, U)) end fun zip _ thms (Bound i) (Bound _) = (nth thms i, []) | zip ctxt thms (Abs (x, T, t)) (Abs (y, U, u)) = let val ([x', y'], ctxt') = Variable.variant_fixes [x, y] ctxt val prop = mk_Rel (rel T U) $ Free (x', T) $ Free (y', U) val cprop = Thm.cterm_of ctxt' (HOLogic.mk_Trueprop prop) val thm0 = Thm.assume cprop val (thm1, hyps) = zip ctxt' (thm0 :: thms) t u val ((r1, x), y) = apfst Thm.dest_comb (Thm.dest_comb (Thm.dest_arg cprop)) val r2 = Thm.dest_fun2 (Thm.dest_arg (Thm.cprop_of thm1)) val (a1, (b1, _)) = apsnd Thm.dest_funT (Thm.dest_funT (Thm.ctyp_of_cterm r1)) val (a2, (b2, _)) = apsnd Thm.dest_funT (Thm.dest_funT (Thm.ctyp_of_cterm r2)) val tinsts = [SOME a1, SOME b1, SOME a2, SOME b2] val insts = [SOME (Thm.dest_arg r1), SOME (Thm.dest_arg r2)] val rule = Thm.instantiate' tinsts insts @{thm Rel_abs} val thm2 = Thm.forall_intr x (Thm.forall_intr y (Thm.implies_intr cprop thm1)) in (thm2 COMP rule, hyps) end | zip ctxt thms (f $ t) (g $ u) = let val (thm1, hyps1) = zip ctxt thms f g val (thm2, hyps2) = zip ctxt thms t u in (thm2 RS (thm1 RS @{thm Rel_app}), hyps1 @ hyps2) end | zip _ _ t u = let val T = fastype_of t val U = fastype_of u val prop = mk_Rel (rel T U) $ t $ u val cprop = Thm.cterm_of ctxt (HOLogic.mk_Trueprop prop) in (Thm.assume cprop, [cprop]) end val r = mk_Rel (rel (fastype_of t) (fastype_of u)) val goal = HOLogic.mk_Trueprop (r $ t $ u) val rename = Thm.trivial (Thm.cterm_of ctxt goal) val (thm, hyps) = zip ctxt [] t u in Drule.implies_intr_list hyps (thm RS rename) end (* create a lambda term of the same shape as the given term *) fun skeleton is_atom = let fun dummy ctxt = let val (c, ctxt') = yield_singleton Variable.variant_fixes "a" ctxt in (Free (c, dummyT), ctxt') end fun skel (Bound i) ctxt = (Bound i, ctxt) | skel (Abs (x, _, t)) ctxt = let val (t', ctxt) = skel t ctxt in (Abs (x, dummyT, t'), ctxt) end | skel (tu as t $ u) ctxt = if is_atom tu andalso not (Term.is_open tu) then dummy ctxt else let val (t', ctxt) = skel t ctxt val (u', ctxt) = skel u ctxt in (t' $ u', ctxt) end | skel _ ctxt = dummy ctxt in fn ctxt => fn t => fst (skel t ctxt) |> Syntax.check_term ctxt (* FIXME avoid syntax operation!? *) end (** Monotonicity analysis **) (* TODO: Put extensible table in theory data *) val monotab = Symtab.make [(\<^const_name>\transfer_implies\, [~1, 1]), (\<^const_name>\transfer_forall\, [1])(*, (@{const_name implies}, [~1, 1]), (@{const_name All}, [1])*)] (* Function bool_insts determines the set of boolean-relation variables that can be instantiated to implies, rev_implies, or iff. Invariants: bool_insts p (t, u) requires that u :: _ => _ => ... => bool, and t is a skeleton of u *) fun bool_insts p (t, u) = let fun strip2 (t1 $ t2, u1 $ u2, tus) = strip2 (t1, u1, (t2, u2) :: tus) | strip2 x = x fun or3 ((a, b, c), (x, y, z)) = (a orelse x, b orelse y, c orelse z) fun go Ts p (Abs (_, T, t), Abs (_, _, u)) tab = go (T :: Ts) p (t, u) tab | go Ts p (t, u) tab = let val (a, _) = dest_TFree (Term.body_type (Term.fastype_of1 (Ts, t))) val (_, tf, tus) = strip2 (t, u, []) val ps_opt = case tf of Const (c, _) => Symtab.lookup monotab c | _ => NONE val tab1 = case ps_opt of SOME ps => let val ps' = map (fn x => p * x) (take (length tus) ps) in fold I (map2 (go Ts) ps' tus) tab end | NONE => tab val tab2 = Symtab.make [(a, (p >= 0, p <= 0, is_none ps_opt))] in Symtab.join (K or3) (tab1, tab2) end val tab = go [] p (t, u) Symtab.empty fun f (a, (true, false, false)) = SOME (a, \<^const>\implies\) | f (a, (false, true, false)) = SOME (a, \<^const>\rev_implies\) | f (a, (true, true, _)) = SOME (a, HOLogic.eq_const HOLogic.boolT) | f _ = NONE in map_filter f (Symtab.dest tab) end fun transfer_rule_of_term ctxt equiv t = let val s = skeleton (is_compound_rhs ctxt) ctxt t val frees = map fst (Term.add_frees s []) val tfrees = map fst (Term.add_tfrees s []) fun prep a = "R" ^ Library.unprefix "'" a val (rnames, ctxt') = Variable.variant_fixes (map prep tfrees) ctxt val tab = tfrees ~~ rnames fun prep a = the (AList.lookup (op =) tab a) val thm = transfer_rule_of_terms fst ctxt' tab s t val binsts = bool_insts (if equiv then 0 else 1) (s, t) val idx = Thm.maxidx_of thm + 1 fun tinst (a, _) = (((a, idx), \<^sort>\type\), \<^ctyp>\bool\) fun inst (a, t) = ((Name.clean_index (prep a, idx), \<^typ>\bool \ bool \ bool\), Thm.cterm_of ctxt' t) in thm |> Thm.generalize (Symtab.make_set tfrees, Symtab.make_set (rnames @ frees)) idx |> Thm.instantiate (map tinst binsts, map inst binsts) end fun transfer_rule_of_lhs ctxt t = let val s = skeleton (is_compound_lhs ctxt) ctxt t val frees = map fst (Term.add_frees s []) val tfrees = map fst (Term.add_tfrees s []) fun prep a = "R" ^ Library.unprefix "'" a val (rnames, ctxt') = Variable.variant_fixes (map prep tfrees) ctxt val tab = tfrees ~~ rnames fun prep a = the (AList.lookup (op =) tab a) val thm = transfer_rule_of_terms snd ctxt' tab t s val binsts = bool_insts 1 (s, t) val idx = Thm.maxidx_of thm + 1 fun tinst (a, _) = (((a, idx), \<^sort>\type\), \<^ctyp>\bool\) fun inst (a, t) = ((Name.clean_index (prep a, idx), \<^typ>\bool \ bool \ bool\), Thm.cterm_of ctxt' t) in thm |> Thm.generalize (Symtab.make_set tfrees, Symtab.make_set (rnames @ frees)) idx |> Thm.instantiate (map tinst binsts, map inst binsts) end fun eq_rules_tac ctxt eq_rules = TRY o REPEAT_ALL_NEW (resolve_tac ctxt eq_rules) THEN_ALL_NEW resolve_tac ctxt @{thms is_equality_eq} fun eq_tac ctxt = eq_rules_tac ctxt (get_relator_eq_raw ctxt) fun transfer_step_tac ctxt = REPEAT_ALL_NEW (resolve_tac ctxt (get_transfer_raw ctxt)) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt (get_relator_eq_raw ctxt)) fun transfer_start_tac equiv ctxt i = let val pre_simps = @{thms transfer_forall_eq transfer_implies_eq} val start_rule = if equiv then @{thm transfer_start} else @{thm transfer_start'} val err_msg = "Transfer failed to convert goal to an object-logic formula" fun main_tac (t, i) = resolve_tac ctxt [start_rule] i THEN (resolve_tac ctxt [transfer_rule_of_term ctxt equiv (HOLogic.dest_Trueprop t)]) (i + 1) handle TERM (_, ts) => raise TERM (err_msg, ts) in EVERY [rewrite_goal_tac ctxt pre_simps i THEN SUBGOAL main_tac i] end; fun transfer_prover_start_tac ctxt = SUBGOAL (fn (t, i) => let val rhs = (snd o Term.dest_comb o HOLogic.dest_Trueprop) t val rule1 = transfer_rule_of_term ctxt false rhs val expand_eqs_in_rel_conv = transfer_rel_conv (unfold_relator_eqs_conv ctxt) in EVERY [CONVERSION prep_conv i, CONVERSION expand_eqs_in_rel_conv i, resolve_tac ctxt @{thms transfer_prover_start} i, resolve_tac ctxt [rule1] (i + 1)] end); fun transfer_end_tac ctxt i = let val post_simps = @{thms transfer_forall_eq [symmetric] transfer_implies_eq [symmetric] transfer_bforall_unfold} in EVERY [rewrite_goal_tac ctxt post_simps i, Goal.norm_hhf_tac ctxt i] end; fun transfer_prover_end_tac ctxt i = resolve_tac ctxt @{thms refl} i local infix 1 THEN_ALL_BUT_FIRST_NEW fun (tac1 THEN_ALL_BUT_FIRST_NEW tac2) i st = st |> (tac1 i THEN (fn st' => Seq.INTERVAL tac2 (i + 1) (i + Thm.nprems_of st' - Thm.nprems_of st) st')); in fun transfer_tac equiv ctxt i = let val rules = get_transfer_raw ctxt val eq_rules = get_relator_eq_raw ctxt (* allow unsolved subgoals only for standard transfer method, not for transfer' *) val end_tac = if equiv then K all_tac else K no_tac fun transfer_search_tac i = (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt eq_rules)) ORELSE' end_tac) i in (transfer_start_tac equiv ctxt THEN_ALL_BUT_FIRST_NEW transfer_search_tac THEN' transfer_end_tac ctxt) i end fun transfer_prover_tac ctxt i = let val rules = get_transfer_raw ctxt val eq_rules = get_relator_eq_raw ctxt fun transfer_prover_search_tac i = (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt eq_rules)) i in (transfer_prover_start_tac ctxt THEN_ALL_BUT_FIRST_NEW transfer_prover_search_tac THEN' transfer_prover_end_tac ctxt) i end end; (** Transfer attribute **) fun transferred ctxt extra_rules thm = let val rules = extra_rules @ get_transfer_raw ctxt val eq_rules = get_relator_eq_raw ctxt val pre_simps = @{thms transfer_forall_eq transfer_implies_eq} - val thm1 = Drule.forall_intr_vars thm + val thm1 = Thm.forall_intr_vars thm val instT = rev (Term.add_tvars (Thm.full_prop_of thm1) []) |> map (fn v as ((a, _), S) => (v, Thm.ctyp_of ctxt (TFree (a, S)))) val thm2 = thm1 |> Thm.instantiate (instT, []) |> Raw_Simplifier.rewrite_rule ctxt pre_simps val ctxt' = Variable.declare_names (Thm.full_prop_of thm2) ctxt val rule = transfer_rule_of_lhs ctxt' (HOLogic.dest_Trueprop (Thm.concl_of thm2)) in Goal.prove_internal ctxt' [] (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (Var (("P", 0), \<^typ>\bool\)))) (fn _ => resolve_tac ctxt' [thm2 RS @{thm transfer_start'}, thm2 RS @{thm transfer_start}] 1 THEN (resolve_tac ctxt' [rule] THEN_ALL_NEW (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt' eq_rules)))) 1 handle TERM (_, ts) => raise TERM ("Transfer failed to convert goal to an object-logic formula", ts)) |> Raw_Simplifier.rewrite_rule ctxt' post_simps |> Simplifier.norm_hhf ctxt' |> Drule.generalize (Symtab.make_set (map (fst o dest_TFree o Thm.typ_of o snd) instT), Symtab.empty) |> Drule.zero_var_indexes end (* handle THM _ => thm *) fun untransferred ctxt extra_rules thm = let val rules = extra_rules @ get_transfer_raw ctxt val eq_rules = get_relator_eq_raw ctxt val pre_simps = @{thms transfer_forall_eq transfer_implies_eq} - val thm1 = Drule.forall_intr_vars thm + val thm1 = Thm.forall_intr_vars thm val instT = rev (Term.add_tvars (Thm.full_prop_of thm1) []) |> map (fn v as ((a, _), S) => (v, Thm.ctyp_of ctxt (TFree (a, S)))) val thm2 = thm1 |> Thm.instantiate (instT, []) |> Raw_Simplifier.rewrite_rule ctxt pre_simps val ctxt' = Variable.declare_names (Thm.full_prop_of thm2) ctxt val t = HOLogic.dest_Trueprop (Thm.concl_of thm2) val rule = transfer_rule_of_term ctxt' true t in Goal.prove_internal ctxt' [] (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (Var (("P", 0), \<^typ>\bool\)))) (fn _ => resolve_tac ctxt' [thm2 RS @{thm untransfer_start}] 1 THEN (resolve_tac ctxt' [rule] THEN_ALL_NEW (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt' eq_rules)))) 1 handle TERM (_, ts) => raise TERM ("Transfer failed to convert goal to an object-logic formula", ts)) |> Raw_Simplifier.rewrite_rule ctxt' post_simps |> Simplifier.norm_hhf ctxt' |> Drule.generalize (Symtab.make_set (map (fst o dest_TFree o Thm.typ_of o snd) instT), Symtab.empty) |> Drule.zero_var_indexes end (** Methods and attributes **) val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) => error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)) val fixing = Scan.optional (Scan.lift (Args.$$$ "fixing" -- Args.colon) |-- Scan.repeat free) [] val reverse_prems = fn _ => PRIMITIVE (fn thm => fold_rev (fn i => Thm.permute_prems i 1) (0 upto Thm.nprems_of thm - 1) thm); fun transfer_start_method equiv : (Proof.context -> Proof.method) context_parser = fixing >> (fn vs => fn ctxt => SIMPLE_METHOD' (gen_frees_tac vs ctxt THEN' transfer_start_tac equiv ctxt THEN' reverse_prems)); fun transfer_method equiv : (Proof.context -> Proof.method) context_parser = fixing >> (fn vs => fn ctxt => SIMPLE_METHOD' (gen_frees_tac vs ctxt THEN' transfer_tac equiv ctxt)) val transfer_prover_start_method : (Proof.context -> Proof.method) context_parser = Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_start_tac ctxt THEN' reverse_prems)) val transfer_prover_method : (Proof.context -> Proof.method) context_parser = Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_tac ctxt)) (* Attribute for transfer rules *) fun prep_rule ctxt = abstract_domains_transfer ctxt o abstract_equalities_transfer ctxt o Conv.fconv_rule prep_conv val transfer_add = Thm.declaration_attribute (fn thm => fn ctxt => (add_transfer_thm o prep_rule (Context.proof_of ctxt)) thm ctxt) val transfer_del = Thm.declaration_attribute (fn thm => fn ctxt => (del_transfer_thm o prep_rule (Context.proof_of ctxt)) thm ctxt) val transfer_attribute = Attrib.add_del transfer_add transfer_del (* Attributes for transfer domain rules *) val transfer_domain_add = Thm.declaration_attribute add_transfer_domain_thm val transfer_domain_del = Thm.declaration_attribute del_transfer_domain_thm val transfer_domain_attribute = Attrib.add_del transfer_domain_add transfer_domain_del (* Attributes for transferred rules *) fun transferred_attribute thms = Thm.rule_attribute thms (fn context => transferred (Context.proof_of context) thms) fun untransferred_attribute thms = Thm.rule_attribute thms (fn context => untransferred (Context.proof_of context) thms) val transferred_attribute_parser = Attrib.thms >> transferred_attribute val untransferred_attribute_parser = Attrib.thms >> untransferred_attribute fun morph_pred_data phi = let val morph_thm = Morphism.thm phi in map_pred_data' morph_thm morph_thm (map morph_thm) end fun lookup_pred_data ctxt type_name = Symtab.lookup (get_pred_data ctxt) type_name |> Option.map (morph_pred_data (Morphism.transfer_morphism' ctxt)) fun update_pred_data type_name qinfo ctxt = Data.map (map_pred_data (Symtab.update (type_name, qinfo))) ctxt (* Theory setup *) val _ = Theory.setup let val name = \<^binding>\relator_eq\ fun add_thm thm context = context |> Data.map (map_relator_eq (Item_Net.update (Thm.trim_context thm))) |> Data.map (map_relator_eq_raw (Item_Net.update (Thm.trim_context (abstract_equalities_relator_eq (Context.proof_of context) thm)))) fun del_thm thm context = context |> Data.map (map_relator_eq (Item_Net.remove thm)) |> Data.map (map_relator_eq_raw (Item_Net.remove (abstract_equalities_relator_eq (Context.proof_of context) thm))) val add = Thm.declaration_attribute add_thm val del = Thm.declaration_attribute del_thm val text = "declaration of relator equality rule (used by transfer method)" val content = Item_Net.content o #relator_eq o Data.get in Attrib.setup name (Attrib.add_del add del) text #> Global_Theory.add_thms_dynamic (name, content) end val _ = Theory.setup let val name = \<^binding>\relator_domain\ fun add_thm thm context = let val thm' = thm |> abstract_domains_relator_domain (Context.proof_of context) |> Thm.trim_context in context |> Data.map (map_relator_domain (Item_Net.update thm')) |> add_transfer_domain_thm thm' end fun del_thm thm context = let val thm' = abstract_domains_relator_domain (Context.proof_of context) thm in context |> Data.map (map_relator_domain (Item_Net.remove thm')) |> del_transfer_domain_thm thm' end val add = Thm.declaration_attribute add_thm val del = Thm.declaration_attribute del_thm val text = "declaration of relator domain rule (used by transfer method)" val content = Item_Net.content o #relator_domain o Data.get in Attrib.setup name (Attrib.add_del add del) text #> Global_Theory.add_thms_dynamic (name, content) end val _ = Theory.setup (Attrib.setup \<^binding>\transfer_rule\ transfer_attribute "transfer rule for transfer method" #> Global_Theory.add_thms_dynamic (\<^binding>\transfer_raw\, Item_Net.content o #transfer_raw o Data.get) #> Attrib.setup \<^binding>\transfer_domain_rule\ transfer_domain_attribute "transfer domain rule for transfer method" #> Attrib.setup \<^binding>\transferred\ transferred_attribute_parser "raw theorem transferred to abstract theorem using transfer rules" #> Attrib.setup \<^binding>\untransferred\ untransferred_attribute_parser "abstract theorem transferred to raw theorem using transfer rules" #> Global_Theory.add_thms_dynamic (\<^binding>\relator_eq_raw\, Item_Net.content o #relator_eq_raw o Data.get) #> Method.setup \<^binding>\transfer_start\ (transfer_start_method true) "firtst step in the transfer algorithm (for debugging transfer)" #> Method.setup \<^binding>\transfer_start'\ (transfer_start_method false) "firtst step in the transfer algorithm (for debugging transfer)" #> Method.setup \<^binding>\transfer_prover_start\ transfer_prover_start_method "firtst step in the transfer_prover algorithm (for debugging transfer_prover)" #> Method.setup \<^binding>\transfer_step\ (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_step_tac ctxt))) "step in the search for transfer rules (for debugging transfer and transfer_prover)" #> Method.setup \<^binding>\transfer_end\ (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_end_tac ctxt))) "last step in the transfer algorithm (for debugging transfer)" #> Method.setup \<^binding>\transfer_prover_end\ (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_end_tac ctxt))) "last step in the transfer_prover algorithm (for debugging transfer_prover)" #> Method.setup \<^binding>\transfer\ (transfer_method true) "generic theorem transfer method" #> Method.setup \<^binding>\transfer'\ (transfer_method false) "generic theorem transfer method" #> Method.setup \<^binding>\transfer_prover\ transfer_prover_method "for proving transfer rules") end diff --git a/src/HOL/Tools/choice_specification.ML b/src/HOL/Tools/choice_specification.ML --- a/src/HOL/Tools/choice_specification.ML +++ b/src/HOL/Tools/choice_specification.ML @@ -1,206 +1,206 @@ (* Title: HOL/Tools/choice_specification.ML Author: Sebastian Skalberg, TU Muenchen Package for defining constants by specification. *) signature CHOICE_SPECIFICATION = sig val close_form : term -> term val add_specification: (bstring * xstring * bool) list -> theory * thm -> theory * thm end structure Choice_Specification: CHOICE_SPECIFICATION = struct local fun mk_definitional [] arg = arg | mk_definitional ((thname, cname, covld) :: cos) (thy, thm) = (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Ex\, _) $ P => let val ctype = domain_type (type_of P) val cname_full = Sign.intern_const thy cname val cdefname = if thname = "" then Thm.def_name (Long_Name.base_name cname) else thname val def_eq = Logic.mk_equals (Const (cname_full, ctype), HOLogic.choice_const ctype $ P) val (thms, thy') = Global_Theory.add_defs covld [((Binding.name cdefname, def_eq), [])] thy val thm' = [thm,hd thms] MRS @{thm exE_some} in mk_definitional cos (thy',thm') end | _ => raise THM ("Bad specification theorem", 0, [thm])) in fun add_specification cos = mk_definitional cos #> apsnd Drule.export_without_context end (* Collect all intances of constants in term *) fun collect_consts (t $ u, tms) = collect_consts (u, collect_consts (t, tms)) | collect_consts (Abs (_, _, t), tms) = collect_consts (t, tms) | collect_consts (tm as Const _, tms) = insert (op aconv) tm tms | collect_consts (_, tms) = tms (* Complementing Type.varify_global... *) fun unvarify_global t fmap = let val fmap' = map Library.swap fmap fun unthaw (f as (a, S)) = (case AList.lookup (op =) fmap' a of NONE => TVar f | SOME (b, _) => TFree (b, S)) in map_types (map_type_tvar unthaw) t end (* The syntactic meddling needed to setup add_specification for work *) fun close_form t = fold_rev (fn (s, T) => fn t => HOLogic.mk_all (s, T, t)) (map dest_Free (Misc_Legacy.term_frees t)) t fun zip3 [] [] [] = [] | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs | zip3 _ _ _ = raise Fail "Choice_Specification.process_spec internal error" fun myfoldr f [x] = x | myfoldr f (x::xs) = f (x,myfoldr f xs) | myfoldr f [] = raise Fail "Choice_Specification.process_spec internal error" fun process_spec cos alt_props thy = let val ctxt = Proof_Context.init_global thy val rew_imps = alt_props |> map (Object_Logic.atomize ctxt o Thm.cterm_of ctxt o Syntax.read_prop_global thy o snd) val props' = rew_imps |> map (HOLogic.dest_Trueprop o Thm.term_of o snd o Thm.dest_equals o Thm.cprop_of) fun proc_single prop = let val frees = Misc_Legacy.term_frees prop val _ = forall (fn v => Sign.of_sort thy (type_of v, \<^sort>\type\)) frees orelse error "Specificaton: Only free variables of sort 'type' allowed" val prop_closed = close_form prop in (prop_closed, frees) end val props'' = map proc_single props' val frees = map snd props'' val prop = myfoldr HOLogic.mk_conj (map fst props'') - val (vmap, prop_thawed) = Type.varify_global [] prop + val (vmap, prop_thawed) = Type.varify_global Term_Subst.TFrees.empty prop val thawed_prop_consts = collect_consts (prop_thawed, []) val (altcos, overloaded) = split_list cos val (names, sconsts) = split_list altcos val consts = map (Syntax.read_term_global thy) sconsts val _ = not (Library.exists (not o Term.is_Const) consts) orelse error "Specification: Non-constant found as parameter" fun proc_const c = let - val (_, c') = Type.varify_global [] c + val (_, c') = Type.varify_global Term_Subst.TFrees.empty c val (cname,ctyp) = dest_Const c' in (case filter (fn t => let val (name, typ) = dest_Const t in name = cname andalso Sign.typ_equiv thy (typ, ctyp) end) thawed_prop_consts of [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found") | [cf] => unvarify_global cf vmap | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")) end val proc_consts = map proc_const consts fun mk_exist c prop = let val T = type_of c val cname = Long_Name.base_name (fst (dest_Const c)) val vname = if Symbol_Pos.is_identifier cname then cname else "x" in HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c, prop)) end val ex_prop = fold_rev mk_exist proc_consts prop val cnames = map (fst o dest_Const) proc_consts fun post_process (arg as (thy,thm)) = let fun inst_all thy v thm = let val cv = Thm.global_cterm_of thy v val cT = Thm.ctyp_of_cterm cv val spec' = Thm.instantiate' [SOME cT] [NONE, SOME cv] spec in thm RS spec' end fun remove_alls frees (thm, thy) = (fold (inst_all thy) frees thm, thy) fun process_single ((name, atts), rew_imp, frees) args = let fun undo_imps thm = Thm.equal_elim (Thm.symmetric rew_imp) thm fun add_final (thm, thy) = if name = "" then (thm, thy) else (writeln (" " ^ name ^ ": " ^ Thm.string_of_thm_global thy thm); Global_Theory.store_thm (Binding.name name, thm) thy) in swap args |> remove_alls frees |> apfst undo_imps |> apfst Drule.export_without_context |-> Thm.theory_attributes (map (Attrib.attribute_cmd_global thy) (@{attributes [nitpick_choice_spec]} @ atts)) |> add_final |> swap end fun process_all [proc_arg] args = process_single proc_arg args | process_all (proc_arg::rest) (thy,thm) = let val single_th = thm RS conjunct1 val rest_th = thm RS conjunct2 val (thy', _) = process_single proc_arg (thy, single_th) in process_all rest (thy', rest_th) end | process_all [] _ = raise Fail "Choice_Specification.process_spec internal error" val alt_names = map fst alt_props val _ = if exists (fn (name, _) => name <> "") alt_names then writeln "specification" else () in arg |> apsnd (Thm.unvarify_global thy) |> process_all (zip3 alt_names rew_imps frees) end fun after_qed [[thm]] = Proof_Context.background_theory (fn thy => #1 (post_process (add_specification (zip3 names cnames overloaded) (thy, thm)))); in thy |> Proof_Context.init_global |> Variable.declare_term ex_prop |> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]] end (* outer syntax *) val opt_name = Scan.optional (Parse.name --| \<^keyword>\:\) "" val opt_overloaded = Parse.opt_keyword "overloaded" val _ = Outer_Syntax.command \<^command_keyword>\specification\ "define constants by specification" (\<^keyword>\(\ |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| \<^keyword>\)\ -- Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop) >> (fn (cos, alt_props) => Toplevel.theory_to_proof (process_spec cos alt_props))) end diff --git a/src/HOL/Tools/datatype_realizer.ML b/src/HOL/Tools/datatype_realizer.ML --- a/src/HOL/Tools/datatype_realizer.ML +++ b/src/HOL/Tools/datatype_realizer.ML @@ -1,244 +1,244 @@ (* Title: HOL/Tools/datatype_realizer.ML Author: Stefan Berghofer, TU Muenchen Program extraction from proofs involving datatypes: realizers for induction and case analysis. *) signature DATATYPE_REALIZER = sig val add_dt_realizers: Old_Datatype_Aux.config -> string list -> theory -> theory end; structure Datatype_Realizer : DATATYPE_REALIZER = struct fun subsets i j = if i <= j then let val is = subsets (i+1) j in map (fn ks => i::ks) is @ is end else [[]]; fun is_unit t = body_type (fastype_of t) = HOLogic.unitT; fun tname_of (Type (s, _)) = s | tname_of _ = ""; fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Old_Datatype_Aux.info) is thy = let val ctxt = Proof_Context.init_global thy; val recTs = Old_Datatype_Aux.get_rec_types descr; val pnames = if length descr = 1 then ["P"] else map (fn i => "P" ^ string_of_int i) (1 upto length descr); val rec_result_Ts = map (fn ((i, _), P) => if member (op =) is i then TFree ("'" ^ P, \<^sort>\type\) else HOLogic.unitT) (descr ~~ pnames); fun make_pred i T U r x = if member (op =) is i then Free (nth pnames i, T --> U --> HOLogic.boolT) $ r $ x else Free (nth pnames i, U --> HOLogic.boolT) $ x; fun mk_all i s T t = if member (op =) is i then Logic.all (Free (s, T)) t else t; val (prems, rec_fns) = split_list (flat (fst (fold_map (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j => let val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) cargs; val tnames = Name.variant_list pnames (Old_Datatype_Prop.make_tnames Ts); val recs = filter (Old_Datatype_Aux.is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts); val frees = tnames ~~ Ts; fun mk_prems vs [] = let val rT = nth (rec_result_Ts) i; val vs' = filter_out is_unit vs; val f = Old_Datatype_Aux.mk_Free "f" (map fastype_of vs' ---> rT) j; val f' = Envir.eta_contract (fold_rev (absfree o dest_Free) vs (if member (op =) is i then list_comb (f, vs') else HOLogic.unit)); in (HOLogic.mk_Trueprop (make_pred i rT T (list_comb (f, vs')) (list_comb (Const (cname, Ts ---> T), map Free frees))), f') end | mk_prems vs (((dt, s), T) :: ds) = let val k = Old_Datatype_Aux.body_index dt; val (Us, U) = strip_type T; val i = length Us; val rT = nth (rec_result_Ts) k; val r = Free ("r" ^ s, Us ---> rT); val (p, f) = mk_prems (vs @ [r]) ds; in (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies (Logic.list_all (map (pair "x") Us, HOLogic.mk_Trueprop (make_pred k rT U (Old_Datatype_Aux.app_bnds r i) (Old_Datatype_Aux.app_bnds (Free (s, T)) i))), p)), f) end; in (apfst (fold_rev (Logic.all o Free) frees) (mk_prems (map Free frees) recs), j + 1) end) constrs) (descr ~~ recTs) 1))); fun mk_proj _ [] t = t | mk_proj j (i :: is) t = if null is then t else if (j: int) = i then HOLogic.mk_fst t else mk_proj j is (HOLogic.mk_snd t); val tnames = Old_Datatype_Prop.make_tnames recTs; val fTs = map fastype_of rec_fns; val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0))) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names); val r = if null is then Extraction.nullt else foldr1 HOLogic.mk_prod (map_filter (fn (((((i, _), T), U), s), tname) => if member (op =) is i then SOME (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T)) else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames)); val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop \<^const_name>\HOL.conj\) (map (fn ((((i, _), T), U), tname) => make_pred i U T (mk_proj i is r) (Free (tname, T))) (descr ~~ recTs ~~ rec_result_Ts ~~ tnames))); val inst = map (#1 o dest_Var o head_of) (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct))) ~~ map (Thm.cterm_of ctxt) ps; val thm = Goal.prove_internal ctxt (map (Thm.cterm_of ctxt) prems) (Thm.cterm_of ctxt concl) (fn prems => EVERY [ rewrite_goals_tac ctxt (map mk_meta_eq [@{thm fst_conv}, @{thm snd_conv}]), resolve_tac ctxt [infer_instantiate ctxt inst induct] 1, ALLGOALS (Object_Logic.atomize_prems_tac ctxt), rewrite_goals_tac ctxt (@{thm o_def} :: map mk_meta_eq rec_rewrites), REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW (fn i => REPEAT (eresolve_tac ctxt [allE] i) THEN assume_tac ctxt i)) 1)]) |> Drule.export_without_context; val ind_name = Thm.derivation_name induct; val vs = map (nth pnames) is; val (thm', thy') = thy |> Sign.root_path |> Global_Theory.store_thm (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm) ||> Sign.restore_naming thy; val ivs = rev (Term.add_vars (Logic.varify_global (Old_Datatype_Prop.make_ind [descr])) []); - val rvs = rev (Thm.fold_terms Term.add_vars thm' []); + val rvs = rev (Thm.fold_terms {hyps = false} Term.add_vars thm' []); val ivs1 = map Var (filter_out (fn (_, T) => \<^type_name>\bool\ = tname_of (body_type T)) ivs); val ivs2 = map (fn (ixn, _) => Var (ixn, the (AList.lookup (op =) rvs ixn))) ivs; val prf = Extraction.abs_corr_shyps thy' induct vs ivs2 (fold_rev (fn (f, p) => fn prf => (case head_of (strip_abs_body f) of Free (s, T) => let val T' = Logic.varifyT_global T in Abst (s, SOME T', Proofterm.prf_abstract_over (Var ((s, 0), T')) (AbsP ("H", SOME p, prf))) end | _ => AbsP ("H", SOME p, prf))) (rec_fns ~~ Thm.prems_of thm) (Proofterm.proof_combP (Thm.reconstruct_proof_of thm', map PBound (length prems - 1 downto 0)))); val r' = if null is then r else Logic.varify_global (fold_rev lambda (map Logic.unvarify_global ivs1 @ filter_out is_unit (map (head_of o strip_abs_body) rec_fns)) r); in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end; fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Old_Datatype_Aux.info) thy = let val ctxt = Proof_Context.init_global thy; val rT = TFree ("'P", \<^sort>\type\); val rT' = TVar (("'P", 0), \<^sort>\type\); fun make_casedist_prem T (cname, cargs) = let val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) cargs; val frees = Name.variant_list ["P", "y"] (Old_Datatype_Prop.make_tnames Ts) ~~ Ts; val free_ts = map Free frees; val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT) in (r, fold_rev Logic.all free_ts (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))), HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $ list_comb (r, free_ts))))) end; val SOME (_, _, constrs) = AList.lookup (op =) descr index; val T = nth (Old_Datatype_Aux.get_rec_types descr) index; val (rs, prems) = split_list (map (make_casedist_prem T) constrs); val r = Const (case_name, map fastype_of rs ---> T --> rT); val y = Var (("y", 0), Logic.varifyT_global T); val y' = Free ("y", T); val thm = Goal.prove_internal ctxt (map (Thm.cterm_of ctxt) prems) (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $ list_comb (r, rs @ [y'])))) (fn prems => EVERY [ resolve_tac ctxt [infer_instantiate ctxt [(#1 (dest_Var y), Thm.cterm_of ctxt y')] exhaust] 1, ALLGOALS (EVERY' [asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps case_rewrites), resolve_tac ctxt prems, asm_simp_tac (put_simpset HOL_basic_ss ctxt)])]) |> Drule.export_without_context; val exh_name = Thm.derivation_name exhaust; val (thm', thy') = thy |> Sign.root_path |> Global_Theory.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm) ||> Sign.restore_naming thy; val P = Var (("P", 0), rT' --> HOLogic.boolT); val prf = Extraction.abs_corr_shyps thy' exhaust ["P"] [y, P] (fold_rev (fn (p, r) => fn prf => Proofterm.forall_intr_proof' (Logic.varify_global r) (AbsP ("H", SOME (Logic.varify_global p), prf))) (prems ~~ rs) (Proofterm.proof_combP (Thm.reconstruct_proof_of thm', map PBound (length prems - 1 downto 0)))); val prf' = Extraction.abs_corr_shyps thy' exhaust [] (map Var (Term.add_vars (Thm.prop_of exhaust) [])) (Thm.reconstruct_proof_of exhaust); val r' = Logic.varify_global (Abs ("y", T, (fold_rev (Term.abs o dest_Free) rs (list_comb (r, map Bound ((length rs - 1 downto 0) @ [length rs])))))); in Extraction.add_realizers_i [(exh_name, (["P"], r', prf)), (exh_name, ([], Extraction.nullt, prf'))] thy' end; fun add_dt_realizers config names thy = if not (Proofterm.proofs_enabled ()) then thy else let val _ = Old_Datatype_Aux.message config "Adding realizers for induction and case analysis ..."; val infos = map (BNF_LFP_Compat.the_info thy []) names; val info :: _ = infos; in thy |> fold_rev (perhaps o try o make_ind info) (subsets 0 (length (#descr info) - 1)) |> fold_rev (perhaps o try o make_casedists) infos end; val _ = Theory.setup (BNF_LFP_Compat.interpretation \<^plugin>\extraction\ [] add_dt_realizers); end; diff --git a/src/HOL/Tools/groebner.ML b/src/HOL/Tools/groebner.ML --- a/src/HOL/Tools/groebner.ML +++ b/src/HOL/Tools/groebner.ML @@ -1,986 +1,985 @@ (* Title: HOL/Tools/groebner.ML Author: Amine Chaieb, TU Muenchen *) signature GROEBNER = sig val ring_and_ideal_conv: {idom: thm list, ring: cterm list * thm list, field: cterm list * thm list, vars: cterm list, semiring: cterm list * thm list, ideal : thm list} -> (cterm -> Rat.rat) -> (Rat.rat -> cterm) -> conv -> conv -> {ring_conv: Proof.context -> conv, simple_ideal: cterm list -> cterm -> cterm ord -> cterm list, multi_ideal: cterm list -> cterm list -> cterm list -> (cterm * cterm) list, poly_eq_ss: simpset, unwind_conv: Proof.context -> conv} val ring_tac: thm list -> thm list -> Proof.context -> int -> tactic val ideal_tac: thm list -> thm list -> Proof.context -> int -> tactic val algebra_tac: thm list -> thm list -> Proof.context -> int -> tactic end structure Groebner : GROEBNER = struct val concl = Thm.cprop_of #> Thm.dest_arg; fun is_binop ct ct' = (case Thm.term_of ct' of c $ _ $ _ => Thm.term_of ct aconv c | _ => false); fun dest_binary ct ct' = if is_binop ct ct' then Thm.dest_binop ct' else raise CTERM ("dest_binary: bad binop", [ct, ct']) val denominator_rat = Rat.dest #> snd #> Rat.of_int; fun int_of_rat a = case Rat.dest a of (i,1) => i | _ => error "int_of_rat: not an int"; val lcm_rat = fn x => fn y => Rat.of_int (Integer.lcm (int_of_rat x) (int_of_rat y)); val (eqF_intr, eqF_elim) = let val [th1,th2] = @{thms PFalse} in (fn th => th COMP th2, fn th => th COMP th1) end; val (PFalse, PFalse') = let val PFalse_eq = nth @{thms simp_thms} 13 in (PFalse_eq RS iffD1, PFalse_eq RS iffD2) end; (* Type for recording history, i.e. how a polynomial was obtained. *) datatype history = Start of int | Mmul of (Rat.rat * int list) * history | Add of history * history; (* Monomial ordering. *) fun morder_lt m1 m2= let fun lexorder l1 l2 = case (l1,l2) of ([],[]) => false | (x1::o1,x2::o2) => x1 > x2 orelse x1 = x2 andalso lexorder o1 o2 | _ => error "morder: inconsistent monomial lengths" val n1 = Integer.sum m1 val n2 = Integer.sum m2 in n1 < n2 orelse n1 = n2 andalso lexorder m1 m2 end; (* Arithmetic on canonical polynomials. *) fun grob_neg l = map (fn (c,m) => (Rat.neg c,m)) l; fun grob_add l1 l2 = case (l1,l2) of ([],l2) => l2 | (l1,[]) => l1 | ((c1,m1)::o1,(c2,m2)::o2) => if m1 = m2 then let val c = c1 + c2 val rest = grob_add o1 o2 in if c = @0 then rest else (c,m1)::rest end else if morder_lt m2 m1 then (c1,m1)::(grob_add o1 l2) else (c2,m2)::(grob_add l1 o2); fun grob_sub l1 l2 = grob_add l1 (grob_neg l2); fun grob_mmul (c1,m1) (c2,m2) = (c1 * c2, ListPair.map (op +) (m1, m2)); fun grob_cmul cm pol = map (grob_mmul cm) pol; fun grob_mul l1 l2 = case l1 of [] => [] | (h1::t1) => grob_add (grob_cmul h1 l2) (grob_mul t1 l2); fun grob_inv l = case l of [(c,vs)] => if (forall (fn x => x = 0) vs) then if c = @0 then error "grob_inv: division by zero" else [(@1 / c,vs)] else error "grob_inv: non-constant divisor polynomial" | _ => error "grob_inv: non-constant divisor polynomial"; fun grob_div l1 l2 = case l2 of [(c,l)] => if (forall (fn x => x = 0) l) then if c = @0 then error "grob_div: division by zero" else grob_cmul (@1 / c,l) l1 else error "grob_div: non-constant divisor polynomial" | _ => error "grob_div: non-constant divisor polynomial"; fun grob_pow vars l n = if n < 0 then error "grob_pow: negative power" else if n = 0 then [(@1,map (K 0) vars)] else grob_mul l (grob_pow vars l (n - 1)); (* Monomial division operation. *) fun mdiv (c1,m1) (c2,m2) = (c1 / c2, map2 (fn n1 => fn n2 => if n1 < n2 then error "mdiv" else n1 - n2) m1 m2); (* Lowest common multiple of two monomials. *) fun mlcm (_,m1) (_,m2) = (@1, ListPair.map Int.max (m1, m2)); (* Reduce monomial cm by polynomial pol, returning replacement for cm. *) fun reduce1 cm (pol,hpol) = case pol of [] => error "reduce1" | cm1::cms => ((let val (c,m) = mdiv cm cm1 in (grob_cmul (~ c, m) cms, Mmul ((~ c,m),hpol)) end) handle ERROR _ => error "reduce1"); (* Try this for all polynomials in a basis. *) fun tryfind f l = case l of [] => error "tryfind" | (h::t) => ((f h) handle ERROR _ => tryfind f t); fun reduceb cm basis = tryfind (fn p => reduce1 cm p) basis; (* Reduction of a polynomial (always picking largest monomial possible). *) fun reduce basis (pol,hist) = case pol of [] => (pol,hist) | cm::ptl => ((let val (q,hnew) = reduceb cm basis in reduce basis (grob_add q ptl,Add(hnew,hist)) end) handle (ERROR _) => (let val (q,hist') = reduce basis (ptl,hist) in (cm::q,hist') end)); (* Check for orthogonality w.r.t. LCM. *) fun orthogonal l p1 p2 = snd l = snd(grob_mmul (hd p1) (hd p2)); (* Compute S-polynomial of two polynomials. *) fun spoly cm ph1 ph2 = case (ph1,ph2) of (([],h),_) => ([],h) | (_,([],h)) => ([],h) | ((cm1::ptl1,his1),(cm2::ptl2,his2)) => (grob_sub (grob_cmul (mdiv cm cm1) ptl1) (grob_cmul (mdiv cm cm2) ptl2), Add(Mmul(mdiv cm cm1,his1), Mmul(mdiv (~ (fst cm),snd cm) cm2,his2))); (* Make a polynomial monic. *) fun monic (pol,hist) = if null pol then (pol,hist) else let val (c',m') = hd pol in (map (fn (c,m) => (c / c',m)) pol, Mmul((@1 / c',map (K 0) m'),hist)) end; (* The most popular heuristic is to order critical pairs by LCM monomial. *) fun forder ((_,m1),_) ((_,m2),_) = morder_lt m1 m2; fun poly_lt p q = case (p,q) of (_,[]) => false | ([],_) => true | ((c1,m1)::o1,(c2,m2)::o2) => c1 < c2 orelse c1 = c2 andalso ((morder_lt m1 m2) orelse m1 = m2 andalso poly_lt o1 o2); fun align ((p,hp),(q,hq)) = if poly_lt p q then ((p,hp),(q,hq)) else ((q,hq),(p,hp)); fun poly_eq p1 p2 = eq_list (fn ((c1, m1), (c2, m2)) => c1 = c2 andalso (m1: int list) = m2) (p1, p2); fun memx ((p1,_),(p2,_)) ppairs = not (exists (fn ((q1,_),(q2,_)) => poly_eq p1 q1 andalso poly_eq p2 q2) ppairs); (* Buchberger's second criterion. *) fun criterion2 basis (lcm,((p1,h1),(p2,h2))) opairs = exists (fn g => not(poly_eq (fst g) p1) andalso not(poly_eq (fst g) p2) andalso can (mdiv lcm) (hd(fst g)) andalso not(memx (align (g,(p1,h1))) (map snd opairs)) andalso not(memx (align (g,(p2,h2))) (map snd opairs))) basis; (* Test for hitting constant polynomial. *) fun constant_poly p = length p = 1 andalso forall (fn x => x = 0) (snd(hd p)); (* Grobner basis algorithm. *) (* FIXME: try to get rid of mergesort? *) fun merge ord l1 l2 = case l1 of [] => l2 | h1::t1 => case l2 of [] => l1 | h2::t2 => if ord h1 h2 then h1::(merge ord t1 l2) else h2::(merge ord l1 t2); fun mergesort ord l = let fun mergepairs l1 l2 = case (l1,l2) of ([s],[]) => s | (l,[]) => mergepairs [] l | (l,[s1]) => mergepairs (s1::l) [] | (l,(s1::s2::ss)) => mergepairs ((merge ord s1 s2)::l) ss in if null l then [] else mergepairs [] (map (fn x => [x]) l) end; fun grobner_basis basis pairs = case pairs of [] => basis | (l,(p1,p2))::opairs => let val (sph as (sp,_)) = monic (reduce basis (spoly l p1 p2)) in if null sp orelse criterion2 basis (l,(p1,p2)) opairs then grobner_basis basis opairs else if constant_poly sp then grobner_basis (sph::basis) [] else let val rawcps = map (fn p => (mlcm (hd(fst p)) (hd sp),align(p,sph))) basis val newcps = filter (fn (l,(p,q)) => not(orthogonal l (fst p) (fst q))) rawcps in grobner_basis (sph::basis) (merge forder opairs (mergesort forder newcps)) end end; (* Interreduce initial polynomials. *) fun grobner_interreduce rpols ipols = case ipols of [] => map monic (rev rpols) | p::ps => let val p' = reduce (rpols @ ps) p in if null (fst p') then grobner_interreduce rpols ps else grobner_interreduce (p'::rpols) ps end; (* Overall function. *) fun grobner pols = let val npols = map_index (fn (n, p) => (p, Start n)) pols val phists = filter (fn (p,_) => not (null p)) npols val bas = grobner_interreduce [] (map monic phists) val prs0 = map_product pair bas bas val prs1 = filter (fn ((x,_),(y,_)) => poly_lt x y) prs0 val prs2 = map (fn (p,q) => (mlcm (hd(fst p)) (hd(fst q)),(p,q))) prs1 val prs3 = filter (fn (l,(p,q)) => not(orthogonal l (fst p) (fst q))) prs2 in grobner_basis bas (mergesort forder prs3) end; (* Get proof of contradiction from Grobner basis. *) fun find p l = case l of [] => error "find" | (h::t) => if p(h) then h else find p t; fun grobner_refute pols = let val gb = grobner pols in snd(find (fn (p,_) => length p = 1 andalso forall (fn x=> x=0) (snd(hd p))) gb) end; (* Turn proof into a certificate as sum of multipliers. *) (* In principle this is very inefficient: in a heavily shared proof it may *) (* make the same calculation many times. Could put in a cache or something. *) fun resolve_proof vars prf = case prf of Start(~1) => [] | Start m => [(m,[(@1,map (K 0) vars)])] | Mmul(pol,lin) => let val lis = resolve_proof vars lin in map (fn (n,p) => (n,grob_cmul pol p)) lis end | Add(lin1,lin2) => let val lis1 = resolve_proof vars lin1 val lis2 = resolve_proof vars lin2 val dom = distinct (op =) (union (op =) (map fst lis1) (map fst lis2)) in map (fn n => let val a = these (AList.lookup (op =) lis1 n) val b = these (AList.lookup (op =) lis2 n) in (n,grob_add a b) end) dom end; (* Run the procedure and produce Weak Nullstellensatz certificate. *) fun grobner_weak vars pols = let val cert = resolve_proof vars (grobner_refute pols) val l = fold_rev (fold_rev (lcm_rat o denominator_rat o fst) o snd) cert @1 in (l,map (fn (i,p) => (i,map (fn (d,m) => (l * d,m)) p)) cert) end; (* Prove a polynomial is in ideal generated by others, using Grobner basis. *) fun grobner_ideal vars pols pol = let val (pol',h) = reduce (grobner pols) (grob_neg pol,Start(~1)) in if not (null pol') then error "grobner_ideal: not in the ideal" else resolve_proof vars h end; (* Produce Strong Nullstellensatz certificate for a power of pol. *) fun grobner_strong vars pols pol = let val vars' = \<^cterm>\True\::vars val grob_z = [(@1, 1::(map (K 0) vars))] val grob_1 = [(@1, (map (K 0) vars'))] fun augment p= map (fn (c,m) => (c,0::m)) p val pols' = map augment pols val pol' = augment pol val allpols = (grob_sub (grob_mul grob_z pol') grob_1)::pols' val (l,cert) = grobner_weak vars' allpols val d = fold (fold (Integer.max o hd o snd) o snd) cert 0 fun transform_monomial (c,m) = grob_cmul (c,tl m) (grob_pow vars pol (d - hd m)) fun transform_polynomial q = fold_rev (grob_add o transform_monomial) q [] val cert' = map (fn (c,q) => (c-1,transform_polynomial q)) (filter (fn (k,_) => k <> 0) cert) in (d,l,cert') end; (* Overall parametrized universal procedure for (semi)rings. *) (* We return an ideal_conv and the actual ring prover. *) fun refute_disj rfn tm = case Thm.term_of tm of Const(\<^const_name>\HOL.disj\,_)$_$_ => Drule.compose (refute_disj rfn (Thm.dest_arg tm), 2, Drule.compose (refute_disj rfn (Thm.dest_arg1 tm), 2, disjE)) | _ => rfn tm ; val notnotD = @{thm notnotD}; fun mk_binop ct x y = Thm.apply (Thm.apply ct x) y fun is_neg t = case Thm.term_of t of (Const(\<^const_name>\Not\,_)$_) => true | _ => false; fun is_eq t = case Thm.term_of t of (Const(\<^const_name>\HOL.eq\,_)$_$_) => true | _ => false; fun end_itlist f l = case l of [] => error "end_itlist" | [x] => x | (h::t) => f h (end_itlist f t); val list_mk_binop = fn b => end_itlist (mk_binop b); val list_dest_binop = fn b => let fun h acc t = ((let val (l,r) = dest_binary b t in h (h acc r) l end) handle CTERM _ => (t::acc)) (* Why had I handle _ => ? *) in h [] end; val strip_exists = let fun h (acc, t) = case Thm.term_of t of Const (\<^const_name>\Ex\, _) $ Abs _ => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc)) | _ => (acc,t) in fn t => h ([],t) end; fun is_forall t = case Thm.term_of t of (Const(\<^const_name>\All\,_)$Abs(_,_,_)) => true | _ => false; val nnf_simps = @{thms nnf_simps}; fun weak_dnf_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms weak_dnf_simps}); val initial_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps nnf_simps addsimps [not_all, not_ex] addsimps map (fn th => th RS sym) (@{thms ex_simps} @ @{thms all_simps})); fun initial_conv ctxt = Simplifier.rewrite (put_simpset initial_ss ctxt); val specl = fold_rev (fn x => fn th => Thm.instantiate' [] [SOME x] (th RS spec)); val cTrp = \<^cterm>\Trueprop\; val cConj = \<^cterm>\HOL.conj\; val (cNot,false_tm) = (\<^cterm>\Not\, \<^cterm>\False\); val assume_Trueprop = Thm.apply cTrp #> Thm.assume; val list_mk_conj = list_mk_binop cConj; val conjs = list_dest_binop cConj; val mk_neg = Thm.apply cNot; fun striplist dest = let fun h acc x = case try dest x of SOME (a,b) => h (h acc b) a | NONE => x::acc in h [] end; fun list_mk_binop b = foldr1 (fn (s,t) => Thm.apply (Thm.apply b s) t); val eq_commute = mk_meta_eq @{thm eq_commute}; fun sym_conv eq = let val (l,r) = Thm.dest_binop eq in Thm.instantiate' [SOME (Thm.ctyp_of_cterm l)] [SOME l, SOME r] eq_commute end; (* FIXME : copied from cqe.ML -- complex QE*) fun conjuncts ct = case Thm.term_of ct of \<^term>\HOL.conj\$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct)) | _ => [ct]; fun fold1 f = foldr1 (uncurry f); fun mk_conj_tab th = let fun h acc th = case Thm.prop_of th of \<^term>\Trueprop\$(\<^term>\HOL.conj\$_$_) => h (h acc (th RS conjunct2)) (th RS conjunct1) | \<^term>\Trueprop\$p => (p,th)::acc in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end; fun is_conj (\<^term>\HOL.conj\$_$_) = true | is_conj _ = false; fun prove_conj tab cjs = case cjs of [c] => if is_conj (Thm.term_of c) then prove_conj tab (conjuncts c) else tab c | c::cs => conjI OF [prove_conj tab [c], prove_conj tab cs]; fun conj_ac_rule eq = let val (l,r) = Thm.dest_equals eq val ctabl = mk_conj_tab (Thm.assume (Thm.apply \<^cterm>\Trueprop\ l)) val ctabr = mk_conj_tab (Thm.assume (Thm.apply \<^cterm>\Trueprop\ r)) fun tabl c = the (Termtab.lookup ctabl (Thm.term_of c)) fun tabr c = the (Termtab.lookup ctabr (Thm.term_of c)) val thl = prove_conj tabl (conjuncts r) |> implies_intr_hyps val thr = prove_conj tabr (conjuncts l) |> implies_intr_hyps val eqI = Thm.instantiate' [] [SOME l, SOME r] @{thm iffI} in Thm.implies_elim (Thm.implies_elim eqI thl) thr |> mk_meta_eq end; (* END FIXME.*) (* Conversion for the equivalence of existential statements where EX quantifiers are rearranged differently *) fun ext ctxt T = Thm.cterm_of ctxt (Const (\<^const_name>\Ex\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) fun mk_ex ctxt v t = Thm.apply (ext ctxt (Thm.typ_of_cterm v)) (Thm.lambda v t) fun choose v th th' = case Thm.concl_of th of \<^term>\Trueprop\ $ (Const(\<^const_name>\Ex\,_)$_) => let val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th val T = Thm.dest_ctyp0 (Thm.ctyp_of_cterm p) val th0 = Conv.fconv_rule (Thm.beta_conversion true) (Thm.instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o Thm.cprop_of) th'] exE) val pv = (Thm.rhs_of o Thm.beta_conversion true) (Thm.apply \<^cterm>\Trueprop\ (Thm.apply p v)) val th1 = Thm.forall_intr v (Thm.implies_intr pv th') in Thm.implies_elim (Thm.implies_elim th0 th) th1 end | _ => error "" (* FIXME ? *) fun simple_choose ctxt v th = choose v (Thm.assume ((Thm.apply \<^cterm>\Trueprop\ o mk_ex ctxt v) (Thm.dest_arg (hd (Thm.chyps_of th))))) th fun mkexi v th = let val p = Thm.lambda v (Thm.dest_arg (Thm.cprop_of th)) in Thm.implies_elim (Conv.fconv_rule (Thm.beta_conversion true) (Thm.instantiate' [SOME (Thm.ctyp_of_cterm v)] [SOME p, SOME v] @{thm exI})) th end fun ex_eq_conv ctxt t = let val (p0,q0) = Thm.dest_binop t val (vs',P) = strip_exists p0 val (vs,_) = strip_exists q0 val th = Thm.assume (Thm.apply \<^cterm>\Trueprop\ P) val th1 = implies_intr_hyps (fold (simple_choose ctxt) vs' (fold mkexi vs th)) val th2 = implies_intr_hyps (fold (simple_choose ctxt) vs (fold mkexi vs' th)) val p = (Thm.dest_arg o Thm.dest_arg1 o Thm.cprop_of) th1 val q = (Thm.dest_arg o Thm.dest_arg o Thm.cprop_of) th1 in Thm.implies_elim (Thm.implies_elim (Thm.instantiate' [] [SOME p, SOME q] iffI) th1) th2 |> mk_meta_eq end; fun getname v = case Thm.term_of v of Free(s,_) => s | Var ((s,_),_) => s | _ => "x" fun mk_eq s t = Thm.apply (Thm.apply \<^cterm>\(\) :: bool \ _\ s) t fun mk_exists ctxt v th = Drule.arg_cong_rule (ext ctxt (Thm.typ_of_cterm v)) (Thm.abstract_rule (getname v) v th) fun simp_ex_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms(39)}) -fun frees t = Drule.cterm_add_frees t []; -fun free_in v t = member op aconvc (frees t) v; +fun free_in v t = member op aconvc (Misc_Legacy.cterm_frees t) v; val vsubst = let fun vsubst (t,v) tm = (Thm.rhs_of o Thm.beta_conversion false) (Thm.apply (Thm.lambda v tm) t) in fold vsubst end; (** main **) fun ring_and_ideal_conv {vars = _, semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), idom, ideal} dest_const mk_const ring_eq_conv ring_normalize_conv = let val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops; val [ring_add_tm, ring_mul_tm, ring_pow_tm] = map Thm.dest_fun2 [add_pat, mul_pat, pow_pat]; val (ring_sub_tm, ring_neg_tm) = (case r_ops of [sub_pat, neg_pat] => (Thm.dest_fun2 sub_pat, Thm.dest_fun neg_pat) |_ => (\<^cterm>\True\, \<^cterm>\True\)); val (field_div_tm, field_inv_tm) = (case f_ops of [div_pat, inv_pat] => (Thm.dest_fun2 div_pat, Thm.dest_fun inv_pat) | _ => (\<^cterm>\True\, \<^cterm>\True\)); val [idom_thm, neq_thm] = idom; val [idl_sub, idl_add0] = if length ideal = 2 then ideal else [eq_commute, eq_commute] fun ring_dest_neg t = let val (l,r) = Thm.dest_comb t in if Term.could_unify(Thm.term_of l, Thm.term_of ring_neg_tm) then r else raise CTERM ("ring_dest_neg", [t]) end fun field_dest_inv t = let val (l,r) = Thm.dest_comb t in if Term.could_unify (Thm.term_of l, Thm.term_of field_inv_tm) then r else raise CTERM ("field_dest_inv", [t]) end val ring_dest_add = dest_binary ring_add_tm; val ring_mk_add = mk_binop ring_add_tm; val ring_dest_sub = dest_binary ring_sub_tm; val ring_dest_mul = dest_binary ring_mul_tm; val ring_mk_mul = mk_binop ring_mul_tm; val field_dest_div = dest_binary field_div_tm; val ring_dest_pow = dest_binary ring_pow_tm; val ring_mk_pow = mk_binop ring_pow_tm ; fun grobvars tm acc = if can dest_const tm then acc else if can ring_dest_neg tm then grobvars (Thm.dest_arg tm) acc else if can ring_dest_pow tm then grobvars (Thm.dest_arg1 tm) acc else if can ring_dest_add tm orelse can ring_dest_sub tm orelse can ring_dest_mul tm then grobvars (Thm.dest_arg1 tm) (grobvars (Thm.dest_arg tm) acc) else if can field_dest_inv tm then let val gvs = grobvars (Thm.dest_arg tm) [] in if null gvs then acc else tm::acc end else if can field_dest_div tm then let val lvs = grobvars (Thm.dest_arg1 tm) acc val gvs = grobvars (Thm.dest_arg tm) [] in if null gvs then lvs else tm::acc end else tm::acc ; fun grobify_term vars tm = ((if not (member (op aconvc) vars tm) then raise CTERM ("Not a variable", [tm]) else [(@1, map (fn i => if i aconvc tm then 1 else 0) vars)]) handle CTERM _ => ((let val x = dest_const tm in if x = @0 then [] else [(x,map (K 0) vars)] end) handle ERROR _ => ((grob_neg(grobify_term vars (ring_dest_neg tm))) handle CTERM _ => ( (grob_inv(grobify_term vars (field_dest_inv tm))) handle CTERM _ => ((let val (l,r) = ring_dest_add tm in grob_add (grobify_term vars l) (grobify_term vars r) end) handle CTERM _ => ((let val (l,r) = ring_dest_sub tm in grob_sub (grobify_term vars l) (grobify_term vars r) end) handle CTERM _ => ((let val (l,r) = ring_dest_mul tm in grob_mul (grobify_term vars l) (grobify_term vars r) end) handle CTERM _ => ( (let val (l,r) = field_dest_div tm in grob_div (grobify_term vars l) (grobify_term vars r) end) handle CTERM _ => ((let val (l,r) = ring_dest_pow tm in grob_pow vars (grobify_term vars l) ((Thm.term_of #> HOLogic.dest_number #> snd) r) end) handle CTERM _ => error "grobify_term: unknown or invalid term"))))))))); val eq_tm = idom_thm |> concl |> Thm.dest_arg |> Thm.dest_arg |> Thm.dest_fun2; val dest_eq = dest_binary eq_tm; fun grobify_equation vars tm = let val (l,r) = dest_binary eq_tm tm in grob_sub (grobify_term vars l) (grobify_term vars r) end; fun grobify_equations tm = let val cjs = conjs tm val rawvars = fold_rev (fn eq => fn a => grobvars (Thm.dest_arg1 eq) (grobvars (Thm.dest_arg eq) a)) cjs [] val vars = sort Thm.term_ord (distinct (op aconvc) rawvars) in (vars,map (grobify_equation vars) cjs) end; val holify_polynomial = let fun holify_varpow (v,n) = if n = 1 then v else ring_mk_pow v (Numeral.mk_cnumber \<^ctyp>\nat\ n) (* FIXME *) fun holify_monomial vars (c,m) = let val xps = map holify_varpow (filter (fn (_,n) => n <> 0) (vars ~~ m)) in end_itlist ring_mk_mul (mk_const c :: xps) end fun holify_polynomial vars p = if null p then mk_const @0 else end_itlist ring_mk_add (map (holify_monomial vars) p) in holify_polynomial end ; fun idom_rule ctxt = simplify (put_simpset HOL_basic_ss ctxt addsimps [idom_thm]); fun prove_nz n = eqF_elim (ring_eq_conv(mk_binop eq_tm (mk_const n) (mk_const @0))); val neq_01 = prove_nz @1; fun neq_rule n th = [prove_nz n, th] MRS neq_thm; fun mk_add th1 = Thm.combination (Drule.arg_cong_rule ring_add_tm th1); fun refute ctxt tm = if tm aconvc false_tm then assume_Trueprop tm else ((let val (nths0,eths0) = List.partition (is_neg o concl) (HOLogic.conj_elims ctxt (assume_Trueprop tm)) val nths = filter (is_eq o Thm.dest_arg o concl) nths0 val eths = filter (is_eq o concl) eths0 in if null eths then let val th1 = end_itlist (fn th1 => fn th2 => idom_rule ctxt (HOLogic.conj_intr ctxt th1 th2)) nths val th2 = Conv.fconv_rule ((Conv.arg_conv #> Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) th1 val conc = th2 |> concl |> Thm.dest_arg val (l,_) = conc |> dest_eq in Thm.implies_intr (Thm.apply cTrp tm) (Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th2)) (HOLogic.mk_obj_eq (Thm.reflexive l))) end else let val (vars,l,cert,noteqth) =( if null nths then let val (vars,pols) = grobify_equations(list_mk_conj(map concl eths)) val (l,cert) = grobner_weak vars pols in (vars,l,cert,neq_01) end else let val nth = end_itlist (fn th1 => fn th2 => idom_rule ctxt (HOLogic.conj_intr ctxt th1 th2)) nths val (vars,pol::pols) = grobify_equations(list_mk_conj(Thm.dest_arg(concl nth)::map concl eths)) val (deg,l,cert) = grobner_strong vars pols pol val th1 = Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) nth val th2 = funpow deg (idom_rule ctxt o HOLogic.conj_intr ctxt th1) neq_01 in (vars,l,cert,th2) end) val cert_pos = map (fn (i,p) => (i,filter (fn (c,_) => c > @0) p)) cert val cert_neg = map (fn (i,p) => (i,map (fn (c,m) => (~ c,m)) (filter (fn (c,_) => c < @0) p))) cert val herts_pos = map (fn (i,p) => (i,holify_polynomial vars p)) cert_pos val herts_neg = map (fn (i,p) => (i,holify_polynomial vars p)) cert_neg fun thm_fn pols = if null pols then Thm.reflexive(mk_const @0) else end_itlist mk_add (map (fn (i,p) => Drule.arg_cong_rule (Thm.apply ring_mul_tm p) (nth eths i |> mk_meta_eq)) pols) val th1 = thm_fn herts_pos val th2 = thm_fn herts_neg val th3 = HOLogic.conj_intr ctxt (HOLogic.mk_obj_eq (mk_add (Thm.symmetric th1) th2)) noteqth val th4 = Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv o Conv.binop_conv) ring_normalize_conv) (neq_rule l th3) val (l, _) = dest_eq(Thm.dest_arg(concl th4)) in Thm.implies_intr (Thm.apply cTrp tm) (Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th4)) (HOLogic.mk_obj_eq (Thm.reflexive l))) end end) handle ERROR _ => raise CTERM ("Groebner-refute: unable to refute",[tm])) fun ring ctxt tm = let fun mk_forall x p = let val T = Thm.typ_of_cterm x; val all = Thm.cterm_of ctxt (Const (\<^const_name>\All\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) in Thm.apply all (Thm.lambda x p) end - val avs = Drule.cterm_add_frees tm [] + val avs = Misc_Legacy.cterm_frees tm val P' = fold mk_forall avs tm val th1 = initial_conv ctxt (mk_neg P') val (evs,bod) = strip_exists(concl th1) in if is_forall bod then raise CTERM("ring: non-universal formula",[tm]) else let val th1a = weak_dnf_conv ctxt bod val boda = concl th1a val th2a = refute_disj (refute ctxt) boda val th2b = [HOLogic.mk_obj_eq th1a, (th2a COMP notI) COMP PFalse'] MRS trans val th2 = fold (fn v => fn th => (Thm.forall_intr v th) COMP allI) evs (th2b RS PFalse) val th3 = Thm.equal_elim (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [not_ex RS sym]) (th2 |> Thm.cprop_of)) th2 in specl avs ([[[HOLogic.mk_obj_eq th1, th3 RS PFalse'] MRS trans] MRS PFalse] MRS notnotD) end end fun ideal tms tm ord = let val rawvars = fold_rev grobvars (tm::tms) [] val vars = sort ord (distinct (fn (x,y) => (Thm.term_of x) aconv (Thm.term_of y)) rawvars) val pols = map (grobify_term vars) tms val pol = grobify_term vars tm val cert = grobner_ideal vars pols pol in map_range (fn n => these (AList.lookup (op =) cert n) |> holify_polynomial vars) (length pols) end fun poly_eq_conv t = let val (a,b) = Thm.dest_binop t in Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv ring_normalize_conv)) (Thm.instantiate' [] [SOME a, SOME b] idl_sub) end val poly_eq_simproc = let fun proc ct = let val th = poly_eq_conv ct in if Thm.is_reflexive th then NONE else SOME th end in Simplifier.cert_simproc (Thm.theory_of_thm idl_sub) "poly_eq_simproc" {lhss = [Thm.term_of (Thm.lhs_of idl_sub)], proc = fn _ => fn _ => proc} end; val poly_eq_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms simp_thms} addsimprocs [poly_eq_simproc]) local fun is_defined v t = let val mons = striplist(dest_binary ring_add_tm) t in member (op aconvc) mons v andalso forall (fn m => v aconvc m - orelse not(member (op aconvc) (Drule.cterm_add_frees m []) v)) mons + orelse not(member (op aconvc) (Misc_Legacy.cterm_frees m) v)) mons end fun isolate_variable vars tm = let val th = poly_eq_conv tm val th' = (sym_conv then_conv poly_eq_conv) tm val (v,th1) = case find_first(fn v=> is_defined v (Thm.dest_arg1 (Thm.rhs_of th))) vars of SOME v => (v,th') | NONE => (the (find_first (fn v => is_defined v (Thm.dest_arg1 (Thm.rhs_of th'))) vars) ,th) val th2 = Thm.transitive th1 (Thm.instantiate' [] [(SOME o Thm.dest_arg1 o Thm.rhs_of) th1, SOME v] idl_add0) in Conv.fconv_rule(funpow 2 Conv.arg_conv ring_normalize_conv) th2 end in fun unwind_polys_conv ctxt tm = let val (vars,bod) = strip_exists tm val cjs = striplist (dest_binary \<^cterm>\HOL.conj\) bod val th1 = (the (get_first (try (isolate_variable vars)) cjs) handle Option.Option => raise CTERM ("unwind_polys_conv",[tm])) val eq = Thm.lhs_of th1 val bod' = list_mk_binop \<^cterm>\HOL.conj\ (eq::(remove (op aconvc) eq cjs)) val th2 = conj_ac_rule (mk_eq bod bod') val th3 = Thm.transitive th2 (Drule.binop_cong_rule \<^cterm>\HOL.conj\ th1 (Thm.reflexive (Thm.dest_arg (Thm.rhs_of th2)))) val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3)) val th4 = Conv.fconv_rule (Conv.arg_conv (simp_ex_conv ctxt)) (mk_exists ctxt v th3) val th5 = ex_eq_conv ctxt (mk_eq tm (fold (mk_ex ctxt) (remove op aconvc v vars) (Thm.lhs_of th4))) in Thm.transitive th5 (fold (mk_exists ctxt) (remove op aconvc v vars) th4) end; end local fun scrub_var v m = let val ps = striplist ring_dest_mul m val ps' = remove op aconvc v ps in if null ps' then one_tm else fold1 ring_mk_mul ps' end fun find_multipliers v mons = let val mons1 = filter (fn m => free_in v m) mons val mons2 = map (scrub_var v) mons1 in if null mons2 then zero_tm else fold1 ring_mk_add mons2 end fun isolate_monomials vars tm = let val (cmons,vmons) = - List.partition (fn m => null (inter (op aconvc) vars (frees m))) + List.partition (fn m => null (inter (op aconvc) vars (Misc_Legacy.cterm_frees m))) (striplist ring_dest_add tm) val cofactors = map (fn v => find_multipliers v vmons) vars val cnc = if null cmons then zero_tm else Thm.apply ring_neg_tm (list_mk_binop ring_add_tm cmons) in (cofactors,cnc) end; fun isolate_variables evs ps eq = let val vars = filter (fn v => free_in v eq) evs val (qs,p) = isolate_monomials vars eq val rs = ideal (qs @ ps) p Thm.term_ord in (eq, take (length qs) rs ~~ vars) end; fun subst_in_poly i p = Thm.rhs_of (ring_normalize_conv (vsubst i p)); in fun solve_idealism evs ps eqs = if null evs then [] else let val (eq,cfs) = get_first (try (isolate_variables evs ps)) eqs |> the val evs' = subtract op aconvc evs (map snd cfs) val eqs' = map (subst_in_poly cfs) (remove op aconvc eq eqs) in cfs @ solve_idealism evs' ps eqs' end; end; in {ring_conv = ring, simple_ideal = ideal, multi_ideal = solve_idealism, poly_eq_ss = poly_eq_ss, unwind_conv = unwind_polys_conv} end; fun find_term bounds tm = (case Thm.term_of tm of Const (\<^const_name>\HOL.eq\, T) $ _ $ _ => if domain_type T = HOLogic.boolT then find_args bounds tm else Thm.dest_arg tm | Const (\<^const_name>\Not\, _) $ _ => find_term bounds (Thm.dest_arg tm) | Const (\<^const_name>\All\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\Ex\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\HOL.conj\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\HOL.disj\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\HOL.implies\, _) $ _ $ _ => find_args bounds tm | \<^term>\Pure.imp\ $_$_ => find_args bounds tm | Const("(==)",_)$_$_ => find_args bounds tm (* FIXME proper const name *) | \<^term>\Trueprop\$_ => find_term bounds (Thm.dest_arg tm) | _ => raise TERM ("find_term", [])) and find_args bounds tm = let val (t, u) = Thm.dest_binop tm in (find_term bounds t handle TERM _ => find_term bounds u) end and find_body bounds b = let val (_, b') = Thm.dest_abs (SOME (Name.bound bounds)) b in find_term (bounds + 1) b' end; fun get_ring_ideal_convs ctxt form = case try (find_term 0) form of NONE => NONE | SOME tm => (case Semiring_Normalizer.match ctxt tm of NONE => NONE | SOME (res as (theory, {is_const = _, dest_const, mk_const, conv = ring_eq_conv})) => SOME (ring_and_ideal_conv theory dest_const (mk_const (Thm.ctyp_of_cterm tm)) (ring_eq_conv ctxt) (Semiring_Normalizer.semiring_normalize_wrapper ctxt res))) fun presimplify ctxt add_thms del_thms = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (Named_Theorems.get ctxt \<^named_theorems>\algebra\) delsimps del_thms addsimps add_thms); fun ring_tac add_ths del_ths ctxt = Object_Logic.full_atomize_tac ctxt THEN' presimplify ctxt add_ths del_ths THEN' CSUBGOAL (fn (p, i) => resolve_tac ctxt [let val form = Object_Logic.dest_judgment ctxt p in case get_ring_ideal_convs ctxt form of NONE => Thm.reflexive form | SOME thy => #ring_conv thy ctxt form end] i handle TERM _ => no_tac | CTERM _ => no_tac | THM _ => no_tac); local fun lhs t = case Thm.term_of t of Const(\<^const_name>\HOL.eq\,_)$_$_ => Thm.dest_arg1 t | _=> raise CTERM ("ideal_tac - lhs",[t]) fun exitac _ NONE = no_tac | exitac ctxt (SOME y) = resolve_tac ctxt [Thm.instantiate' [SOME (Thm.ctyp_of_cterm y)] [NONE,SOME y] exI] 1 val claset = claset_of \<^context> in fun ideal_tac add_ths del_ths ctxt = presimplify ctxt add_ths del_ths THEN' CSUBGOAL (fn (p, i) => case get_ring_ideal_convs ctxt p of NONE => no_tac | SOME thy => let fun poly_exists_tac {asms = asms, concl = concl, prems = prems, params = _, context = ctxt, schematics = _} = let val (evs,bod) = strip_exists (Thm.dest_arg concl) val ps = map_filter (try (lhs o Thm.dest_arg)) asms val cfs = (map swap o #multi_ideal thy evs ps) (map Thm.dest_arg1 (conjuncts bod)) val ws = map (exitac ctxt o AList.lookup op aconvc cfs) evs in EVERY (rev ws) THEN Method.insert_tac ctxt prems 1 THEN ring_tac add_ths del_ths ctxt 1 end in clarify_tac (put_claset claset ctxt) i THEN Object_Logic.full_atomize_tac ctxt i THEN asm_full_simp_tac (put_simpset (#poly_eq_ss thy) ctxt) i THEN clarify_tac (put_claset claset ctxt) i THEN (REPEAT (CONVERSION (#unwind_conv thy ctxt) i)) THEN SUBPROOF poly_exists_tac ctxt i end handle TERM _ => no_tac | CTERM _ => no_tac | THM _ => no_tac); end; fun algebra_tac add_ths del_ths ctxt i = ring_tac add_ths del_ths ctxt i ORELSE ideal_tac add_ths del_ths ctxt i end; diff --git a/src/Pure/General/buffer.ML b/src/Pure/General/buffer.ML --- a/src/Pure/General/buffer.ML +++ b/src/Pure/General/buffer.ML @@ -1,41 +1,46 @@ (* Title: Pure/General/buffer.ML Author: Makarius Scalable text buffers. *) signature BUFFER = sig type T val empty: T val is_empty: T -> bool + val add: string -> T -> T val content: T -> string - val add: string -> T -> T + val build: (T -> T) -> T + val build_content: (T -> T) -> string val output: T -> (string -> unit) -> unit val markup: Markup.T -> (T -> T) -> T -> T end; structure Buffer: BUFFER = struct abstype T = Buffer of string list with val empty = Buffer []; fun is_empty (Buffer xs) = null xs; fun add "" buf = buf | add x (Buffer xs) = Buffer (x :: xs); fun content (Buffer xs) = implode (rev xs); +fun build (f: T -> T) = f empty; +fun build_content f = build f |> content; + fun output (Buffer xs) out = List.app out (rev xs); end; fun markup m body = let val (bg, en) = Markup.output m in add bg #> body #> add en end; end; diff --git a/src/Pure/General/graph_display.ML b/src/Pure/General/graph_display.ML --- a/src/Pure/General/graph_display.ML +++ b/src/Pure/General/graph_display.ML @@ -1,101 +1,102 @@ (* Title: Pure/General/graph_display.ML Author: Makarius Support for graph display. *) signature GRAPH_DISPLAY = sig type node val content_node: string -> Pretty.T list -> node val session_node: {name: string, unfold: bool, directory: string, path: string} -> node type entry = (string * node) * string list val display_graph: entry list -> unit val display_graph_old: entry list -> unit end; structure Graph_Display: GRAPH_DISPLAY = struct (* graph entries *) datatype node = Node of {name: string, content: Pretty.T list, unfold: bool, directory: string, path: string}; fun content_node name content = Node {name = name, content = content, unfold = true, directory = "", path = ""}; fun session_node {name, unfold, directory, path} = Node {name = name, content = [], unfold = unfold, directory = directory, path = path}; type entry = (string * node) * string list; (* display graph *) local fun encode_node (Node {name, content, ...}) = (name, content) |> let open XML.Encode in pair string (YXML.parse_body o Pretty.symbolic_string_of o Pretty.chunks) end; val encode_graph = let open XML.Encode in list (pair (pair string encode_node) (list string)) end; in fun display_graph entries = let val ((bg1, bg2), en) = YXML.output_markup_elem (Active.make_markup Markup.graphviewN {implicit = false, properties = []}); in writeln ("See " ^ bg1 ^ YXML.string_of_body (encode_graph entries) ^ bg2 ^ "graph" ^ en) end; end; (* support for old browser *) local structure Graph = Graph(type key = string * string val ord = prod_ord string_ord string_ord); fun build_graph entries = let val ident_names = - fold (fn ((ident, Node {name, ...}), _) => Symtab.update_new (ident, (name, ident))) - entries Symtab.empty; + Symtab.build + (entries |> fold (fn ((ident, Node {name, ...}), _) => + Symtab.update_new (ident, (name, ident)))); val the_key = the o Symtab.lookup ident_names; in Graph.empty |> fold (fn ((ident, node), _) => Graph.new_node (the_key ident, node)) entries |> fold (fn ((ident, _), parents) => fold (fn parent => Graph.add_edge (the_key parent, the_key ident)) parents) entries end; val sort_graph = build_graph #> (fn graph => Graph.topological_order graph |> map (fn key => let val (_, (node, (preds, _))) = Graph.get_entry graph key in ((#2 key, node), map #2 (Graph.Keys.dest preds)) end)); val encode_browser = sort_graph #> map (fn ((key, Node {name, unfold, content = _, directory, path}), parents) => "\"" ^ name ^ "\" \"" ^ key ^ "\" \"" ^ directory ^ (if unfold then "\" + \"" else "\" \"") ^ path ^ "\" > " ^ space_implode " " (map quote parents) ^ " ;") #> cat_lines; in fun display_graph_old entries = let val ((bg1, bg2), en) = YXML.output_markup_elem (Active.make_markup Markup.browserN {implicit = false, properties = []}); in writeln ("See " ^ bg1 ^ encode_browser entries ^ bg2 ^ "graph" ^ en) end; end; end; diff --git a/src/Pure/General/pretty.ML b/src/Pure/General/pretty.ML --- a/src/Pure/General/pretty.ML +++ b/src/Pure/General/pretty.ML @@ -1,427 +1,427 @@ (* Title: Pure/General/pretty.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Markus Wenzel, TU Munich Generic pretty printing module. Loosely based on D. C. Oppen, "Pretty Printing", ACM Transactions on Programming Languages and Systems (1980), 465-483. The object to be printed is given as a tree with indentation and line breaking information. A "break" inserts a newline if the text until the next break is too long to fit on the current line. After the newline, text is indented to the level of the enclosing block. Normally, if a block is broken then all enclosing blocks will also be broken. The stored length of a block is used in break_dist (to treat each inner block as a unit for breaking). *) signature PRETTY = sig val default_indent: string -> int -> Output.output val add_mode: string -> (string -> int -> Output.output) -> unit type T val make_block: {markup: Markup.output, consistent: bool, indent: int} -> T list -> T val markup_block: {markup: Markup.T, consistent: bool, indent: int} -> T list -> T val str: string -> T val brk: int -> T val brk_indent: int -> int -> T val fbrk: T val breaks: T list -> T list val fbreaks: T list -> T list val blk: int * T list -> T val block: T list -> T val strs: string list -> T val markup: Markup.T -> T list -> T val mark: Markup.T -> T -> T val mark_str: Markup.T * string -> T val marks_str: Markup.T list * string -> T val item: T list -> T val text_fold: T list -> T val keyword1: string -> T val keyword2: string -> T val text: string -> T list val paragraph: T list -> T val para: string -> T val quote: T -> T val cartouche: T -> T val separate: string -> T list -> T list val commas: T list -> T list val enclose: string -> string -> T list -> T val enum: string -> string -> string -> T list -> T val position: Position.T -> T val here: Position.T -> T list val list: string -> string -> T list -> T val str_list: string -> string -> string list -> T val big_list: string -> T list -> T val indent: int -> T -> T val unbreakable: T -> T val margin_default: int Unsynchronized.ref val regularN: string val symbolicN: string val output_buffer: int option -> T -> Buffer.T val output: int option -> T -> Output.output val string_of_margin: int -> T -> string val string_of: T -> string val writeln: T -> unit val symbolic_output: T -> Output.output val symbolic_string_of: T -> string val unformatted_string_of: T -> string val markup_chunks: Markup.T -> T list -> T val chunks: T list -> T val chunks2: T list -> T val block_enclose: T * T -> T list -> T val writeln_chunks: T list -> unit val writeln_chunks2: T list -> unit val to_ML: FixedInt.int -> T -> ML_Pretty.pretty val from_ML: ML_Pretty.pretty -> T val to_polyml: T -> PolyML.pretty val from_polyml: PolyML.pretty -> T end; structure Pretty: PRETTY = struct (** print mode operations **) fun default_indent (_: string) = Symbol.spaces; local val default = {indent = default_indent}; val modes = Synchronized.var "Pretty.modes" (Symtab.make [("", default)]); in fun add_mode name indent = Synchronized.change modes (fn tab => (if not (Symtab.defined tab name) then () else warning ("Redefining pretty mode " ^ quote name); Symtab.update (name, {indent = indent}) tab)); fun get_mode () = the_default default (Library.get_first (Symtab.lookup (Synchronized.value modes)) (print_mode_value ())); end; fun mode_indent x y = #indent (get_mode ()) x y; val output_spaces = Output.output o Symbol.spaces; val add_indent = Buffer.add o output_spaces; (** printing items: compound phrases, strings, and breaks **) val force_nat = Integer.max 0; abstype T = Block of Markup.output * bool * int * T list * int (*markup output, consistent, indentation, body, length*) | Break of bool * int * int (*mandatory flag, width if not taken, extra indentation if taken*) | Str of Output.output * int (*text, length*) with fun length (Block (_, _, _, _, len)) = len | length (Break (_, wd, _)) = wd | length (Str (_, len)) = len; fun make_block {markup, consistent, indent} body = let val indent' = force_nat indent; fun body_length prts len = let val (line, rest) = chop_prefix (fn Break (true, _, _) => false | _ => true) prts; val len' = Int.max (fold (Integer.add o length) line 0, len); in (case rest of Break (true, _, ind) :: rest' => body_length (Break (false, indent' + ind, 0) :: rest') len' | [] => len') end; in Block (markup, consistent, indent', body, body_length body 0) end; fun markup_block {markup, consistent, indent} es = make_block {markup = Markup.output markup, consistent = consistent, indent = indent} es; (** derived operations to create formatting expressions **) val str = Output.output_width ##> force_nat #> Str; fun brk wd = Break (false, force_nat wd, 0); fun brk_indent wd ind = Break (false, force_nat wd, ind); val fbrk = Break (true, 1, 0); fun breaks prts = Library.separate (brk 1) prts; fun fbreaks prts = Library.separate fbrk prts; fun blk (indent, es) = markup_block {markup = Markup.empty, consistent = false, indent = indent} es; fun block prts = blk (2, prts); val strs = block o breaks o map str; fun markup m = markup_block {markup = m, consistent = false, indent = 0}; fun mark m prt = if m = Markup.empty then prt else markup m [prt]; fun mark_str (m, s) = mark m (str s); fun marks_str (ms, s) = fold_rev mark ms (str s); val item = markup Markup.item; val text_fold = markup Markup.text_fold; fun keyword1 name = mark_str (Markup.keyword1, name); fun keyword2 name = mark_str (Markup.keyword2, name); val text = breaks o map str o Symbol.explode_words; val paragraph = markup Markup.paragraph; val para = paragraph o text; fun quote prt = blk (1, [str "\"", prt, str "\""]); fun cartouche prt = blk (1, [str Symbol.open_, prt, str Symbol.close]); fun separate sep prts = flat (Library.separate [str sep, brk 1] (map single prts)); val commas = separate ","; fun enclose lpar rpar prts = block (str lpar :: (prts @ [str rpar])); fun enum sep lpar rpar prts = enclose lpar rpar (separate sep prts); val position = enum "," "{" "}" o map (fn (x, y) => str (x ^ "=" ^ y)) o Position.properties_of; fun here pos = let val props = Position.properties_of pos; val (s1, s2) = Position.here_strs pos; in if s2 = "" then [] else [str s1, mark_str (Markup.properties props Markup.position, s2)] end; val list = enum ","; fun str_list lpar rpar strs = list lpar rpar (map str strs); fun big_list name prts = block (fbreaks (str name :: prts)); fun indent 0 prt = prt | indent n prt = blk (0, [str (Symbol.spaces n), prt]); fun unbreakable (Block (m, consistent, indent, es, len)) = Block (m, consistent, indent, map unbreakable es, len) | unbreakable (Break (_, wd, _)) = Str (output_spaces wd, wd) | unbreakable (e as Str _) = e; (** formatting **) (* formatted output *) local type text = {tx: Buffer.T, ind: Buffer.T, pos: int, nl: int}; val empty: text = {tx = Buffer.empty, ind = Buffer.empty, pos = 0, nl = 0}; fun newline {tx, ind = _, pos = _, nl} : text = {tx = Buffer.add (Output.output "\n") tx, ind = Buffer.empty, pos = 0, nl = nl + 1}; fun control s {tx, ind, pos: int, nl} : text = {tx = Buffer.add s tx, ind = ind, pos = pos, nl = nl}; fun string (s, len) {tx, ind, pos: int, nl} : text = {tx = Buffer.add s tx, ind = Buffer.add s ind, pos = pos + len, nl = nl}; fun blanks wd = string (output_spaces wd, wd); fun indentation (buf, len) {tx, ind, pos, nl} : text = let val s = Buffer.content buf in {tx = Buffer.add (mode_indent s len) tx, ind = Buffer.add s ind, pos = pos + len, nl = nl} end; (*Add the lengths of the expressions until the next Break; if no Break then include "after", to account for text following this block.*) fun break_dist (Break _ :: _, _) = 0 | break_dist (e :: es, after) = length e + break_dist (es, after) | break_dist ([], after) = after; val force_break = fn Break (false, wd, ind) => Break (true, wd, ind) | e => e; val force_all = map force_break; (*Search for the next break (at this or higher levels) and force it to occur.*) fun force_next [] = [] | force_next ((e as Break _) :: es) = force_break e :: es | force_next (e :: es) = e :: force_next es; in fun formatted margin input = let val breakgain = margin div 20; (*minimum added space required of a break*) val emergencypos = margin div 2; (*position too far to right*) (*es is list of expressions to print; blockin is the indentation of the current block; after is the width of the following context until next break.*) fun format ([], _, _) text = text | format (e :: es, block as (_, blockin), after) (text as {ind, pos, nl, ...}) = (case e of Block ((bg, en), consistent, indent, bes, blen) => let val pos' = pos + indent; val pos'' = pos' mod emergencypos; val block' = if pos' < emergencypos then (ind |> add_indent indent, pos') else (add_indent pos'' Buffer.empty, pos''); val d = break_dist (es, after) val bes' = if consistent andalso pos + blen > margin - d then force_all bes else bes; val btext: text = text |> control bg |> format (bes', block', d) |> control en; (*if this block was broken then force the next break*) val es' = if nl < #nl btext then force_next es else es; in format (es', block, after) btext end | Break (force, wd, ind) => (*no break if text to next break fits on this line or if breaking would add only breakgain to space*) format (es, block, after) (if not force andalso pos + wd <= Int.max (margin - break_dist (es, after), blockin + breakgain) then text |> blanks wd (*just insert wd blanks*) else text |> newline |> indentation block |> blanks ind) | Str str => format (es, block, after) (string str text)); in #tx (format ([input], (Buffer.empty, 0), 0) empty) end; end; (* special output *) (*symbolic markup -- no formatting*) -fun symbolic prt = +val symbolic = let fun out (Block ((bg, en), _, _, [], _)) = Buffer.add bg #> Buffer.add en | out (Block ((bg, en), consistent, indent, prts, _)) = Buffer.add bg #> Buffer.markup (Markup.block consistent indent) (fold out prts) #> Buffer.add en | out (Break (false, wd, ind)) = Buffer.markup (Markup.break wd ind) (Buffer.add (output_spaces wd)) | out (Break (true, _, _)) = Buffer.add (Output.output "\n") | out (Str (s, _)) = Buffer.add s; - in out prt Buffer.empty end; + in Buffer.build o out end; (*unformatted output*) -fun unformatted prt = +val unformatted = let fun out (Block ((bg, en), _, _, prts, _)) = Buffer.add bg #> fold out prts #> Buffer.add en | out (Break (_, wd, _)) = Buffer.add (output_spaces wd) | out (Str (s, _)) = Buffer.add s; - in out prt Buffer.empty end; + in Buffer.build o out end; (* output interfaces *) val margin_default = Unsynchronized.ref ML_Pretty.default_margin; (*right margin, or page width*) val regularN = "pretty_regular"; val symbolicN = "pretty_symbolic"; fun output_buffer margin prt = if print_mode_active symbolicN andalso not (print_mode_active regularN) then symbolic prt else formatted (the_default (! margin_default) margin) prt; val output = Buffer.content oo output_buffer; fun string_of_margin margin = Output.escape o output (SOME margin); val string_of = Output.escape o output NONE; val writeln = Output.writeln o string_of; val symbolic_output = Buffer.content o symbolic; val symbolic_string_of = Output.escape o symbolic_output; val unformatted_string_of = Output.escape o Buffer.content o unformatted; (* chunks *) fun markup_chunks m prts = markup m (fbreaks (map (text_fold o single) prts)); val chunks = markup_chunks Markup.empty; fun chunks2 prts = (case try split_last prts of NONE => blk (0, []) | SOME (prefix, last) => blk (0, maps (fn prt => [text_fold [prt, fbrk], fbrk]) prefix @ [text_fold [last]])); fun block_enclose (prt1, prt2) prts = chunks [block (fbreaks (prt1 :: prts)), prt2]; fun string_of_text_fold prt = string_of prt |> Markup.markup Markup.text_fold; fun writeln_chunks prts = Output.writelns (Library.separate "\n" (map string_of_text_fold prts)); fun writeln_chunks2 prts = (case try split_last prts of NONE => () | SOME (prefix, last) => (map (fn prt => Markup.markup Markup.text_fold (string_of prt ^ "\n") ^ "\n") prefix @ [string_of_text_fold last]) |> Output.writelns); (** toplevel pretty printing **) fun to_ML 0 (Block _) = ML_Pretty.str "..." | to_ML depth (Block (m, consistent, indent, prts, _)) = ML_Pretty.Block (m, consistent, FixedInt.fromInt indent, map (to_ML (depth - 1)) prts) | to_ML _ (Break (force, wd, ind)) = ML_Pretty.Break (force, FixedInt.fromInt wd, FixedInt.fromInt ind) | to_ML _ (Str (s, len)) = ML_Pretty.String (s, FixedInt.fromInt len); fun from_ML (ML_Pretty.Block (m, consistent, indent, prts)) = make_block {markup = m, consistent = consistent, indent = FixedInt.toInt indent} (map from_ML prts) | from_ML (ML_Pretty.Break (force, wd, ind)) = Break (force, FixedInt.toInt wd, FixedInt.toInt ind) | from_ML (ML_Pretty.String (s, len)) = Str (s, force_nat (FixedInt.toInt len)); val to_polyml = to_ML ~1 #> ML_Pretty.to_polyml; val from_polyml = ML_Pretty.from_polyml #> from_ML; end; val _ = ML_system_pp (fn d => fn _ => ML_Pretty.to_polyml o to_ML (d + 1) o quote); val _ = ML_system_pp (fn _ => fn _ => to_polyml o position); end; structure ML_Pretty = struct open ML_Pretty; val string_of_polyml = Pretty.string_of o Pretty.from_polyml; end; diff --git a/src/Pure/General/table.ML b/src/Pure/General/table.ML --- a/src/Pure/General/table.ML +++ b/src/Pure/General/table.ML @@ -1,458 +1,479 @@ (* Title: Pure/General/table.ML Author: Markus Wenzel and Stefan Berghofer, TU Muenchen Generic tables. Efficient purely functional implementation using balanced 2-3 trees. *) signature KEY = sig type key val ord: key ord end; signature TABLE = sig type key type 'a table exception DUP of key exception SAME exception UNDEF of key val empty: 'a table + val build: ('a table -> 'a table) -> 'a table val is_empty: 'a table -> bool val is_single: 'a table -> bool val map: (key -> 'a -> 'b) -> 'a table -> 'b table val fold: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a val fold_rev: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a + val size: 'a table -> int val dest: 'a table -> (key * 'a) list val keys: 'a table -> key list val min: 'a table -> (key * 'a) option val max: 'a table -> (key * 'a) option - val get_first: (key * 'a -> 'b option) -> 'a table -> 'b option val exists: (key * 'a -> bool) -> 'a table -> bool val forall: (key * 'a -> bool) -> 'a table -> bool + val get_first: (key * 'a -> 'b option) -> 'a table -> 'b option val lookup_key: 'a table -> key -> (key * 'a) option val lookup: 'a table -> key -> 'a option val defined: 'a table -> key -> bool val update: key * 'a -> 'a table -> 'a table val update_new: key * 'a -> 'a table -> 'a table (*exception DUP*) val default: key * 'a -> 'a table -> 'a table val map_entry: key -> ('a -> 'a) (*exception SAME*) -> 'a table -> 'a table val map_default: key * 'a -> ('a -> 'a) -> 'a table -> 'a table val make: (key * 'a) list -> 'a table (*exception DUP*) val join: (key -> 'a * 'a -> 'a) (*exception SAME*) -> 'a table * 'a table -> 'a table (*exception DUP*) val merge: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table (*exception DUP*) val delete: key -> 'a table -> 'a table (*exception UNDEF*) val delete_safe: key -> 'a table -> 'a table val member: ('b * 'a -> bool) -> 'a table -> key * 'b -> bool val insert: ('a * 'a -> bool) -> key * 'a -> 'a table -> 'a table (*exception DUP*) val remove: ('b * 'a -> bool) -> key * 'b -> 'a table -> 'a table val lookup_list: 'a list table -> key -> 'a list val cons_list: key * 'a -> 'a list table -> 'a list table val insert_list: ('a * 'a -> bool) -> key * 'a -> 'a list table -> 'a list table val remove_list: ('b * 'a -> bool) -> key * 'b -> 'a list table -> 'a list table val update_list: ('a * 'a -> bool) -> key * 'a -> 'a list table -> 'a list table val make_list: (key * 'a) list -> 'a list table val dest_list: 'a list table -> (key * 'a) list val merge_list: ('a * 'a -> bool) -> 'a list table * 'a list table -> 'a list table type set = unit table val insert_set: key -> set -> set val remove_set: key -> set -> set val make_set: key list -> set + val subset: set * set -> bool + val eq_set: set * set -> bool end; functor Table(Key: KEY): TABLE = struct (* datatype table *) type key = Key.key; datatype 'a table = Empty | Branch2 of 'a table * (key * 'a) * 'a table | Branch3 of 'a table * (key * 'a) * 'a table * (key * 'a) * 'a table; exception DUP of key; (* empty and single *) val empty = Empty; +fun build (f: 'a table -> 'a table) = f empty; + fun is_empty Empty = true | is_empty _ = false; fun is_single (Branch2 (Empty, _, Empty)) = true | is_single _ = false; (* map and fold combinators *) fun map_table f = let fun map Empty = Empty | map (Branch2 (left, (k, x), right)) = Branch2 (map left, (k, f k x), map right) | map (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = Branch3 (map left, (k1, f k1 x1), map mid, (k2, f k2 x2), map right); in map end; fun fold_table f = let fun fold Empty x = x | fold (Branch2 (left, p, right)) x = fold right (f p (fold left x)) | fold (Branch3 (left, p1, mid, p2, right)) x = fold right (f p2 (fold mid (f p1 (fold left x)))); in fold end; fun fold_rev_table f = let fun fold Empty x = x | fold (Branch2 (left, p, right)) x = fold left (f p (fold right x)) | fold (Branch3 (left, p1, mid, p2, right)) x = fold left (f p1 (fold mid (f p2 (fold right x)))); in fold end; +fun size tab = fold_table (fn _ => fn n => n + 1) tab 0; fun dest tab = fold_rev_table cons tab []; fun keys tab = fold_rev_table (cons o #1) tab []; (* min/max entries *) fun min Empty = NONE | min (Branch2 (Empty, p, _)) = SOME p | min (Branch3 (Empty, p, _, _, _)) = SOME p | min (Branch2 (left, _, _)) = min left | min (Branch3 (left, _, _, _, _)) = min left; fun max Empty = NONE | max (Branch2 (_, p, Empty)) = SOME p | max (Branch3 (_, _, _, p, Empty)) = SOME p | max (Branch2 (_, _, right)) = max right | max (Branch3 (_, _, _, _, right)) = max right; +(* exists and forall *) + +fun exists pred = + let + fun ex Empty = false + | ex (Branch2 (left, (k, x), right)) = + ex left orelse pred (k, x) orelse ex right + | ex (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = + ex left orelse pred (k1, x1) orelse ex mid orelse pred (k2, x2) orelse ex right; + in ex end; + +fun forall pred = not o exists (not o pred); + + (* get_first *) fun get_first f = let fun get Empty = NONE | get (Branch2 (left, (k, x), right)) = (case get left of NONE => (case f (k, x) of NONE => get right | some => some) | some => some) | get (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = (case get left of NONE => (case f (k1, x1) of NONE => (case get mid of NONE => (case f (k2, x2) of NONE => get right | some => some) | some => some) | some => some) | some => some); in get end; -fun exists pred = is_some o get_first (fn entry => if pred entry then SOME () else NONE); -fun forall pred = not o exists (not o pred); - (* lookup *) fun lookup tab key = let fun look Empty = NONE | look (Branch2 (left, (k, x), right)) = (case Key.ord (key, k) of LESS => look left | EQUAL => SOME x | GREATER => look right) | look (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = (case Key.ord (key, k1) of LESS => look left | EQUAL => SOME x1 | GREATER => (case Key.ord (key, k2) of LESS => look mid | EQUAL => SOME x2 | GREATER => look right)); in look tab end; fun lookup_key tab key = let fun look Empty = NONE | look (Branch2 (left, (k, x), right)) = (case Key.ord (key, k) of LESS => look left | EQUAL => SOME (k, x) | GREATER => look right) | look (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = (case Key.ord (key, k1) of LESS => look left | EQUAL => SOME (k1, x1) | GREATER => (case Key.ord (key, k2) of LESS => look mid | EQUAL => SOME (k2, x2) | GREATER => look right)); in look tab end; fun defined tab key = let fun def Empty = false | def (Branch2 (left, (k, x), right)) = (case Key.ord (key, k) of LESS => def left | EQUAL => true | GREATER => def right) | def (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = (case Key.ord (key, k1) of LESS => def left | EQUAL => true | GREATER => (case Key.ord (key, k2) of LESS => def mid | EQUAL => true | GREATER => def right)); in def tab end; (* modify *) datatype 'a growth = Stay of 'a table | Sprout of 'a table * (key * 'a) * 'a table; exception SAME; fun modify key f tab = let fun modfy Empty = Sprout (Empty, (key, f NONE), Empty) | modfy (Branch2 (left, p as (k, x), right)) = (case Key.ord (key, k) of LESS => (case modfy left of Stay left' => Stay (Branch2 (left', p, right)) | Sprout (left1, q, left2) => Stay (Branch3 (left1, q, left2, p, right))) | EQUAL => Stay (Branch2 (left, (k, f (SOME x)), right)) | GREATER => (case modfy right of Stay right' => Stay (Branch2 (left, p, right')) | Sprout (right1, q, right2) => Stay (Branch3 (left, p, right1, q, right2)))) | modfy (Branch3 (left, p1 as (k1, x1), mid, p2 as (k2, x2), right)) = (case Key.ord (key, k1) of LESS => (case modfy left of Stay left' => Stay (Branch3 (left', p1, mid, p2, right)) | Sprout (left1, q, left2) => Sprout (Branch2 (left1, q, left2), p1, Branch2 (mid, p2, right))) | EQUAL => Stay (Branch3 (left, (k1, f (SOME x1)), mid, p2, right)) | GREATER => (case Key.ord (key, k2) of LESS => (case modfy mid of Stay mid' => Stay (Branch3 (left, p1, mid', p2, right)) | Sprout (mid1, q, mid2) => Sprout (Branch2 (left, p1, mid1), q, Branch2 (mid2, p2, right))) | EQUAL => Stay (Branch3 (left, p1, mid, (k2, f (SOME x2)), right)) | GREATER => (case modfy right of Stay right' => Stay (Branch3 (left, p1, mid, p2, right')) | Sprout (right1, q, right2) => Sprout (Branch2 (left, p1, mid), p2, Branch2 (right1, q, right2))))); in (case modfy tab of Stay tab' => tab' | Sprout br => Branch2 br) handle SAME => tab end; fun update (key, x) tab = modify key (fn _ => x) tab; fun update_new (key, x) tab = modify key (fn NONE => x | SOME _ => raise DUP key) tab; fun default (key, x) tab = modify key (fn NONE => x | SOME _ => raise SAME) tab; fun map_entry key f = modify key (fn NONE => raise SAME | SOME x => f x); fun map_default (key, x) f = modify key (fn NONE => f x | SOME y => f y); (* delete *) exception UNDEF of key; local fun compare NONE (k2, _) = LESS | compare (SOME k1) (k2, _) = Key.ord (k1, k2); fun if_eq EQUAL x y = x | if_eq _ x y = y; fun del (SOME k) Empty = raise UNDEF k | del NONE (Branch2 (Empty, p, Empty)) = (p, (true, Empty)) | del NONE (Branch3 (Empty, p, Empty, q, Empty)) = (p, (false, Branch2 (Empty, q, Empty))) | del k (Branch2 (Empty, p, Empty)) = (case compare k p of EQUAL => (p, (true, Empty)) | _ => raise UNDEF (the k)) | del k (Branch3 (Empty, p, Empty, q, Empty)) = (case compare k p of EQUAL => (p, (false, Branch2 (Empty, q, Empty))) | _ => (case compare k q of EQUAL => (q, (false, Branch2 (Empty, p, Empty))) | _ => raise UNDEF (the k))) | del k (Branch2 (l, p, r)) = (case compare k p of LESS => (case del k l of (p', (false, l')) => (p', (false, Branch2 (l', p, r))) | (p', (true, l')) => (p', case r of Branch2 (rl, rp, rr) => (true, Branch3 (l', p, rl, rp, rr)) | Branch3 (rl, rp, rm, rq, rr) => (false, Branch2 (Branch2 (l', p, rl), rp, Branch2 (rm, rq, rr))))) | ord => (case del (if_eq ord NONE k) r of (p', (false, r')) => (p', (false, Branch2 (l, if_eq ord p' p, r'))) | (p', (true, r')) => (p', case l of Branch2 (ll, lp, lr) => (true, Branch3 (ll, lp, lr, if_eq ord p' p, r')) | Branch3 (ll, lp, lm, lq, lr) => (false, Branch2 (Branch2 (ll, lp, lm), lq, Branch2 (lr, if_eq ord p' p, r')))))) | del k (Branch3 (l, p, m, q, r)) = (case compare k q of LESS => (case compare k p of LESS => (case del k l of (p', (false, l')) => (p', (false, Branch3 (l', p, m, q, r))) | (p', (true, l')) => (p', (false, case (m, r) of (Branch2 (ml, mp, mr), Branch2 _) => Branch2 (Branch3 (l', p, ml, mp, mr), q, r) | (Branch3 (ml, mp, mm, mq, mr), _) => Branch3 (Branch2 (l', p, ml), mp, Branch2 (mm, mq, mr), q, r) | (Branch2 (ml, mp, mr), Branch3 (rl, rp, rm, rq, rr)) => Branch3 (Branch2 (l', p, ml), mp, Branch2 (mr, q, rl), rp, Branch2 (rm, rq, rr))))) | ord => (case del (if_eq ord NONE k) m of (p', (false, m')) => (p', (false, Branch3 (l, if_eq ord p' p, m', q, r))) | (p', (true, m')) => (p', (false, case (l, r) of (Branch2 (ll, lp, lr), Branch2 _) => Branch2 (Branch3 (ll, lp, lr, if_eq ord p' p, m'), q, r) | (Branch3 (ll, lp, lm, lq, lr), _) => Branch3 (Branch2 (ll, lp, lm), lq, Branch2 (lr, if_eq ord p' p, m'), q, r) | (_, Branch3 (rl, rp, rm, rq, rr)) => Branch3 (l, if_eq ord p' p, Branch2 (m', q, rl), rp, Branch2 (rm, rq, rr)))))) | ord => (case del (if_eq ord NONE k) r of (q', (false, r')) => (q', (false, Branch3 (l, p, m, if_eq ord q' q, r'))) | (q', (true, r')) => (q', (false, case (l, m) of (Branch2 _, Branch2 (ml, mp, mr)) => Branch2 (l, p, Branch3 (ml, mp, mr, if_eq ord q' q, r')) | (_, Branch3 (ml, mp, mm, mq, mr)) => Branch3 (l, p, Branch2 (ml, mp, mm), mq, Branch2 (mr, if_eq ord q' q, r')) | (Branch3 (ll, lp, lm, lq, lr), Branch2 (ml, mp, mr)) => Branch3 (Branch2 (ll, lp, lm), lq, Branch2 (lr, p, ml), mp, Branch2 (mr, if_eq ord q' q, r')))))); in fun delete key tab = snd (snd (del (SOME key) tab)); fun delete_safe key tab = if defined tab key then delete key tab else tab; end; (* membership operations *) fun member eq tab (key, x) = (case lookup tab key of NONE => false | SOME y => eq (x, y)); fun insert eq (key, x) = modify key (fn NONE => x | SOME y => if eq (x, y) then raise SAME else raise DUP key); fun remove eq (key, x) tab = (case lookup tab key of NONE => tab | SOME y => if eq (x, y) then delete key tab else tab); (* simultaneous modifications *) fun make entries = fold update_new entries empty; fun join f (table1, table2) = let fun add (key, y) tab = modify key (fn NONE => y | SOME x => f key (x, y)) tab; in if pointer_eq (table1, table2) then table1 else if is_empty table1 then table2 else fold_table add table2 table1 end; fun merge eq = join (fn key => fn xy => if eq xy then raise SAME else raise DUP key); (* list tables *) fun lookup_list tab key = these (lookup tab key); fun cons_list (key, x) tab = modify key (fn NONE => [x] | SOME xs => x :: xs) tab; fun insert_list eq (key, x) = modify key (fn NONE => [x] | SOME xs => if Library.member eq xs x then raise SAME else x :: xs); fun remove_list eq (key, x) tab = map_entry key (fn xs => (case Library.remove eq x xs of [] => raise UNDEF key | ys => ys)) tab handle UNDEF _ => delete key tab; fun update_list eq (key, x) = modify key (fn NONE => [x] | SOME [] => [x] | SOME (xs as y :: _) => if eq (x, y) then raise SAME else Library.update eq x xs); fun make_list args = fold_rev cons_list args empty; fun dest_list tab = maps (fn (key, xs) => map (pair key) xs) (dest tab); fun merge_list eq = join (fn _ => Library.merge eq); -(* unit tables *) +(* set operations *) type set = unit table; fun insert_set x = default (x, ()); fun remove_set x : set -> set = delete_safe x; fun make_set entries = fold insert_set entries empty; +fun subset (A: set, B: set) = forall (defined B o #1) A; +fun eq_set (A: set, B: set) = size A = size B andalso subset (A, B); + (* ML pretty-printing *) val _ = ML_system_pp (fn depth => fn pretty => fn tab => ML_Pretty.to_polyml (ML_Pretty.enum "," "{" "}" (ML_Pretty.pair (ML_Pretty.from_polyml o ML_system_pretty) (ML_Pretty.from_polyml o pretty)) (dest tab, depth))); (*final declarations of this structure!*) val map = map_table; val fold = fold_table; val fold_rev = fold_rev_table; end; structure Inttab = Table(type key = int val ord = int_ord); structure Symtab = Table(type key = string val ord = fast_string_ord); structure Symreltab = Table(type key = string * string val ord = prod_ord fast_string_ord fast_string_ord); diff --git a/src/Pure/Isar/class.ML b/src/Pure/Isar/class.ML --- a/src/Pure/Isar/class.ML +++ b/src/Pure/Isar/class.ML @@ -1,845 +1,843 @@ (* Title: Pure/Isar/class.ML Author: Florian Haftmann, TU Muenchen Type classes derived from primitive axclasses and locales. *) signature CLASS = sig (*classes*) val is_class: theory -> class -> bool val these_params: theory -> sort -> (string * (class * (string * typ))) list val base_sort: theory -> class -> sort val rules: theory -> class -> thm option * thm val these_defs: theory -> sort -> thm list val these_operations: theory -> sort -> (string * (class * ((typ * term) * bool))) list val print_classes: Proof.context -> unit val init: class -> theory -> Proof.context val begin: class list -> sort -> Proof.context -> Proof.context val const: class -> (binding * mixfix) * term -> term list * term list -> local_theory -> local_theory val abbrev: class -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val redeclare_operations: theory -> sort -> Proof.context -> Proof.context val class_prefix: string -> string val register: class -> class list -> ((string * typ) * (string * typ)) list -> sort -> morphism -> morphism -> thm option -> thm option -> thm -> theory -> theory (*instances*) val instantiation: string list * (string * sort) list * sort -> theory -> local_theory val instantiation_instance: (local_theory -> local_theory) -> local_theory -> Proof.state val prove_instantiation_instance: (Proof.context -> tactic) -> local_theory -> local_theory val prove_instantiation_exit: (Proof.context -> tactic) -> local_theory -> theory val prove_instantiation_exit_result: (morphism -> 'a -> 'b) -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory val read_multi_arity: theory -> xstring list * xstring list * xstring -> string list * (string * sort) list * sort val instantiation_cmd: xstring list * xstring list * xstring -> theory -> local_theory val instance_arity_cmd: xstring list * xstring list * xstring -> theory -> Proof.state val theory_map_result: string list * (string * sort) list * sort -> (morphism -> 'a -> 'b) -> (local_theory -> 'a * local_theory) -> (Proof.context -> 'b -> tactic) -> theory -> 'b * theory (*subclasses*) val classrel: class * class -> theory -> Proof.state val classrel_cmd: xstring * xstring -> theory -> Proof.state val register_subclass: class * class -> morphism option -> Element.witness option -> morphism -> local_theory -> local_theory (*tactics*) val intro_classes_tac: Proof.context -> thm list -> tactic val standard_intro_classes_tac: Proof.context -> thm list -> tactic (*diagnostics*) val pretty_specification: theory -> class -> Pretty.T list end; structure Class: CLASS = struct (** class data **) datatype class_data = Class_Data of { (* static part *) consts: (string * string) list (*locale parameter ~> constant name*), base_sort: sort, base_morph: morphism (*static part of canonical morphism*), export_morph: morphism, assm_intro: thm option, of_class: thm, axiom: thm option, (* dynamic part *) defs: thm Item_Net.T, operations: (string * (class * ((typ * term) * bool))) list (* n.b. params = logical parameters of class operations = operations participating in user-space type system *) }; fun make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), (defs, operations)) = Class_Data {consts = consts, base_sort = base_sort, base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, of_class = of_class, axiom = axiom, defs = defs, operations = operations}; fun map_class_data f (Class_Data {consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom, defs, operations}) = make_class_data (f ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), (defs, operations))); fun merge_class_data _ (Class_Data {consts = consts, base_sort = base_sort, base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, of_class = of_class, axiom = axiom, defs = defs1, operations = operations1}, Class_Data {consts = _, base_sort = _, base_morph = _, export_morph = _, assm_intro = _, of_class = _, axiom = _, defs = defs2, operations = operations2}) = make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), (Item_Net.merge (defs1, defs2), AList.merge (op =) (K true) (operations1, operations2))); structure Class_Data = Theory_Data ( type T = class_data Graph.T val empty = Graph.empty; val extend = I; val merge = Graph.join merge_class_data; ); (* queries *) fun lookup_class_data thy class = (case try (Graph.get_node (Class_Data.get thy)) class of SOME (Class_Data data) => SOME data | NONE => NONE); fun the_class_data thy class = (case lookup_class_data thy class of NONE => error ("Undeclared class " ^ quote class) | SOME data => data); val is_class = is_some oo lookup_class_data; val ancestry = Graph.all_succs o Class_Data.get; val heritage = Graph.all_preds o Class_Data.get; fun these_params thy = let fun params class = let val const_typs = (#params o Axclass.get_info thy) class; val const_names = (#consts o the_class_data thy) class; in (map o apsnd) (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names end; in maps params o ancestry thy end; val base_sort = #base_sort oo the_class_data; fun rules thy class = let val {axiom, of_class, ...} = the_class_data thy class in (axiom, of_class) end; fun all_assm_intros thy = Graph.fold (fn (_, (Class_Data {assm_intro, ...}, _)) => fold (insert Thm.eq_thm) (the_list assm_intro)) (Class_Data.get thy) []; fun these_defs thy = maps (Item_Net.content o #defs o the_class_data thy) o ancestry thy; fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy; val base_morphism = #base_morph oo the_class_data; fun morphism thy class = (case Element.eq_morphism thy (these_defs thy [class]) of SOME eq_morph => base_morphism thy class $> eq_morph | NONE => base_morphism thy class); val export_morphism = #export_morph oo the_class_data; fun pretty_param ctxt (c, ty) = Pretty.block [Name_Space.pretty ctxt (Proof_Context.const_space ctxt) c, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt ty]; fun print_classes ctxt = let val thy = Proof_Context.theory_of ctxt; val algebra = Sign.classes_of thy; val class_space = Proof_Context.class_space ctxt; val type_space = Proof_Context.type_space ctxt; val arities = - Symtab.empty - |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => - Symtab.map_default (class, []) (insert (op =) tyco)) arities) - (Sorts.arities_of algebra); + Symtab.build (Sorts.arities_of algebra |> Symtab.fold (fn (tyco, arities) => + fold (fn (class, _) => Symtab.map_default (class, []) (insert (op =) tyco)) arities)); fun prt_supersort class = Syntax.pretty_sort ctxt (Sign.minimize_sort thy (Sign.super_classes thy class)); fun prt_arity class tyco = let val Ss = Sorts.mg_domain algebra tyco [class]; in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end; fun prt_param (c, ty) = pretty_param ctxt (c, Type.strip_sorts_dummy ty); fun prt_entry class = Pretty.block ([Pretty.keyword1 "class", Pretty.brk 1, Name_Space.pretty ctxt class_space class, Pretty.str ":", Pretty.fbrk, Pretty.block [Pretty.str "supersort: ", prt_supersort class]] @ (case (these o Option.map #params o try (Axclass.get_info thy)) class of [] => [] | params => [Pretty.fbrk, Pretty.big_list "parameters:" (map prt_param params)]) @ (case (these o Symtab.lookup arities) class of [] => [] | ars => [Pretty.fbrk, Pretty.big_list "instances:" (map (prt_arity class) (sort (Name_Space.extern_ord ctxt type_space) ars))])); in Sorts.all_classes algebra |> sort (Name_Space.extern_ord ctxt class_space) |> map prt_entry |> Pretty.writeln_chunks2 end; (* updaters *) fun register class sups params base_sort base_morph export_morph some_axiom some_assm_intro of_class thy = let val operations = map (fn (v_ty as (_, ty), (c, _)) => (c, (class, ((ty, Free v_ty), false)))) params; val add_class = Graph.new_node (class, make_class_data (((map o apply2) fst params, base_sort, base_morph, export_morph, Option.map Thm.trim_context some_assm_intro, Thm.trim_context of_class, Option.map Thm.trim_context some_axiom), (Thm.item_net, operations))) #> fold (curry Graph.add_edge class) sups; in Class_Data.map add_class thy end; fun activate_defs class thms thy = (case Element.eq_morphism thy thms of SOME eq_morph => fold (fn cls => fn thy => Context.theory_map (Locale.amend_registration {inst = (cls, base_morphism thy cls), mixin = SOME (eq_morph, true), export = export_morphism thy cls}) thy) (heritage thy [class]) thy | NONE => thy); fun register_operation class (c, t) input_only thy = let val base_sort = base_sort thy class; val prep_typ = map_type_tfree (fn (v, sort) => if Name.aT = v then TFree (v, base_sort) else TVar ((v, 0), sort)); val t' = map_types prep_typ t; val ty' = Term.fastype_of t'; in thy |> (Class_Data.map o Graph.map_node class o map_class_data o apsnd o apsnd) (cons (c, (class, ((ty', t'), input_only)))) end; fun register_def class def_thm thy = let val sym_thm = Thm.trim_context (Thm.symmetric def_thm) in thy |> (Class_Data.map o Graph.map_node class o map_class_data o apsnd o apfst) (Item_Net.update sym_thm) |> activate_defs class [sym_thm] end; (** classes and class target **) (* class context syntax *) fun make_rewrite t c_ty = let val vs = strip_abs_vars t; val vts = map snd vs |> Name.invent_names Name.context Name.uu |> map (fn (v, T) => Var ((v, 0), T)); in (betapplys (t, vts), betapplys (Const c_ty, vts)) end; fun these_unchecks thy = these_operations thy #> map_filter (fn (c, (_, ((ty, t), input_only))) => if input_only then NONE else SOME (make_rewrite t (c, ty))); fun these_unchecks_reversed thy = these_operations thy #> map (fn (c, (_, ((ty, t), _))) => (Const (c, ty), t)); fun redeclare_const thy c = let val b = Long_Name.base_name c in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end; fun synchronize_class_syntax sort base_sort ctxt = let val thy = Proof_Context.theory_of ctxt; val algebra = Sign.classes_of thy; val operations = these_operations thy sort; fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort))); val primary_constraints = (map o apsnd) (subst_class_typ base_sort o fst o fst o snd) operations; val secondary_constraints = (map o apsnd) (fn (class, ((ty, _), _)) => subst_class_typ [class] ty) operations; fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0) of SOME (_, ty' as TVar (vi, sort)) => if Type_Infer.is_param vi andalso Sorts.sort_le algebra (base_sort, sort) then SOME (ty', Term.aT base_sort) else NONE | _ => NONE) | NONE => NONE) | NONE => NONE); fun subst (c, _) = Option.map (fst o snd) (AList.lookup (op =) operations c); val unchecks = these_unchecks thy sort; in ctxt |> fold (redeclare_const thy o fst) primary_constraints |> Overloading.map_improvable_syntax (K {primary_constraints = primary_constraints, secondary_constraints = secondary_constraints, improve = improve, subst = subst, no_subst_in_abbrev_mode = true, unchecks = unchecks}) |> Overloading.set_primary_constraints end; fun synchronize_class_syntax_target class lthy = lthy |> Local_Theory.map_contexts (K (synchronize_class_syntax [class] (base_sort (Proof_Context.theory_of lthy) class))); fun redeclare_operations thy sort = fold (redeclare_const thy o fst) (these_operations thy sort); fun begin sort base_sort ctxt = ctxt |> Variable.declare_term (Logic.mk_type (Term.aT base_sort)) |> synchronize_class_syntax sort base_sort |> Overloading.activate_improvable_syntax; fun init class thy = thy |> Locale.init class |> begin [class] (base_sort thy class); (* class target *) val class_prefix = Logic.const_of_class o Long_Name.base_name; fun guess_morphism_identity (b, rhs) phi1 phi2 = let (*FIXME proper concept to identify morphism instead of educated guess*) val name_of_binding = Name_Space.full_name Name_Space.global_naming; val n1 = (name_of_binding o Morphism.binding phi1) b; val n2 = (name_of_binding o Morphism.binding phi2) b; val rhs1 = Morphism.term phi1 rhs; val rhs2 = Morphism.term phi2 rhs; in n1 = n2 andalso Term.aconv_untyped (rhs1, rhs2) end; fun target_const class phi0 prmode (b, rhs) lthy = let val export = Variable.export_morphism lthy (Local_Theory.target_of lthy); val guess_identity = guess_morphism_identity (b, rhs) export; val guess_canonical = guess_morphism_identity (b, rhs) (export $> phi0); in lthy |> Generic_Target.locale_target_const class (not o (guess_identity orf guess_canonical)) prmode ((b, NoSyn), rhs) end; local fun dangling_params_for lthy class (type_params, term_params) = let val class_param_names = map fst (these_params (Proof_Context.theory_of lthy) [class]); val dangling_term_params = subtract (fn (v, Free (w, _)) => v = w | _ => false) class_param_names term_params; in (type_params, dangling_term_params) end; fun global_def (b, eq) thy = let val ((_, def_thm), thy') = thy |> Thm.add_def_global false false (b, eq); val def_thm' = def_thm |> Thm.forall_intr_frees |> Thm.forall_elim_vars 0 |> Thm.varifyT_global; val (_, thy'') = thy' |> Global_Theory.store_thm (b, def_thm'); in (def_thm', thy'') end; fun canonical_const class phi dangling_params ((b, mx), rhs) thy = let val b_def = Binding.suffix_name "_dict" b; val c = Sign.full_name thy b; val ty = map Term.fastype_of dangling_params ---> Term.fastype_of rhs; val def_eq = Logic.mk_equals (list_comb (Const (c, ty), dangling_params), rhs) |> map_types Type.strip_sorts; in thy |> Sign.declare_const_global ((b, Type.strip_sorts ty), mx) |> snd |> global_def (b_def, def_eq) |-> (fn def_thm => register_def class def_thm) |> null dangling_params ? register_operation class (c, rhs) false |> Sign.add_const_constraint (c, SOME ty) end; in fun const class ((b, mx), lhs) params lthy = let val phi = morphism (Proof_Context.theory_of lthy) class; val dangling_params = map (Morphism.term phi) (uncurry append (dangling_params_for lthy class params)); in lthy |> target_const class phi Syntax.mode_default (b, lhs) |> Local_Theory.raw_theory (canonical_const class phi dangling_params ((Morphism.binding phi b, if null dangling_params then mx else NoSyn), Morphism.term phi lhs)) |> Generic_Target.standard_const (fn (this, other) => other <> 0 andalso this <> other) Syntax.mode_default ((b, if null dangling_params then NoSyn else mx), lhs) |> synchronize_class_syntax_target class end; end; local fun canonical_abbrev class phi prmode with_syntax ((b, mx), rhs) thy = let val c = Sign.full_name thy b; val constrain = map_atyps (fn T as TFree (v, _) => if v = Name.aT then TFree (v, [class]) else T | T => T); val rhs' = map_types constrain rhs; in thy |> Sign.add_abbrev (#1 prmode) (b, Logic.varify_types_global rhs') |> snd |> with_syntax ? Sign.notation true prmode [(Const (c, fastype_of rhs), mx)] |> with_syntax ? register_operation class (c, rhs) (#1 prmode = Print_Mode.input) |> Sign.add_const_constraint (c, SOME (fastype_of rhs')) end; fun canonical_abbrev_target class phi prmode ((b, mx), rhs) lthy = let val thy = Proof_Context.theory_of lthy; val preprocess = perhaps (try (Pattern.rewrite_term_top thy (these_unchecks thy [class]) [])); val (global_rhs, (_, (_, term_params))) = Generic_Target.export_abbrev lthy preprocess rhs; val mx' = Generic_Target.check_mixfix_global (b, null term_params) mx; in lthy |> Local_Theory.raw_theory (canonical_abbrev class phi prmode (null term_params) ((Morphism.binding phi b, mx'), Logic.unvarify_types_global global_rhs)) end; fun further_abbrev_target class phi prmode (b, mx) rhs params = Generic_Target.background_abbrev (b, rhs) (snd params) #-> (fn (lhs, _) => target_const class phi prmode (b, lhs) #> Generic_Target.standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), lhs)) in fun abbrev class prmode ((b, mx), rhs) lthy = let val thy = Proof_Context.theory_of lthy; val phi = morphism thy class; val rhs_generic = perhaps (try (Pattern.rewrite_term_top thy (these_unchecks_reversed thy [class]) [])) rhs; in lthy |> canonical_abbrev_target class phi prmode ((b, mx), rhs) |> Generic_Target.abbrev (further_abbrev_target class phi) prmode ((b, mx), rhs_generic) ||> synchronize_class_syntax_target class end; end; (* subclasses *) fun register_subclass (sub, sup) some_dep_morph some_witn export lthy = let val thy = Proof_Context.theory_of lthy; val intros = (snd o rules thy) sup :: map_filter I [Option.map (Drule.export_without_context_open o Element.conclude_witness lthy) some_witn, (fst o rules thy) sub]; val classrel = Goal.prove_sorry_global thy [] [] (Logic.mk_classrel (sub, sup)) (fn {context = ctxt, ...} => EVERY (map (TRYALL o resolve_tac ctxt o single) intros)); val diff_sort = Sign.complete_sort thy [sup] |> subtract (op =) (Sign.complete_sort thy [sub]) |> filter (is_class thy); val add_dependency = (case some_dep_morph of SOME dep_morph => Generic_Target.locale_dependency sub {inst = (sup, dep_morph $> Element.satisfy_morphism (the_list some_witn)), mixin = NONE, export = export} | NONE => I); in lthy |> Local_Theory.raw_theory (Axclass.add_classrel classrel #> Class_Data.map (Graph.add_edge (sub, sup)) #> activate_defs sub (these_defs thy diff_sort)) |> add_dependency |> synchronize_class_syntax_target sub end; local fun gen_classrel mk_prop classrel thy = let fun after_qed results = Proof_Context.background_theory ((fold o fold) Axclass.add_classrel results); in thy |> Proof_Context.init_global |> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]] end; in val classrel = gen_classrel (Logic.mk_classrel oo Axclass.cert_classrel); val classrel_cmd = gen_classrel (Logic.mk_classrel oo Axclass.read_classrel); end; (*local*) (** instantiation target **) (* bookkeeping *) datatype instantiation = Instantiation of { arities: string list * (string * sort) list * sort, params: ((string * string) * (string * typ)) list (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*) } fun make_instantiation (arities, params) = Instantiation {arities = arities, params = params}; val empty_instantiation = make_instantiation (([], [], []), []); structure Instantiation = Proof_Data ( type T = instantiation; fun init _ = empty_instantiation; ); val get_instantiation = (fn Instantiation data => data) o Instantiation.get o Local_Theory.target_of; fun map_instantiation f = (Local_Theory.target o Instantiation.map) (fn Instantiation {arities, params} => make_instantiation (f (arities, params))); fun the_instantiation lthy = (case get_instantiation lthy of {arities = ([], [], []), ...} => error "No instantiation target" | data => data); val instantiation_params = #params o get_instantiation; fun instantiation_param lthy b = instantiation_params lthy |> find_first (fn (_, (v, _)) => Binding.name_of b = v) |> Option.map (fst o fst); fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) = let val ctxt = Proof_Context.init_global thy; val all_arities = map (fn raw_tyco => Proof_Context.read_arity ctxt (raw_tyco, raw_sorts, raw_sort)) raw_tycos; val tycos = map #1 all_arities; val (_, sorts, sort) = hd all_arities; val vs = Name.invent_names Name.context Name.aT sorts; in (tycos, vs, sort) end; (* syntax *) fun synchronize_inst_syntax ctxt = let val Instantiation {params, ...} = Instantiation.get ctxt; val lookup_inst_param = Axclass.lookup_inst_param (Sign.consts_of (Proof_Context.theory_of ctxt)) params; fun subst (c, ty) = (case lookup_inst_param (c, ty) of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) | NONE => NONE); val unchecks = map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params; in ctxt |> Overloading.map_improvable_syntax (fn {primary_constraints, improve, ...} => {primary_constraints = primary_constraints, secondary_constraints = [], improve = improve, subst = subst, no_subst_in_abbrev_mode = false, unchecks = unchecks}) end; fun resort_terms ctxt algebra consts constraints ts = let fun matchings (Const (c_ty as (c, _))) = (case constraints c of NONE => I | SOME sorts => fold2 (curry (Sorts.meet_sort algebra)) (Consts.typargs consts c_ty) sorts) | matchings _ = I; - val tvartab = (fold o fold_aterms) matchings ts Vartab.empty + val tvartab = Vartab.build ((fold o fold_aterms) matchings ts) handle Sorts.CLASS_ERROR e => error (Sorts.class_error (Context.Proof ctxt) e); val inst = map_type_tvar (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi))); in if Vartab.is_empty tvartab then ts else (map o map_types) inst ts end; (* target *) fun define_overloaded (c, U) b (b_def, rhs) lthy = let val name = Binding.name_of b; val pos = Binding.pos_of b; val _ = if Context_Position.is_reported lthy pos then Position.report_text pos Markup.class_parameter (Pretty.string_of (Pretty.block [Pretty.keyword1 "class", Pretty.brk 1, Pretty.str "parameter", Pretty.brk 1, pretty_param lthy (c, U)])) else (); in lthy |> Local_Theory.background_theory_result (Axclass.declare_overloaded (c, U) ##>> Axclass.define_overloaded b_def (c, rhs)) ||> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = name)) ||> Local_Theory.map_contexts (K synchronize_inst_syntax) end; fun foundation (((b, U), mx), (b_def, rhs)) params lthy = (case instantiation_param lthy b of SOME c => if Mixfix.is_empty mx then lthy |> define_overloaded (c, U) b (b_def, rhs) else error ("Illegal mixfix syntax for overloaded constant " ^ quote c) | NONE => lthy |> Generic_Target.theory_target_foundation (((b, U), mx), (b_def, rhs)) params); fun pretty lthy = let val {arities = (tycos, vs, sort), params} = the_instantiation lthy; fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); fun pr_param ((c, _), (v, ty)) = Pretty.block (Pretty.breaks [Pretty.str v, Pretty.str "==", Proof_Context.pretty_const lthy c, Pretty.str "::", Syntax.pretty_typ lthy ty]); in [Pretty.block (Pretty.fbreaks (Pretty.keyword1 "instantiation" :: map pr_arity tycos @ map pr_param params))] end; fun conclude lthy = let val (tycos, vs, sort) = #arities (the_instantiation lthy); val thy = Proof_Context.theory_of lthy; val _ = tycos |> List.app (fn tyco => if Sign.of_sort thy (Type (tyco, map TFree vs), sort) then () else error ("Missing instance proof for type " ^ quote (Proof_Context.markup_type lthy tyco))); in lthy end; fun registration thy_ctxt {inst, mixin, export} lthy = lthy |> Generic_Target.theory_registration {inst = inst, mixin = mixin, export = export $> Proof_Context.export_morphism lthy thy_ctxt} (*handle fixed types variables on target context properly*); fun instantiation (tycos, vs, sort) thy = let val _ = if null tycos then error "At least one arity must be given" else (); val class_params = these_params thy (filter (can (Axclass.get_info thy)) sort); fun get_param tyco (param, (_, (c, ty))) = if can (Axclass.param_of_inst thy) (c, tyco) then NONE else SOME ((c, tyco), (param ^ "_" ^ Long_Name.base_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty)); val params = map_product get_param tycos class_params |> map_filter I; val _ = if null params andalso forall (fn tyco => can (Sign.arity_sorts thy tyco) sort) tycos then error "No parameters and no pending instance proof obligations in instantiation." else (); val primary_constraints = map (apsnd (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) class_params; val algebra = Sign.classes_of thy |> fold (fn tyco => Sorts.add_arities (Context.Theory thy) (tyco, map (fn class => (class, map snd vs)) sort)) tycos; val consts = Sign.consts_of thy; val improve_constraints = AList.lookup (op =) (map (fn (_, (class, (c, _))) => (c, [[class]])) class_params); fun resort_check ctxt ts = resort_terms ctxt algebra consts improve_constraints ts; val lookup_inst_param = Axclass.lookup_inst_param consts params; fun improve (c, ty) = (case lookup_inst_param (c, ty) of SOME (_, ty') => if Sign.typ_instance thy (ty', ty) then SOME (ty, ty') else NONE | NONE => NONE); in thy |> Local_Theory.init {background_naming = Sign.naming_of thy, setup = Proof_Context.init_global #> Instantiation.put (make_instantiation ((tycos, vs, sort), params)) #> fold (Variable.declare_typ o TFree) vs #> fold (Variable.declare_names o Free o snd) params #> (Overloading.map_improvable_syntax) (K {primary_constraints = primary_constraints, secondary_constraints = [], improve = improve, subst = K NONE, no_subst_in_abbrev_mode = false, unchecks = []}) #> Overloading.activate_improvable_syntax #> Context.proof_map (Syntax_Phases.term_check 0 "resorting" resort_check) #> synchronize_inst_syntax, conclude = conclude} {define = Generic_Target.define foundation, notes = Generic_Target.notes Generic_Target.theory_target_notes, abbrev = Generic_Target.abbrev Generic_Target.theory_target_abbrev, declaration = K Generic_Target.theory_declaration, theory_registration = registration (Proof_Context.init_global thy), locale_dependency = fn _ => error "Not possible in instantiation target", pretty = pretty} end; fun instantiation_cmd arities thy = instantiation (read_multi_arity thy arities) thy; fun gen_instantiation_instance do_proof after_qed lthy = let val (tycos, vs, sort) = (#arities o the_instantiation) lthy; val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos; fun after_qed' results = Local_Theory.background_theory (fold (Axclass.add_arity o Thm.varifyT_global) results) #> after_qed; in lthy |> do_proof after_qed' arities_proof end; val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts => Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts)); fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed => fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t (fn {context, ...} => tac context)) ts) lthy) I; fun prove_instantiation_exit tac = prove_instantiation_instance tac #> Local_Theory.exit_global; fun prove_instantiation_exit_result f tac x lthy = let val morph = Proof_Context.export_morphism lthy (Proof_Context.init_global (Proof_Context.theory_of lthy)); val y = f morph x; in lthy |> prove_instantiation_exit (fn ctxt => tac ctxt y) |> pair y end; fun theory_map_result arities f g tac = instantiation arities #> g #-> prove_instantiation_exit_result f tac; (* simplified instantiation interface with no class parameter *) fun instance_arity_cmd raw_arities thy = let val (tycos, vs, sort) = read_multi_arity thy raw_arities; val sorts = map snd vs; val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos; fun after_qed results = Proof_Context.background_theory ((fold o fold) Axclass.add_arity results); in thy |> Proof_Context.init_global |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities) end; (** tactics and methods **) fun intro_classes_tac ctxt facts st = let val thy = Proof_Context.theory_of ctxt; val classes = Sign.all_classes thy; val class_trivs = map (Thm.class_triv thy) classes; val class_intros = map_filter (try (#intro o Axclass.get_info thy)) classes; val assm_intros = all_assm_intros thy; in Method.intros_tac ctxt (class_trivs @ class_intros @ assm_intros) facts st end; fun standard_intro_classes_tac ctxt facts st = if null facts andalso not (Thm.no_prems st) then (intro_classes_tac ctxt [] ORELSE Locale.intro_locales_tac {strict = true, eager = true} ctxt []) st else no_tac st; fun standard_tac ctxt facts = HEADGOAL (Method.some_rule_tac ctxt [] facts) ORELSE standard_intro_classes_tac ctxt facts; val _ = Theory.setup (Method.setup \<^binding>\intro_classes\ (Scan.succeed (METHOD o intro_classes_tac)) "back-chain introduction rules of classes" #> Method.setup \<^binding>\standard\ (Scan.succeed (METHOD o standard_tac)) "standard proof step: Pure intro/elim rule or class introduction"); (** diagnostics **) fun pretty_specification thy class = if is_class thy class then let val class_ctxt = init class thy; val prt_class = Name_Space.pretty class_ctxt (Proof_Context.class_space class_ctxt); val super_classes = Sign.minimize_sort thy (Sign.super_classes thy class); val fix_args = #params (Axclass.get_info thy class) |> map (fn (c, T) => (Binding.name (Long_Name.base_name c), SOME T, NoSyn)); val fixes = if null fix_args then [] else [Element.Fixes fix_args]; val assumes = Locale.hyp_spec_of thy class; val header = [Pretty.keyword1 "class", Pretty.brk 1, prt_class class, Pretty.str " =", Pretty.brk 1] @ Pretty.separate " +" (map prt_class super_classes) @ (if null super_classes then [] else [Pretty.str " +"]); val body = if null fixes andalso null assumes then [] else maps (Element.pretty_ctxt_no_attribs class_ctxt) (fixes @ assumes) |> maps (fn prt => [Pretty.fbrk, prt]); in if null body then [] else [Pretty.block (header @ body)] end else []; end; diff --git a/src/Pure/Isar/code.ML b/src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML +++ b/src/Pure/Isar/code.ML @@ -1,1584 +1,1584 @@ (* Title: Pure/Isar/code.ML Author: Florian Haftmann, TU Muenchen Abstract executable ingredients of theory. Management of data dependent on executable ingredients as synchronized cache; purged on any change of underlying executable ingredients. *) signature CODE = sig (*constants*) val check_const: theory -> term -> string val read_const: theory -> string -> string val string_of_const: theory -> string -> string val args_number: theory -> string -> int (*constructor sets*) val constrset_of_consts: theory -> (string * typ) list -> string * ((string * sort) list * (string * ((string * sort) list * typ list)) list) (*code equations and certificates*) val assert_eqn: theory -> thm * bool -> thm * bool val assert_abs_eqn: theory -> string option -> thm -> thm * (string * string) type cert val constrain_cert: theory -> sort list -> cert -> cert val conclude_cert: cert -> cert val typargs_deps_of_cert: theory -> cert -> (string * sort) list * (string * typ list) list val equations_of_cert: theory -> cert -> ((string * sort) list * typ) * (((term * string option) list * (term * string option)) * (thm option * bool)) list option val pretty_cert: theory -> cert -> Pretty.T list (*executable code*) type constructors type abs_type val type_interpretation: (string -> theory -> theory) -> theory -> theory val datatype_interpretation: (string * constructors -> theory -> theory) -> theory -> theory val abstype_interpretation: (string * abs_type -> theory -> theory) -> theory -> theory val declare_datatype_global: (string * typ) list -> theory -> theory val declare_datatype_cmd: string list -> theory -> theory val declare_abstype: thm -> local_theory -> local_theory val declare_abstype_global: thm -> theory -> theory val declare_default_eqns: (thm * bool) list -> local_theory -> local_theory val declare_default_eqns_global: (thm * bool) list -> theory -> theory val declare_eqns: (thm * bool) list -> local_theory -> local_theory val declare_eqns_global: (thm * bool) list -> theory -> theory val add_eqn_global: thm * bool -> theory -> theory val del_eqn_global: thm -> theory -> theory val declare_abstract_eqn: thm -> local_theory -> local_theory val declare_abstract_eqn_global: thm -> theory -> theory val declare_aborting_global: string -> theory -> theory val declare_unimplemented_global: string -> theory -> theory val declare_case_global: thm -> theory -> theory val declare_undefined_global: string -> theory -> theory val get_type: theory -> string -> constructors * bool val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option val is_constr: theory -> string -> bool val is_abstr: theory -> string -> bool val get_cert: Proof.context -> ((thm * bool) list -> (thm * bool) list option) list -> string -> cert type case_schema val get_case_schema: theory -> string -> case_schema option val get_case_cong: theory -> string -> thm option val is_undefined: theory -> string -> bool val print_codesetup: theory -> unit end; signature CODE_DATA_ARGS = sig type T val empty: T end; signature CODE_DATA = sig type T val change: theory option -> (T -> T) -> T val change_yield: theory option -> (T -> 'a * T) -> 'a * T end; signature PRIVATE_CODE = sig include CODE val declare_data: Any.T -> serial val change_yield_data: serial * ('a -> Any.T) * (Any.T -> 'a) -> theory -> ('a -> 'b * 'a) -> 'b * 'a end; structure Code : PRIVATE_CODE = struct (** auxiliary **) (* printing *) fun string_of_typ thy = Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy)); fun string_of_const thy c = let val ctxt = Proof_Context.init_global thy in case Axclass.inst_of_param thy c of SOME (c, tyco) => Proof_Context.extern_const ctxt c ^ " " ^ enclose "[" "]" (Proof_Context.extern_type ctxt tyco) | NONE => Proof_Context.extern_const ctxt c end; (* constants *) fun const_typ thy = Type.strip_sorts o Sign.the_const_type thy; fun args_number thy = length o binder_types o const_typ thy; fun devarify ty = let val tys = fold_atyps (fn TVar vi_sort => AList.update (op =) vi_sort) ty []; val vs = Name.invent Name.context Name.aT (length tys); val mapping = map2 (fn v => fn (vi, sort) => (vi, TFree (v, sort))) vs tys; in Term.typ_subst_TVars mapping ty end; fun typscheme thy (c, ty) = (map dest_TFree (Sign.const_typargs thy (c, ty)), Type.strip_sorts ty); fun typscheme_equiv (ty1, ty2) = Type.raw_instance (devarify ty1, ty2) andalso Type.raw_instance (devarify ty2, ty1); fun check_bare_const thy t = case try dest_Const t of SOME c_ty => c_ty | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t); fun check_unoverload thy (c, ty) = let val c' = Axclass.unoverload_const thy (c, ty); val ty_decl = const_typ thy c'; in if typscheme_equiv (ty_decl, Logic.varifyT_global ty) then c' else error ("Type\n" ^ string_of_typ thy ty ^ "\nof constant " ^ quote c ^ "\nis too specific compared to declared type\n" ^ string_of_typ thy ty_decl) end; fun check_const thy = check_unoverload thy o check_bare_const thy; fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; fun read_const thy = check_unoverload thy o read_bare_const thy; (** executable specifications **) (* types *) datatype type_spec = Constructors of { constructors: (string * ((string * sort) list * typ list)) list, case_combinators: string list} | Abstractor of { abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string, more_abstract_functions: string list}; fun concrete_constructors_of (Constructors {constructors, ...}) = constructors | concrete_constructors_of _ = []; fun constructors_of (Constructors {constructors, ...}) = (constructors, false) | constructors_of (Abstractor {abstractor = (co, (vs, ty)), ...}) = ([(co, (vs, [ty]))], true); fun case_combinators_of (Constructors {case_combinators, ...}) = case_combinators | case_combinators_of (Abstractor _) = []; fun add_case_combinator c (vs, Constructors {constructors, case_combinators}) = (vs, Constructors {constructors = constructors, case_combinators = insert (op =) c case_combinators}); fun projection_of (Constructors _) = NONE | projection_of (Abstractor {projection, ...}) = SOME projection; fun abstract_functions_of (Constructors _) = [] | abstract_functions_of (Abstractor {more_abstract_functions, projection, ...}) = projection :: more_abstract_functions; fun add_abstract_function c (vs, Abstractor {abs_rep, abstractor, projection, more_abstract_functions}) = (vs, Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, more_abstract_functions = insert (op =) c more_abstract_functions}); fun join_same_types' (Constructors {constructors, case_combinators = case_combinators1}, Constructors {case_combinators = case_combinators2, ...}) = Constructors {constructors = constructors, case_combinators = merge (op =) (case_combinators1, case_combinators2)} | join_same_types' (Abstractor {abs_rep, abstractor, projection, more_abstract_functions = more_abstract_functions1}, Abstractor {more_abstract_functions = more_abstract_functions2, ...}) = Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, more_abstract_functions = merge (op =) (more_abstract_functions1, more_abstract_functions2)}; fun join_same_types ((vs, spec1), (_, spec2)) = (vs, join_same_types' (spec1, spec2)); (* functions *) datatype fun_spec = Eqns of bool * (thm * bool) list | Proj of term * (string * string) | Abstr of thm * (string * string); val unimplemented = Eqns (true, []); fun is_unimplemented (Eqns (true, [])) = true | is_unimplemented _ = false; fun is_default (Eqns (true, _)) = true | is_default _ = false; val aborting = Eqns (false, []); fun associated_abstype (Proj (_, tyco_abs)) = SOME tyco_abs | associated_abstype (Abstr (_, tyco_abs)) = SOME tyco_abs | associated_abstype _ = NONE; (* cases *) type case_schema = int * (int * string option list); datatype case_spec = No_Case | Case of {schema: case_schema, tycos: string list, cong: thm} | Undefined; fun associated_datatypes (Case {tycos, schema = (_, (_, raw_cos)), ...}) = (tycos, map_filter I raw_cos) | associated_datatypes _ = ([], []); (** background theory data store **) (* historized declaration data *) structure History = struct type 'a T = { entry: 'a, suppressed: bool, (*incompatible entries are merely suppressed after theory merge but sustain*) history: serial list (*explicit trace of declaration history supports non-monotonic declarations*) } Symtab.table; fun some_entry (SOME {suppressed = false, entry, ...}) = SOME entry | some_entry _ = NONE; fun lookup table = Symtab.lookup table #> some_entry; fun register key entry table = if is_some (Symtab.lookup table key) then Symtab.map_entry key (fn {history, ...} => {entry = entry, suppressed = false, history = serial () :: history}) table else Symtab.update (key, {entry = entry, suppressed = false, history = [serial ()]}) table; fun modify_entry key f = Symtab.map_entry key (fn {entry, suppressed, history} => {entry = f entry, suppressed = suppressed, history = history}); fun all table = Symtab.dest table |> map_filter (fn (key, {entry, suppressed = false, ...}) => SOME (key, entry) | _ => NONE); local fun merge_history join_same ({entry = entry1, history = history1, ...}, {entry = entry2, history = history2, ...}) = let val history = merge (op =) (history1, history2); val entry = if hd history1 = hd history2 then join_same (entry1, entry2) else if hd history = hd history1 then entry1 else entry2; in {entry = entry, suppressed = false, history = history} end; in fun join join_same tables = Symtab.join (K (merge_history join_same)) tables; fun suppress key = Symtab.map_entry key (fn {entry, history, ...} => {entry = entry, suppressed = true, history = history}); fun suppress_except f = Symtab.map (fn key => fn {entry, suppressed, history} => {entry = entry, suppressed = suppressed orelse (not o f) (key, entry), history = history}); end; end; datatype specs = Specs of { types: ((string * sort) list * type_spec) History.T, pending_eqns: (thm * bool) list Symtab.table, functions: fun_spec History.T, cases: case_spec History.T }; fun types_of (Specs {types, ...}) = types; fun pending_eqns_of (Specs {pending_eqns, ...}) = pending_eqns; fun functions_of (Specs {functions, ...}) = functions; fun cases_of (Specs {cases, ...}) = cases; fun make_specs (types, ((pending_eqns, functions), cases)) = Specs {types = types, pending_eqns = pending_eqns, functions = functions, cases = cases}; val empty_specs = make_specs (Symtab.empty, ((Symtab.empty, Symtab.empty), Symtab.empty)); fun map_specs f (Specs {types = types, pending_eqns = pending_eqns, functions = functions, cases = cases}) = make_specs (f (types, ((pending_eqns, functions), cases))); fun merge_specs (Specs {types = types1, pending_eqns = _, functions = functions1, cases = cases1}, Specs {types = types2, pending_eqns = _, functions = functions2, cases = cases2}) = let val types = History.join join_same_types (types1, types2); val all_types = map (snd o snd) (History.all types); fun check_abstype (c, fun_spec) = case associated_abstype fun_spec of NONE => true | SOME (tyco, abs) => (case History.lookup types tyco of NONE => false | SOME (_, Constructors _) => false | SOME (_, Abstractor {abstractor = (abs', _), projection, more_abstract_functions, ...}) => abs = abs' andalso (c = projection orelse member (op =) more_abstract_functions c)); fun check_datatypes (_, case_spec) = let val (tycos, required_constructors) = associated_datatypes case_spec; val allowed_constructors = tycos |> maps (these o Option.map (concrete_constructors_of o snd) o History.lookup types) |> map fst; in subset (op =) (required_constructors, allowed_constructors) end; val all_constructors = maps (fst o constructors_of) all_types; val functions = History.join fst (functions1, functions2) |> fold (History.suppress o fst) all_constructors |> History.suppress_except check_abstype; val cases = History.join fst (cases1, cases2) |> History.suppress_except check_datatypes; in make_specs (types, ((Symtab.empty, functions), cases)) end; val map_types = map_specs o apfst; val map_pending_eqns = map_specs o apsnd o apfst o apfst; val map_functions = map_specs o apsnd o apfst o apsnd; val map_cases = map_specs o apsnd o apsnd; (* data slots dependent on executable code *) (*private copy avoids potential conflict of table exceptions*) structure Datatab = Table(type key = int val ord = int_ord); local type kind = {empty: Any.T}; val kinds = Synchronized.var "Code_Data" (Datatab.empty: kind Datatab.table); fun invoke f k = (case Datatab.lookup (Synchronized.value kinds) k of SOME kind => f kind | NONE => raise Fail "Invalid code data identifier"); in fun declare_data empty = let val k = serial (); val kind = {empty = empty}; val _ = Synchronized.change kinds (Datatab.update (k, kind)); in k end; fun invoke_init k = invoke (fn kind => #empty kind) k; end; (*local*) (* global theory store *) local type data = Any.T Datatab.table; fun make_dataref thy = (Context.theory_long_name thy, Synchronized.var "code data" (NONE : (data * Context.theory_id) option)); structure Code_Data = Theory_Data ( type T = specs * (string * (data * Context.theory_id) option Synchronized.var); val empty = (empty_specs, make_dataref (Context.the_global_context ())); val extend = I; fun merge ((specs1, dataref), (specs2, _)) = (merge_specs (specs1, specs2), dataref); ); fun init_dataref thy = if #1 (#2 (Code_Data.get thy)) = Context.theory_long_name thy then NONE else SOME ((Code_Data.map o apsnd) (fn _ => make_dataref thy) thy) in val _ = Theory.setup (Theory.at_begin init_dataref); (* access to executable specifications *) val specs_of : theory -> specs = fst o Code_Data.get; fun modify_specs f thy = Code_Data.map (fn (specs, _) => (f specs, make_dataref thy)) thy; (* access to data dependent on executable specifications *) fun change_yield_data (kind, mk, dest) theory f = let val dataref = #2 (#2 (Code_Data.get theory)); val (datatab, thy_id) = case Synchronized.value dataref of SOME (datatab, thy_id) => if Context.eq_thy_id (Context.theory_id theory, thy_id) then (datatab, thy_id) else (Datatab.empty, Context.theory_id theory) | NONE => (Datatab.empty, Context.theory_id theory) val data = case Datatab.lookup datatab kind of SOME data => data | NONE => invoke_init kind; val result as (_, data') = f (dest data); val _ = Synchronized.change dataref ((K o SOME) (Datatab.update (kind, mk data') datatab, thy_id)); in result end; end; (*local*) (* pending function equations *) (* Ideally, *all* equations implementing a functions would be treated as *one* atomic declaration; unfortunately, we cannot implement this: the too-well-established declaration interface are Isar attributes which operate on *one* single theorem. Hence we treat such Isar declarations as "pending" and historize them as proper declarations at the end of each theory. *) fun modify_pending_eqns c f specs = let val existing_eqns = case History.lookup (functions_of specs) c of SOME (Eqns (false, eqns)) => eqns | _ => []; in specs |> map_pending_eqns (Symtab.map_default (c, existing_eqns) f) end; fun register_fun_spec c spec = map_pending_eqns (Symtab.delete_safe c) #> map_functions (History.register c spec); fun lookup_fun_spec specs c = case Symtab.lookup (pending_eqns_of specs) c of SOME eqns => Eqns (false, eqns) | NONE => (case History.lookup (functions_of specs) c of SOME spec => spec | NONE => unimplemented); fun lookup_proper_fun_spec specs c = let val spec = lookup_fun_spec specs c in if is_unimplemented spec then NONE else SOME spec end; fun all_fun_specs specs = map_filter (fn c => Option.map (pair c) (lookup_proper_fun_spec specs c)) (union (op =) ((Symtab.keys o pending_eqns_of) specs) ((Symtab.keys o functions_of) specs)); fun historize_pending_fun_specs thy = let val pending_eqns = (pending_eqns_of o specs_of) thy; in if Symtab.is_empty pending_eqns then NONE else thy |> modify_specs (map_functions (Symtab.fold (fn (c, eqs) => History.register c (Eqns (false, eqs))) pending_eqns) #> map_pending_eqns (K Symtab.empty)) |> SOME end; val _ = Theory.setup (Theory.at_end historize_pending_fun_specs); (** foundation **) (* types *) fun no_constr thy s (c, ty) = error ("Not a datatype constructor:\n" ^ string_of_const thy c ^ " :: " ^ string_of_typ thy ty ^ "\n" ^ enclose "(" ")" s); fun analyze_constructor thy (c, ty) = let val _ = Thm.global_cterm_of thy (Const (c, ty)); val ty_decl = devarify (const_typ thy c); fun last_typ c_ty ty = let val tfrees = Term.add_tfreesT ty []; val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type ty)) handle TYPE _ => no_constr thy "bad type" c_ty val _ = if tyco = "fun" then no_constr thy "bad type" c_ty else (); val _ = if has_duplicates (eq_fst (op =)) vs then no_constr thy "duplicate type variables in datatype" c_ty else (); val _ = if length tfrees <> length vs then no_constr thy "type variables missing in datatype" c_ty else (); in (tyco, vs) end; val (tyco, _) = last_typ (c, ty) ty_decl; val (_, vs) = last_typ (c, ty) ty; in ((tyco, map snd vs), (c, (map fst vs, ty))) end; fun constrset_of_consts thy consts = let val _ = map (fn (c, _) => if (is_some o Axclass.class_of_param thy) c then error ("Is a class parameter: " ^ string_of_const thy c) else ()) consts; val raw_constructors = map (analyze_constructor thy) consts; val tyco = case distinct (op =) (map (fst o fst) raw_constructors) of [tyco] => tyco | [] => error "Empty constructor set" | tycos => error ("Different type constructors in constructor set: " ^ commas_quote tycos) val vs = Name.invent Name.context Name.aT (Sign.arity_number thy tyco); fun inst vs' (c, (vs, ty)) = let val the_v = the o AList.lookup (op =) (vs ~~ vs'); val ty' = map_type_tfree (fn (v, _) => TFree (the_v v, [])) ty; val (vs'', ty'') = typscheme thy (c, ty'); in (c, (vs'', binder_types ty'')) end; val constructors = map (inst vs o snd) raw_constructors; in (tyco, (map (rpair []) vs, constructors)) end; fun lookup_vs_type_spec thy = History.lookup ((types_of o specs_of) thy); type constructors = (string * sort) list * (string * ((string * sort) list * typ list)) list; fun get_type thy tyco = case lookup_vs_type_spec thy tyco of SOME (vs, type_spec) => apfst (pair vs) (constructors_of type_spec) | NONE => Sign.arity_number thy tyco |> Name.invent Name.context Name.aT |> map (rpair []) |> rpair [] |> rpair false; type abs_type = (string * sort) list * {abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string}; fun get_abstype_spec thy tyco = case lookup_vs_type_spec thy tyco of SOME (vs, Abstractor {abs_rep, abstractor, projection, ...}) => (vs, {abs_rep = abs_rep, abstractor = abstractor, projection = projection}) | _ => error ("Not an abstract type: " ^ tyco); fun get_type_of_constr_or_abstr thy c = case (body_type o const_typ thy) c of Type (tyco, _) => let val ((_, cos), abstract) = get_type thy tyco in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end | _ => NONE; fun is_constr thy c = case get_type_of_constr_or_abstr thy c of SOME (_, false) => true | _ => false; fun is_abstr thy c = case get_type_of_constr_or_abstr thy c of SOME (_, true) => true | _ => false; (* bare code equations *) (* convention for variables: ?x ?'a for free-floating theorems (e.g. in the data store) ?x 'a for certificates x 'a for final representation of equations *) exception BAD_THM of string; fun bad_thm msg = raise BAD_THM msg; datatype strictness = Silent | Liberal | Strict fun handle_strictness thm_of f strictness thy x = SOME (f x) handle BAD_THM msg => case strictness of Silent => NONE | Liberal => (warning (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); NONE) | Strict => error (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); fun is_linear thm = let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm in not (has_duplicates (op =) ((fold o fold_aterms) (fn Var (v, _) => cons v | _ => I) args [])) end; fun check_decl_ty thy (c, ty) = let val ty_decl = const_typ thy c; in if typscheme_equiv (ty_decl, ty) then () else bad_thm ("Type\n" ^ string_of_typ thy ty ^ "\nof constant " ^ quote c ^ "\nis too specific compared to declared type\n" ^ string_of_typ thy ty_decl) end; fun check_eqn thy {allow_nonlinear, allow_consts, allow_pats} thm (lhs, rhs) = let fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v | Free _ => bad_thm "Illegal free variable" | _ => I) t []; fun tvars_of t = fold_term_types (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v | TFree _ => bad_thm "Illegal free type variable")) t []; val lhs_vs = vars_of lhs; val rhs_vs = vars_of rhs; val lhs_tvs = tvars_of lhs; val rhs_tvs = tvars_of rhs; val _ = if null (subtract (op =) lhs_vs rhs_vs) then () else bad_thm "Free variables on right hand side of equation"; val _ = if null (subtract (op =) lhs_tvs rhs_tvs) then () else bad_thm "Free type variables on right hand side of equation"; val (head, args) = strip_comb lhs; val (c, ty) = case head of Const (c_ty as (_, ty)) => (Axclass.unoverload_const thy c_ty, ty) | _ => bad_thm "Equation not headed by constant"; fun check _ (Abs _) = bad_thm "Abstraction on left hand side of equation" | check 0 (Var _) = () | check _ (Var _) = bad_thm "Variable with application on left hand side of equation" | check n (t1 $ t2) = (check (n+1) t1; check 0 t2) | check n (Const (c_ty as (c, ty))) = if allow_pats then let val c' = Axclass.unoverload_const thy c_ty in if n = (length o binder_types) ty then if allow_consts orelse is_constr thy c' then () else bad_thm (quote c ^ " is not a constructor, on left hand side of equation") else bad_thm ("Partially applied constant " ^ quote c ^ " on left hand side of equation") end else bad_thm ("Pattern not allowed here, but constant " ^ quote c ^ " encountered on left hand side of equation") val _ = map (check 0) args; val _ = if allow_nonlinear orelse is_linear thm then () else bad_thm "Duplicate variables on left hand side of equation"; val _ = if (is_none o Axclass.class_of_param thy) c then () else bad_thm "Overloaded constant as head in equation"; val _ = if not (is_constr thy c) then () else bad_thm "Constructor as head in equation"; val _ = if not (is_abstr thy c) then () else bad_thm "Abstractor as head in equation"; val _ = check_decl_ty thy (c, ty); val _ = case strip_type ty of (Type (tyco, _) :: _, _) => (case lookup_vs_type_spec thy tyco of SOME (_, type_spec) => (case projection_of type_spec of SOME proj => if c = proj then bad_thm "Projection as head in equation" else () | _ => ()) | _ => ()) | _ => (); in () end; local fun raw_assert_eqn thy check_patterns (thm, proper) = let val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val _ = check_eqn thy {allow_nonlinear = not proper, allow_consts = not (proper andalso check_patterns), allow_pats = true} thm (lhs, rhs); in (thm, proper) end; fun raw_assert_abs_eqn thy some_tyco thm = let val (full_lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val (proj_t, lhs) = dest_comb full_lhs handle TERM _ => bad_thm "Not an abstract equation"; val (proj, ty) = dest_Const proj_t handle TERM _ => bad_thm "Not an abstract equation"; val (tyco, Ts) = (dest_Type o domain_type) ty handle TERM _ => bad_thm "Not an abstract equation" | TYPE _ => bad_thm "Not an abstract equation"; val _ = case some_tyco of SOME tyco' => if tyco = tyco' then () else bad_thm ("Abstract type mismatch:" ^ quote tyco ^ " vs. " ^ quote tyco') | NONE => (); val (vs, proj', (abs', _)) = case lookup_vs_type_spec thy tyco of SOME (vs, Abstractor spec) => (vs, #projection spec, #abstractor spec) | _ => bad_thm ("Not an abstract type: " ^ tyco); val _ = if proj = proj' then () else bad_thm ("Projection mismatch: " ^ quote proj ^ " vs. " ^ quote proj'); val _ = check_eqn thy {allow_nonlinear = false, allow_consts = false, allow_pats = false} thm (lhs, rhs); val _ = if ListPair.all (fn (T, (_, sort)) => Sign.of_sort thy (T, sort)) (Ts, vs) then () else error ("Type arguments do not satisfy sort constraints of abstype certificate."); in (thm, (tyco, abs')) end; in fun generic_assert_eqn strictness thy check_patterns eqn = handle_strictness fst (raw_assert_eqn thy check_patterns) strictness thy eqn; fun generic_assert_abs_eqn strictness thy check_patterns thm = handle_strictness I (raw_assert_abs_eqn thy check_patterns) strictness thy thm; end; fun assert_eqn thy = the o generic_assert_eqn Strict thy true; fun assert_abs_eqn thy some_tyco = the o generic_assert_abs_eqn Strict thy some_tyco; val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; fun const_typ_eqn thy thm = let val (c, ty) = head_eqn thm; val c' = Axclass.unoverload_const thy (c, ty); (*permissive wrt. to overloaded constants!*) in (c', ty) end; fun const_eqn thy = fst o const_typ_eqn thy; fun const_abs_eqn thy = Axclass.unoverload_const thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.plain_prop_of; fun mk_proj tyco vs ty abs rep = let val ty_abs = Type (tyco, map TFree vs); val xarg = Var (("x", 0), ty); in Logic.mk_equals (Const (rep, ty_abs --> ty) $ (Const (abs, ty --> ty_abs) $ xarg), xarg) end; (* technical transformations of code equations *) fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy); fun expand_eta thy k thm = let val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm; val (_, args) = strip_comb lhs; val l = if k = ~1 then (length o fst o strip_abs) rhs else Int.max (0, k - length args); val (raw_vars, _) = Term.strip_abs_eta l rhs; val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) raw_vars; fun expand (v, ty) thm = Drule.fun_cong_rule thm (Thm.global_cterm_of thy (Var ((v, 0), ty))); in thm |> fold expand vars |> Conv.fconv_rule Drule.beta_eta_conversion end; fun same_arity thy thms = let val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals; val k = fold (Integer.max o num_args_of o Thm.prop_of) thms 0; in map (expand_eta thy k) thms end; fun mk_desymbolization pre post mk vs = let val names = map (pre o fst o fst) vs |> map (Name.desymbolize (SOME false)) |> Name.variant_list [] |> map post; in map_filter (fn (((v, i), x), v') => if v = v' andalso i = 0 then NONE else SOME (((v, i), x), mk ((v', 0), x))) (vs ~~ names) end; fun desymbolize_tvars thy thms = let val tvs = fold (Term.add_tvars o Thm.prop_of) thms []; val instT = mk_desymbolization (unprefix "'") (prefix "'") (Thm.global_ctyp_of thy o TVar) tvs; in map (Thm.instantiate (instT, [])) thms end; fun desymbolize_vars thy thm = let val vs = Term.add_vars (Thm.prop_of thm) []; val inst = mk_desymbolization I I (Thm.global_cterm_of thy o Var) vs; in Thm.instantiate ([], inst) thm end; fun canonize_thms thy = desymbolize_tvars thy #> same_arity thy #> map (desymbolize_vars thy); (* preparation and classification of code equations *) fun prep_eqn strictness thy = apfst (meta_rewrite thy) #> generic_assert_eqn strictness thy false #> Option.map (fn eqn => (const_eqn thy (fst eqn), eqn)); fun prep_eqns strictness thy = map_filter (prep_eqn strictness thy) #> AList.group (op =); fun prep_abs_eqn strictness thy = meta_rewrite thy #> generic_assert_abs_eqn strictness thy NONE #> Option.map (fn abs_eqn => (const_abs_eqn thy (fst abs_eqn), abs_eqn)); fun prep_maybe_abs_eqn thy raw_thm = let val thm = meta_rewrite thy raw_thm; val some_abs_thm = generic_assert_abs_eqn Silent thy NONE thm; in case some_abs_thm of SOME (thm, tyco) => SOME (const_abs_eqn thy thm, ((thm, true), SOME tyco)) | NONE => generic_assert_eqn Liberal thy false (thm, false) |> Option.map (fn (thm, _) => (const_eqn thy thm, ((thm, is_linear thm), NONE))) end; (* abstype certificates *) local fun raw_abstype_cert thy proto_thm = let val thm = (Axclass.unoverload (Proof_Context.init_global thy) o meta_rewrite thy) proto_thm; val (lhs, rhs) = Logic.dest_equals (Thm.plain_prop_of thm) handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val ((abs, raw_ty), ((rep, rep_ty), param)) = (apsnd (apfst dest_Const o dest_comb) o apfst dest_Const o dest_comb) lhs handle TERM _ => bad_thm "Not an abstype certificate"; val _ = apply2 (fn c => if (is_some o Axclass.class_of_param thy) c then error ("Is a class parameter: " ^ string_of_const thy c) else ()) (abs, rep); val _ = check_decl_ty thy (abs, raw_ty); val _ = check_decl_ty thy (rep, rep_ty); val _ = if length (binder_types raw_ty) = 1 then () else bad_thm "Bad type for abstract constructor"; val _ = (fst o dest_Var) param handle TERM _ => bad_thm "Not an abstype certificate"; val _ = if param = rhs then () else bad_thm "Not an abstype certificate"; val ((tyco, sorts), (abs, (vs, ty'))) = analyze_constructor thy (abs, devarify raw_ty); val ty = domain_type ty'; val (vs', _) = typscheme thy (abs, ty'); in (tyco, (vs ~~ sorts, ((abs, (vs', ty)), (rep, thm)))) end; in fun check_abstype_cert strictness thy proto_thm = handle_strictness I (raw_abstype_cert thy) strictness thy proto_thm; end; (* code equation certificates *) fun build_head thy (c, ty) = Thm.global_cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty))); fun get_head thy cert_thm = let val [head] = Thm.chyps_of cert_thm; val (_, Const (c, ty)) = (Logic.dest_equals o Thm.term_of) head; in (typscheme thy (c, ty), head) end; fun typscheme_projection thy = typscheme thy o dest_Const o fst o dest_comb o fst o Logic.dest_equals; fun typscheme_abs thy = typscheme thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.prop_of; fun constrain_thm thy vs sorts thm = let val mapping = map2 (fn (v, sort) => fn sort' => (v, Sorts.inter_sort (Sign.classes_of thy) (sort, sort'))) vs sorts; val inst = map2 (fn (v, sort) => fn (_, sort') => (((v, 0), sort), Thm.global_ctyp_of thy (TFree (v, sort')))) vs mapping; val subst = (Term.map_types o map_type_tfree) (fn (v, _) => TFree (v, the (AList.lookup (op =) mapping v))); in thm |> Thm.varifyT_global |> Thm.instantiate (inst, []) |> pair subst end; fun concretify_abs thy tyco abs_thm = let val (_, {abstractor = (c_abs, _), abs_rep, ...}) = get_abstype_spec thy tyco; val lhs = (fst o Logic.dest_equals o Thm.prop_of) abs_thm val ty = fastype_of lhs; val ty_abs = (fastype_of o snd o dest_comb) lhs; val abs = Thm.global_cterm_of thy (Const (c_abs, ty --> ty_abs)); val raw_concrete_thm = Drule.transitive_thm OF [Thm.symmetric abs_rep, Thm.combination (Thm.reflexive abs) abs_thm]; in (c_abs, (Thm.varifyT_global o zero_var_indexes) raw_concrete_thm) end; fun add_rhss_of_eqn thy t = let val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals) t; fun add_const (Const (c, ty)) = insert (op =) (c, Sign.const_typargs thy (c, ty)) | add_const _ = I val add_consts = fold_aterms add_const in add_consts rhs o fold add_consts args end; val dest_eqn = apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify_global; abstype cert = Nothing of thm | Equations of thm * bool list | Projection of term * string | Abstract of thm * string with fun dummy_thm ctxt c = let val thy = Proof_Context.theory_of ctxt; val raw_ty = devarify (const_typ thy c); val (vs, _) = typscheme thy (c, raw_ty); val sortargs = case Axclass.class_of_param thy c of SOME class => [[class]] | NONE => (case get_type_of_constr_or_abstr thy c of SOME (tyco, _) => (map snd o fst o the) (AList.lookup (op =) ((snd o fst o get_type thy) tyco) c) | NONE => replicate (length vs) []); val the_sort = the o AList.lookup (op =) (map fst vs ~~ sortargs); val ty = map_type_tfree (fn (v, _) => TFree (v, the_sort v)) raw_ty val chead = build_head thy (c, ty); in Thm.weaken chead Drule.dummy_thm end; fun nothing_cert ctxt c = Nothing (dummy_thm ctxt c); fun cert_of_eqns ctxt c [] = Equations (dummy_thm ctxt c, []) | cert_of_eqns ctxt c raw_eqns = let val thy = Proof_Context.theory_of ctxt; val eqns = burrow_fst (canonize_thms thy) raw_eqns; val _ = map (assert_eqn thy) eqns; val (thms, propers) = split_list eqns; val _ = map (fn thm => if c = const_eqn thy thm then () else error ("Wrong head of code equation,\nexpected constant " ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy thm)) thms; - fun tvars_of T = rev (Term.add_tvarsT T []); + val tvars_of = build_rev o Term.add_tvarsT; val vss = map (tvars_of o snd o head_eqn) thms; fun inter_sorts vs = fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd) vs []; val sorts = map_transpose inter_sorts vss; val vts = Name.invent_names Name.context Name.aT sorts; val thms' = map2 (fn vs => Thm.instantiate (vs ~~ map (Thm.ctyp_of ctxt o TFree) vts, [])) vss thms; val head_thm = Thm.symmetric (Thm.assume (build_head thy (head_eqn (hd thms')))); fun head_conv ct = if can Thm.dest_comb ct then Conv.fun_conv head_conv ct else Conv.rewr_conv head_thm ct; val rewrite_head = Conv.fconv_rule (Conv.arg1_conv head_conv); val cert_thm = Conjunction.intr_balanced (map rewrite_head thms'); in Equations (cert_thm, propers) end; fun cert_of_proj ctxt proj tyco = let val thy = Proof_Context.theory_of ctxt val (vs, {abstractor = (abs, (_, ty)), projection = proj', ...}) = get_abstype_spec thy tyco; val _ = if proj = proj' then () else error ("Wrong head of projection,\nexpected constant " ^ string_of_const thy proj); in Projection (mk_proj tyco vs ty abs proj, tyco) end; fun cert_of_abs ctxt tyco c raw_abs_thm = let val thy = Proof_Context.theory_of ctxt; val abs_thm = singleton (canonize_thms thy) raw_abs_thm; val _ = assert_abs_eqn thy (SOME tyco) abs_thm; val _ = if c = const_abs_eqn thy abs_thm then () else error ("Wrong head of abstract code equation,\nexpected constant " ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy abs_thm); in Abstract (Thm.legacy_freezeT abs_thm, tyco) end; fun constrain_cert_thm thy sorts cert_thm = let val ((vs, _), head) = get_head thy cert_thm; val (subst, cert_thm') = cert_thm |> Thm.implies_intr head |> constrain_thm thy vs sorts; val head' = Thm.term_of head |> subst |> Thm.global_cterm_of thy; val cert_thm'' = cert_thm' |> Thm.elim_implies (Thm.assume head'); in cert_thm'' end; fun constrain_cert thy sorts (Nothing cert_thm) = Nothing (constrain_cert_thm thy sorts cert_thm) | constrain_cert thy sorts (Equations (cert_thm, propers)) = Equations (constrain_cert_thm thy sorts cert_thm, propers) | constrain_cert _ _ (cert as Projection _) = cert | constrain_cert thy sorts (Abstract (abs_thm, tyco)) = Abstract (snd (constrain_thm thy (fst (typscheme_abs thy abs_thm)) sorts abs_thm), tyco); fun conclude_cert (Nothing cert_thm) = Nothing (Thm.close_derivation \<^here> cert_thm) | conclude_cert (Equations (cert_thm, propers)) = Equations (Thm.close_derivation \<^here> cert_thm, propers) | conclude_cert (cert as Projection _) = cert | conclude_cert (Abstract (abs_thm, tyco)) = Abstract (Thm.close_derivation \<^here> abs_thm, tyco); fun typscheme_of_cert thy (Nothing cert_thm) = fst (get_head thy cert_thm) | typscheme_of_cert thy (Equations (cert_thm, _)) = fst (get_head thy cert_thm) | typscheme_of_cert thy (Projection (proj, _)) = typscheme_projection thy proj | typscheme_of_cert thy (Abstract (abs_thm, _)) = typscheme_abs thy abs_thm; fun typargs_deps_of_cert thy (Nothing cert_thm) = let val vs = (fst o fst) (get_head thy cert_thm); in (vs, []) end | typargs_deps_of_cert thy (Equations (cert_thm, propers)) = let val vs = (fst o fst) (get_head thy cert_thm); val equations = if null propers then [] else Thm.prop_of cert_thm |> Logic.dest_conjunction_balanced (length propers); in (vs, fold (add_rhss_of_eqn thy) equations []) end | typargs_deps_of_cert thy (Projection (t, _)) = (fst (typscheme_projection thy t), add_rhss_of_eqn thy t []) | typargs_deps_of_cert thy (Abstract (abs_thm, tyco)) = let val vs = fst (typscheme_abs thy abs_thm); val (_, concrete_thm) = concretify_abs thy tyco abs_thm; in (vs, add_rhss_of_eqn thy (Logic.unvarify_types_global (Thm.prop_of concrete_thm)) []) end; fun equations_of_cert thy (cert as Nothing _) = (typscheme_of_cert thy cert, NONE) | equations_of_cert thy (cert as Equations (cert_thm, propers)) = let val tyscm = typscheme_of_cert thy cert; val thms = if null propers then [] else cert_thm |> Local_Defs.expand [snd (get_head thy cert_thm)] |> Thm.varifyT_global |> Conjunction.elim_balanced (length propers); fun abstractions (args, rhs) = (map (rpair NONE) args, (rhs, NONE)); in (tyscm, SOME (map (abstractions o dest_eqn o Thm.prop_of) thms ~~ (map SOME thms ~~ propers))) end | equations_of_cert thy (Projection (t, tyco)) = let val (_, {abstractor = (abs, _), ...}) = get_abstype_spec thy tyco; val tyscm = typscheme_projection thy t; val t' = Logic.varify_types_global t; fun abstractions (args, rhs) = (map (rpair (SOME abs)) args, (rhs, NONE)); in (tyscm, SOME [((abstractions o dest_eqn) t', (NONE, true))]) end | equations_of_cert thy (Abstract (abs_thm, tyco)) = let val tyscm = typscheme_abs thy abs_thm; val (abs, concrete_thm) = concretify_abs thy tyco abs_thm; fun abstractions (args, rhs) = (map (rpair NONE) args, (rhs, (SOME abs))); in (tyscm, SOME [((abstractions o dest_eqn o Thm.prop_of) concrete_thm, (SOME (Thm.varifyT_global abs_thm), true))]) end; fun pretty_cert _ (Nothing _) = [] | pretty_cert thy (cert as Equations _) = (map_filter (Option.map (Thm.pretty_thm_global thy o Axclass.overload (Proof_Context.init_global thy)) o fst o snd) o these o snd o equations_of_cert thy) cert | pretty_cert thy (Projection (t, _)) = [Syntax.pretty_term_global thy (Logic.varify_types_global t)] | pretty_cert thy (Abstract (abs_thm, _)) = [(Thm.pretty_thm_global thy o Axclass.overload (Proof_Context.init_global thy) o Thm.varifyT_global) abs_thm]; end; (* code certificate access with preprocessing *) fun eqn_conv conv ct = let fun lhs_conv ct = if can Thm.dest_comb ct then Conv.combination_conv lhs_conv conv ct else Conv.all_conv ct; in Conv.combination_conv (Conv.arg_conv lhs_conv) conv ct end; fun rewrite_eqn conv ctxt = singleton (Variable.trade (K (map (Conv.fconv_rule (conv (Simplifier.rewrite ctxt))))) ctxt) fun preprocess conv ctxt = Thm.transfer' ctxt #> rewrite_eqn conv ctxt #> Axclass.unoverload ctxt; fun cert_of_eqns_preprocess ctxt functrans c = let fun trace_eqns s eqns = (Pretty.writeln o Pretty.chunks) (Pretty.str s :: map (Thm.pretty_thm ctxt o fst) eqns); val tracing = if Config.get ctxt simp_trace then trace_eqns else (K o K) (); in tap (tracing "before function transformation") #> (perhaps o perhaps_loop o perhaps_apply) functrans #> tap (tracing "after function transformation") #> (map o apfst) (preprocess eqn_conv ctxt) #> cert_of_eqns ctxt c end; fun get_cert ctxt functrans c = case lookup_proper_fun_spec (specs_of (Proof_Context.theory_of ctxt)) c of NONE => nothing_cert ctxt c | SOME (Eqns (_, eqns)) => eqns |> cert_of_eqns_preprocess ctxt functrans c | SOME (Proj (_, (tyco, _))) => cert_of_proj ctxt c tyco | SOME (Abstr (abs_thm, (tyco, _))) => abs_thm |> preprocess Conv.arg_conv ctxt |> cert_of_abs ctxt tyco c; (* case certificates *) local fun raw_case_cert thm = let val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm; val _ = case head of Free _ => () | Var _ => () | _ => raise TERM ("case_cert", []); val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; val (Const (case_const, _), raw_params) = strip_comb case_expr; val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; val _ = if n = ~1 then raise TERM ("case_cert", []) else (); val params = map (fst o dest_Var) (nth_drop n raw_params); fun dest_case t = let val (head' $ t_co, rhs) = Logic.dest_equals t; val _ = if head' = head then () else raise TERM ("case_cert", []); val (Const (co, _), args) = strip_comb t_co; val (Var (param, _), args') = strip_comb rhs; val _ = if args' = args then () else raise TERM ("case_cert", []); in (param, co) end; fun analyze_cases cases = let val co_list = fold (AList.update (op =) o dest_case) cases []; in map (AList.lookup (op =) co_list) params end; fun analyze_let t = let val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; val _ = if head' = head then () else raise TERM ("case_cert", []); val _ = if arg' = arg then () else raise TERM ("case_cert", []); val _ = if [param'] = params then () else raise TERM ("case_cert", []); in [] end; fun analyze (cases as [let_case]) = (analyze_cases cases handle Bind => analyze_let let_case) | analyze cases = analyze_cases cases; in (case_const, (n, analyze cases)) end; in fun case_cert thm = raw_case_cert thm handle Bind => error "bad case certificate" | TERM _ => error "bad case certificate"; end; fun lookup_case_spec thy = History.lookup ((cases_of o specs_of) thy); fun get_case_schema thy c = case lookup_case_spec thy c of SOME (Case {schema, ...}) => SOME schema | _ => NONE; fun get_case_cong thy c = case lookup_case_spec thy c of SOME (Case {cong, ...}) => SOME cong | _ => NONE; fun is_undefined thy c = case lookup_case_spec thy c of SOME Undefined => true | _ => false; (* diagnostic *) fun print_codesetup thy = let val ctxt = Proof_Context.init_global thy; val specs = specs_of thy; fun pretty_equations const thms = (Pretty.block o Pretty.fbreaks) (Pretty.str (string_of_const thy const) :: map (Thm.pretty_thm_item ctxt) thms); fun pretty_function (const, Eqns (_, eqns)) = pretty_equations const (map fst eqns) | pretty_function (const, Proj (proj, _)) = Pretty.block [Pretty.str (string_of_const thy const), Pretty.fbrk, Syntax.pretty_term ctxt proj] | pretty_function (const, Abstr (thm, _)) = pretty_equations const [thm]; fun pretty_typ (tyco, vs) = Pretty.str (string_of_typ thy (Type (tyco, map TFree vs))); fun pretty_type_spec (typ, (cos, abstract)) = if null cos then pretty_typ typ else (Pretty.block o Pretty.breaks) ( pretty_typ typ :: Pretty.str "=" :: (if abstract then [Pretty.str "(abstract)"] else []) @ separate (Pretty.str "|") (map (fn (c, (_, [])) => Pretty.str (string_of_const thy c) | (c, (_, tys)) => (Pretty.block o Pretty.breaks) (Pretty.str (string_of_const thy c) :: Pretty.str "of" :: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos) ); fun pretty_case_param NONE = "" | pretty_case_param (SOME c) = string_of_const thy c fun pretty_case (const, Case {schema = (_, (_, [])), ...}) = Pretty.str (string_of_const thy const) | pretty_case (const, Case {schema = (_, (_, cos)), ...}) = (Pretty.block o Pretty.breaks) [ Pretty.str (string_of_const thy const), Pretty.str "with", (Pretty.block o Pretty.commas o map (Pretty.str o pretty_case_param)) cos] | pretty_case (const, Undefined) = (Pretty.block o Pretty.breaks) [ Pretty.str (string_of_const thy const), Pretty.str ""]; val functions = all_fun_specs specs |> sort (string_ord o apply2 fst); val types = History.all (types_of specs) |> map (fn (tyco, (vs, spec)) => ((tyco, vs), constructors_of spec)) |> sort (string_ord o apply2 (fst o fst)); val cases = History.all (cases_of specs) |> filter (fn (_, No_Case) => false | _ => true) |> sort (string_ord o apply2 fst); in Pretty.writeln_chunks [ Pretty.block ( Pretty.str "types:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_type_spec) types ), Pretty.block ( Pretty.str "functions:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_function) functions ), Pretty.block ( Pretty.str "cases:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_case) cases ) ] end; (** declaration of executable ingredients **) (* plugins for dependent applications *) structure Codetype_Plugin = Plugin(type T = string); val codetype_plugin = Plugin_Name.declare_setup \<^binding>\codetype\; fun type_interpretation f = Codetype_Plugin.interpretation codetype_plugin (fn tyco => Local_Theory.background_theory (fn thy => thy |> Sign.root_path |> Sign.add_path (Long_Name.qualifier tyco) |> f tyco |> Sign.restore_naming thy)); fun datatype_interpretation f = type_interpretation (fn tyco => fn thy => case get_type thy tyco of (spec, false) => f (tyco, spec) thy | (_, true) => thy ); fun abstype_interpretation f = type_interpretation (fn tyco => fn thy => case try (get_abstype_spec thy) tyco of SOME spec => f (tyco, spec) thy | NONE => thy ); fun register_tyco_for_plugin tyco = Named_Target.theory_map (Codetype_Plugin.data_default tyco); (* abstract code declarations *) local fun generic_code_declaration strictness lift_phi f x = Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => Context.mapping (f strictness (lift_phi phi x)) I); in fun silent_code_declaration lift_phi = generic_code_declaration Silent lift_phi; fun code_declaration lift_phi = generic_code_declaration Liberal lift_phi; end; (* types *) fun invalidate_constructors_of (_, type_spec) = fold (fn (c, _) => History.register c unimplemented) (fst (constructors_of type_spec)); fun invalidate_abstract_functions_of (_, type_spec) = fold (fn c => History.register c unimplemented) (abstract_functions_of type_spec); fun invalidate_case_combinators_of (_, type_spec) = fold (fn c => History.register c No_Case) (case_combinators_of type_spec); fun register_type (tyco, vs_typ_spec) specs = let val olds = the_list (History.lookup (types_of specs) tyco); in specs |> map_functions (fold invalidate_abstract_functions_of olds #> invalidate_constructors_of vs_typ_spec) |> map_cases (fold invalidate_case_combinators_of olds) |> map_types (History.register tyco vs_typ_spec) end; fun declare_datatype_global proto_constrs thy = let fun unoverload_const_typ (c, ty) = (Axclass.unoverload_const thy (c, ty), ty); val constrs = map unoverload_const_typ proto_constrs; val (tyco, (vs, cos)) = constrset_of_consts thy constrs; in thy |> modify_specs (register_type (tyco, (vs, Constructors {constructors = cos, case_combinators = []}))) |> register_tyco_for_plugin tyco end; fun declare_datatype_cmd raw_constrs thy = declare_datatype_global (map (read_bare_const thy) raw_constrs) thy; fun generic_declare_abstype strictness proto_thm thy = case check_abstype_cert strictness thy proto_thm of SOME (tyco, (vs, (abstractor as (abs, (_, ty)), (proj, abs_rep)))) => thy |> modify_specs (register_type (tyco, (vs, Abstractor {abstractor = abstractor, projection = proj, abs_rep = abs_rep, more_abstract_functions = []})) #> register_fun_spec proj (Proj (Logic.varify_types_global (mk_proj tyco vs ty abs proj), (tyco, abs)))) |> register_tyco_for_plugin tyco | NONE => thy; val declare_abstype_global = generic_declare_abstype Strict; val declare_abstype = code_declaration Morphism.thm generic_declare_abstype; (* functions *) (* strictness wrt. shape of theorem propositions: * default equations: silent * using declarations and attributes: warnings (after morphism application!) * using global declarations (... -> thy -> thy): strict * internal processing after storage: strict *) local fun subsumptive_add thy verbose (thm, proper) eqns = let val args_of = drop_prefix is_Var o rev o snd o strip_comb o Term.map_types Type.strip_sorts o fst o Logic.dest_equals o Thm.plain_prop_of o Thm.transfer thy; val args = args_of thm; val incr_idx = Logic.incr_indexes ([], [], Thm.maxidx_of thm + 1); fun matches_args args' = let val k = length args' - length args in if k >= 0 then Pattern.matchess thy (args, (map incr_idx o drop k) args') else false end; fun drop (thm', proper') = if (proper orelse not proper') andalso matches_args (args_of thm') then (if verbose then warning ("Code generator: dropping subsumed code equation\n" ^ Thm.string_of_thm_global thy thm') else (); true) else false; in (thm |> Thm.close_derivation \<^here> |> Thm.trim_context, proper) :: filter_out drop eqns end; fun add_eqn_for (c, eqn) thy = thy |> modify_specs (modify_pending_eqns c (subsumptive_add thy true eqn)); fun add_eqns_for default (c, proto_eqns) thy = thy |> modify_specs (fn specs => if is_default (lookup_fun_spec specs c) orelse not default then let val eqns = [] |> fold_rev (subsumptive_add thy (not default)) proto_eqns; in specs |> register_fun_spec c (Eqns (default, eqns)) end else specs); fun add_abstract_for (c, (thm, tyco_abs as (tyco, _))) = modify_specs (register_fun_spec c (Abstr (Thm.close_derivation \<^here> thm, tyco_abs)) #> map_types (History.modify_entry tyco (add_abstract_function c))) in fun generic_declare_eqns default strictness raw_eqns thy = fold (add_eqns_for default) (prep_eqns strictness thy raw_eqns) thy; fun generic_add_eqn strictness raw_eqn thy = fold add_eqn_for (the_list (prep_eqn strictness thy raw_eqn)) thy; fun generic_declare_abstract_eqn strictness raw_abs_eqn thy = fold add_abstract_for (the_list (prep_abs_eqn strictness thy raw_abs_eqn)) thy; fun add_maybe_abs_eqn_liberal thm thy = case prep_maybe_abs_eqn thy thm of SOME (c, (eqn, NONE)) => add_eqn_for (c, eqn) thy | SOME (c, ((thm, _), SOME tyco)) => add_abstract_for (c, (thm, tyco)) thy | NONE => thy; end; val declare_default_eqns_global = generic_declare_eqns true Silent; val declare_default_eqns = silent_code_declaration (map o apfst o Morphism.thm) (generic_declare_eqns true); val declare_eqns_global = generic_declare_eqns false Strict; val declare_eqns = code_declaration (map o apfst o Morphism.thm) (generic_declare_eqns false); val add_eqn_global = generic_add_eqn Strict; fun del_eqn_global thm thy = case prep_eqn Liberal thy (thm, false) of SOME (c, (thm, _)) => modify_specs (modify_pending_eqns c (filter_out (fn (thm', _) => Thm.eq_thm_prop (thm, thm')))) thy | NONE => thy; val declare_abstract_eqn_global = generic_declare_abstract_eqn Strict; val declare_abstract_eqn = code_declaration Morphism.thm generic_declare_abstract_eqn; fun declare_aborting_global c = modify_specs (register_fun_spec c aborting); fun declare_unimplemented_global c = modify_specs (register_fun_spec c unimplemented); (* cases *) fun case_cong thy case_const (num_args, (pos, _)) = let val ([x, y], ctxt) = fold_map Name.variant ["A", "A'"] Name.context; val (zs, _) = fold_map Name.variant (replicate (num_args - 1) "") ctxt; val (ws, vs) = chop pos zs; val T = devarify (const_typ thy case_const); val Ts = binder_types T; val T_cong = nth Ts pos; fun mk_prem z = Free (z, T_cong); fun mk_concl z = list_comb (Const (case_const, T), map2 (curry Free) (ws @ z :: vs) Ts); val (prem, concl) = apply2 Logic.mk_equals (apply2 mk_prem (x, y), apply2 mk_concl (x, y)); in Goal.prove_sorry_global thy (x :: y :: zs) [prem] concl (fn {context = ctxt', prems} => Simplifier.rewrite_goals_tac ctxt' prems THEN ALLGOALS (Proof_Context.fact_tac ctxt' [Drule.reflexive_thm])) end; fun declare_case_global thm thy = let val (case_const, (k, cos)) = case_cert thm; fun get_type_of_constr c = case get_type_of_constr_or_abstr thy c of SOME (c, false) => SOME c | _ => NONE; val cos_with_tycos = (map_filter o Option.map) (fn c => (c, get_type_of_constr c)) cos; val _ = case map_filter (fn (c, NONE) => SOME c | _ => NONE) cos_with_tycos of [] => () | cs => error ("Non-constructor(s) in case certificate: " ^ commas_quote cs); val tycos = distinct (op =) (map_filter snd cos_with_tycos); val schema = (1 + Int.max (1, length cos), (k, cos)); val cong = case_cong thy case_const schema; in thy |> modify_specs (map_cases (History.register case_const (Case {schema = schema, tycos = tycos, cong = cong})) #> map_types (fold (fn tyco => History.modify_entry tyco (add_case_combinator case_const)) tycos)) end; fun declare_undefined_global c = (modify_specs o map_cases) (History.register c Undefined); (* attributes *) fun code_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); fun code_thm_attribute g f = Scan.lift (g |-- Scan.succeed (code_attribute f)); fun code_const_attribute g f = Scan.lift (g -- Args.colon) |-- Scan.repeat1 Args.term >> (fn ts => code_attribute (K (fold (fn t => fn thy => f ((check_const thy o Logic.unvarify_types_global) t) thy) ts))); val _ = Theory.setup (let val code_attribute_parser = code_thm_attribute (Args.$$$ "equation") (fn thm => generic_add_eqn Liberal (thm, true)) || code_thm_attribute (Args.$$$ "nbe") (fn thm => generic_add_eqn Liberal (thm, false)) || code_thm_attribute (Args.$$$ "abstract") (generic_declare_abstract_eqn Liberal) || code_thm_attribute (Args.$$$ "abstype") (generic_declare_abstype Liberal) || code_thm_attribute Args.del del_eqn_global || code_const_attribute (Args.$$$ "abort") declare_aborting_global || code_const_attribute (Args.$$$ "drop") declare_unimplemented_global || Scan.succeed (code_attribute add_maybe_abs_eqn_liberal); in Attrib.setup \<^binding>\code\ code_attribute_parser "declare theorems for code generation" end); end; (*struct*) (* type-safe interfaces for data dependent on executable code *) functor Code_Data(Data: CODE_DATA_ARGS): CODE_DATA = struct type T = Data.T; exception Data of T; fun dest (Data x) = x val kind = Code.declare_data (Data Data.empty); val data_op = (kind, Data, dest); fun change_yield (SOME thy) f = Code.change_yield_data data_op thy f | change_yield NONE f = f Data.empty fun change some_thy f = snd (change_yield some_thy (pair () o f)); end; structure Code : CODE = struct open Code; end; diff --git a/src/Pure/Isar/generic_target.ML b/src/Pure/Isar/generic_target.ML --- a/src/Pure/Isar/generic_target.ML +++ b/src/Pure/Isar/generic_target.ML @@ -1,468 +1,473 @@ (* Title: Pure/Isar/generic_target.ML Author: Makarius Author: Florian Haftmann, TU Muenchen Common target infrastructure. *) signature GENERIC_TARGET = sig (*auxiliary*) val export_abbrev: Proof.context -> (term -> term) -> term -> term * ((string * sort) list * (term list * term list)) val check_mixfix: Proof.context -> binding * (string * sort) list -> mixfix -> mixfix val check_mixfix_global: binding * bool -> mixfix -> mixfix (*background primitives*) val background_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val background_declaration: declaration -> local_theory -> local_theory val background_abbrev: binding * term -> term list -> local_theory -> (term * term) * local_theory val add_foundation_interpretation: (binding * (term * term list) -> Context.generic -> Context.generic) -> theory -> theory (*nested local theories primitives*) val standard_facts: local_theory -> Proof.context -> Attrib.fact list -> Attrib.fact list val standard_notes: (int * int -> bool) -> string -> Attrib.fact list -> local_theory -> local_theory val standard_declaration: (int * int -> bool) -> (morphism -> Context.generic -> Context.generic) -> local_theory -> local_theory val standard_const: (int * int -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val local_interpretation: Locale.registration -> local_theory -> local_theory (*lifting target primitives to local theory operations*) val define: (((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory) -> bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val notes: (string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory) -> string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: (Syntax.mode -> binding * mixfix -> term -> term list * term list -> local_theory -> local_theory) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory (*theory target primitives*) val theory_target_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val theory_target_notes: string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val theory_target_abbrev: Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory (*theory target operations*) val theory_abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val theory_declaration: declaration -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory (*locale target primitives*) val locale_target_notes: string -> string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val locale_target_abbrev: string -> Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory val locale_target_declaration: string -> bool -> declaration -> local_theory -> local_theory val locale_target_const: string -> (morphism -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory (*locale operations*) val locale_abbrev: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val locale_declaration: string -> {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory val locale_const: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val locale_dependency: string -> Locale.registration -> local_theory -> local_theory end structure Generic_Target: GENERIC_TARGET = struct (** consts **) fun export_abbrev lthy preprocess rhs = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); val rhs' = rhs |> Assumption.export_term lthy (Local_Theory.target_of lthy) |> preprocess; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) (Variable.add_fixed lthy rhs' [])); val u = fold_rev lambda term_params rhs'; val global_rhs = singleton (Variable.polymorphic thy_ctxt) u; + val type_tfrees = Term_Subst.TFrees.build (Term_Subst.add_tfreesT (Term.fastype_of u)); val extra_tfrees = - subtract (op =) (Term.add_tfreesT (Term.fastype_of u) []) (Term.add_tfrees u []); + build_rev (u |> (Term.fold_types o Term.fold_atyps) + (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I)); val type_params = map (Logic.mk_type o TFree) extra_tfrees; in (global_rhs, (extra_tfrees, (type_params, term_params))) end; fun check_mixfix ctxt (b, extra_tfrees) mx = if null extra_tfrees then mx else (if Context_Position.is_visible ctxt then warning ("Additional type variable(s) in specification of " ^ Binding.print b ^ ": " ^ commas (map (Syntax.string_of_typ ctxt o TFree) (sort_by #1 extra_tfrees)) ^ (if Mixfix.is_empty mx then "" else "\nDropping mixfix syntax " ^ Pretty.string_of (Mixfix.pretty_mixfix mx))) else (); NoSyn); fun check_mixfix_global (b, no_params) mx = if no_params orelse Mixfix.is_empty mx then mx else (warning ("Dropping global mixfix syntax: " ^ Binding.print b ^ " " ^ Pretty.string_of (Mixfix.pretty_mixfix mx)); NoSyn); fun same_const (Const (c, _), Const (c', _)) = c = c' | same_const (t $ _, t' $ _) = same_const (t, t') | same_const (_, _) = false; fun const_decl phi_pred prmode ((b, mx), rhs) phi context = if phi_pred phi then let val b' = Morphism.binding phi b; val rhs' = Morphism.term phi rhs; val same_shape = Term.aconv_untyped (rhs, rhs'); val same_stem = same_shape orelse same_const (rhs, rhs'); val const_alias = if same_shape then (case rhs' of Const (c, T) => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; in (case Type_Infer_Context.const_type ctxt c of SOME T' => if Sign.typ_equiv thy (T, T') then SOME c else NONE | NONE => NONE) end | _ => NONE) else NONE; in (case const_alias of SOME c => context |> Context.mapping (Sign.const_alias b' c) (Proof_Context.const_alias b' c) |> Morphism.form (Proof_Context.generic_notation true prmode [(rhs', mx)]) | NONE => context |> Proof_Context.generic_add_abbrev Print_Mode.internal (b', Term.close_schematic_term rhs') |-> (fn (const as Const (c, _), _) => same_stem ? (Proof_Context.generic_revert_abbrev (#1 prmode) c #> same_shape ? Morphism.form (Proof_Context.generic_notation true prmode [(const, mx)])))) end else context; (** background primitives **) structure Foundation_Interpretations = Theory_Data ( type T = (binding * (term * term list) -> Context.generic -> Context.generic) Inttab.table; val empty = Inttab.empty; val extend = I; val merge = Inttab.merge (K true); ); fun add_foundation_interpretation f = Foundation_Interpretations.map (Inttab.update_new (serial (), f)); fun foundation_interpretation binding_const_params lthy = let val interps = Foundation_Interpretations.get (Proof_Context.theory_of lthy); val interp = Inttab.fold (fn (_, f) => f binding_const_params) interps; in lthy |> Local_Theory.background_theory (Context.theory_map interp) |> Local_Theory.map_contexts (K (Context.proof_map interp)) end; fun background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy = let val params = type_params @ term_params; val target_params = type_params @ take_prefix is_Free (Variable.export_terms lthy (Local_Theory.target_of lthy) term_params); val mx' = check_mixfix_global (b, null params) mx; val (const, lthy2) = lthy |> Local_Theory.background_theory_result (Sign.declare_const lthy ((b, U), mx')); val lhs = Term.list_comb (const, params); val ((_, def), lthy3) = lthy2 |> Local_Theory.background_theory_result (Thm.add_def (Proof_Context.defs_context lthy2) false false (Thm.def_binding_optional b b_def, Logic.mk_equals (lhs, rhs))) ||> foundation_interpretation (b, (const, target_params)); in ((lhs, def), lthy3) end; fun background_declaration decl lthy = let fun theory_decl context = Local_Theory.standard_form lthy (Proof_Context.init_global (Context.theory_of context)) decl context; in Local_Theory.background_theory (Context.theory_map theory_decl) lthy end; fun background_abbrev (b, global_rhs) params = Local_Theory.background_theory_result (Sign.add_abbrev Print_Mode.internal (b, global_rhs)) #>> apply2 (fn t => Term.list_comb (Logic.unvarify_global t, params)) (** nested local theories primitives **) fun standard_facts lthy ctxt = Attrib.transform_facts (Local_Theory.standard_morphism lthy ctxt); fun standard_notes pred kind facts lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Attrib.local_notes kind (standard_facts lthy ctxt facts) ctxt |> snd else ctxt) lthy; fun standard_declaration pred decl lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Context.proof_map (Local_Theory.standard_form lthy ctxt decl) ctxt else ctxt) lthy; fun standard_const pred prmode ((b, mx), rhs) = standard_declaration pred (const_decl (K true) prmode ((b, mx), rhs)); fun standard_registration pred registration lthy = Local_Theory.map_contexts (fn level => if pred (Local_Theory.level lthy, level) then Context.proof_map (Locale.add_registration registration) else I) lthy; val local_interpretation = standard_registration (fn (n, level) => level = n - 1); (** lifting target primitives to local theory operations **) (* define *) fun define foundation internal ((b, mx), ((b_def, atts), rhs)) lthy = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); (*term and type parameters*) val ((defs, _), rhs') = Thm.cterm_of lthy rhs |> Local_Defs.export_cterm lthy thy_ctxt ||> Thm.term_of; val xs = Variable.add_fixed lthy rhs' []; val T = Term.fastype_of rhs; - val tfreesT = Term.add_tfreesT T (fold (Term.add_tfreesT o #2) xs []); - val extra_tfrees = rev (subtract (op =) tfreesT (Term.add_tfrees rhs [])); + val type_tfrees = + Term_Subst.TFrees.build (Term_Subst.add_tfreesT T #> fold (Term_Subst.add_tfreesT o #2) xs); + val extra_tfrees = + build_rev (rhs |> (Term.fold_types o Term.fold_atyps) + (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I)); val mx' = check_mixfix lthy (b, extra_tfrees) mx; val type_params = map (Logic.mk_type o TFree) extra_tfrees; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) xs); val params = type_params @ term_params; val U = map Term.fastype_of params ---> T; (*foundation*) val ((lhs', global_def), lthy2) = lthy |> foundation (((b, U), mx'), (b_def, rhs')) (type_params, term_params); (*local definition*) val ([(lhs, (_, local_def))], lthy3) = lthy2 |> Context_Position.set_visible false |> Local_Defs.define [((b, NoSyn), (Binding.empty_atts, lhs'))] ||> Context_Position.restore_visible lthy2; (*result*) val def = Thm.transitive local_def global_def |> Local_Defs.contract lthy3 defs (Thm.cterm_of lthy3 (Logic.mk_equals (lhs, rhs))); val ([(res_name, [res])], lthy4) = lthy3 |> Local_Theory.notes [((if internal then Binding.empty else b_def, atts), [([def], [])])]; in ((lhs, (res_name, res)), lthy4) end; (* notes *) local fun import_export_proof ctxt (name, raw_th) = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of ctxt); (*export assumes/defines*) val th = Goal.norm_result ctxt raw_th; val ((defs, asms), th') = Local_Defs.export ctxt thy_ctxt th; val asms' = map (rewrite_rule ctxt (Drule.norm_hhf_eqs @ defs)) asms; (*export fixes*) - val tfrees = map TFree (Thm.fold_terms Term.add_tfrees th' []); - val frees = map Free (Thm.fold_terms Term.add_frees th' []); + val tfrees = map TFree (Thm.fold_terms {hyps = true} Term.add_tfrees th' []); + val frees = map Free (Thm.fold_terms {hyps = true} Term.add_frees th' []); val (th'' :: vs) = (th' :: map (Drule.mk_term o Thm.cterm_of ctxt) (map Logic.mk_type tfrees @ frees)) |> Variable.export ctxt thy_ctxt |> Drule.zero_var_indexes_list; (*thm definition*) val result = Global_Theory.name_thm Global_Theory.official1 name th''; (*import fixes*) val (tvars, vars) = chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs) |>> map Logic.dest_type; val instT = - fold2 (fn a => fn b => (case a of TVar v => Term_Subst.TVars.add (v, b) | _ => I)) - tvars tfrees Term_Subst.TVars.empty; + Term_Subst.TVars.build (fold2 (fn a => fn b => + (case a of TVar v => Term_Subst.TVars.add (v, b) | _ => I)) tvars tfrees); val cinstT = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt T)) instT []; val cinst = map_filter (fn (Var (xi, T), t) => SOME ((xi, Term_Subst.instantiateT instT T), Thm.cterm_of ctxt (Term.map_types (Term_Subst.instantiateT instT) t)) | _ => NONE) (vars ~~ frees); val result' = Thm.instantiate (cinstT, cinst) result; (*import assumes/defines*) val result'' = (fold (curry op COMP) asms' result' handle THM _ => raise THM ("Failed to re-import result", 0, result' :: asms')) |> Local_Defs.contract ctxt defs (Thm.cprop_of th) |> Goal.norm_result ctxt |> Global_Theory.name_thm Global_Theory.unofficial2 name; in (result'', result) end; fun bind_name lthy b = (Local_Theory.full_name lthy b, Binding.default_pos_of b); fun map_facts f = map (apsnd (map (apfst (map f)))); in fun notes target_notes kind facts lthy = let val facts' = facts |> map (fn (a, bs) => (a, Global_Theory.burrow_fact (Global_Theory.name_multi (bind_name lthy (fst a))) bs)) |> map_facts (import_export_proof lthy); val local_facts = map_facts #1 facts'; val global_facts = map_facts #2 facts'; in lthy |> target_notes kind global_facts (Attrib.partial_evaluation lthy local_facts) |> Attrib.local_notes kind local_facts end; end; (* abbrev *) fun abbrev target_abbrev prmode ((b, mx), rhs) lthy = let val (global_rhs, (extra_tfrees, (type_params, term_params))) = export_abbrev lthy I rhs; val mx' = check_mixfix lthy (b, extra_tfrees) mx; in lthy |> target_abbrev prmode (b, mx') global_rhs (type_params, term_params) |> Context_Position.set_visible false |> Proof_Context.add_abbrev Print_Mode.internal (b, rhs) |> snd |> Local_Defs.fixed_abbrev ((b, NoSyn), rhs) ||> Context_Position.restore_visible lthy end; (** theory target primitives **) fun theory_target_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) = background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) #-> (fn (lhs, def) => standard_const (op <>) Syntax.mode_default ((b, mx), lhs) #> pair (lhs, def)); fun theory_target_notes kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind global_facts #> snd) #> standard_notes (op <>) kind local_facts; fun theory_target_abbrev prmode (b, mx) global_rhs params = Local_Theory.background_theory_result (Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> (fn (lhs, _) => (* FIXME type_params!? *) Sign.notation true prmode [(lhs, check_mixfix_global (b, null (snd params)) mx)] #> pair lhs)) #-> (fn lhs => standard_const (op <>) prmode ((b, if null (snd params) then NoSyn else mx), Term.list_comb (Logic.unvarify_global lhs, snd params))); (** theory operations **) val theory_abbrev = abbrev theory_target_abbrev; fun theory_declaration decl = background_declaration decl #> standard_declaration (K true) decl; fun target_registration lthy {inst, mixin, export} = {inst = inst, mixin = mixin, export = export $> Proof_Context.export_morphism lthy (Local_Theory.target_of lthy)}; fun theory_registration registration lthy = lthy |> (Local_Theory.raw_theory o Context.theory_map) (Locale.add_registration (target_registration lthy registration)) |> standard_registration (K true) registration; (** locale target primitives **) fun locale_target_notes locale kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind (Attrib.map_facts (K []) global_facts) #> snd) #> (fn lthy => lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_facts locale kind (standard_facts lthy ctxt local_facts))) #> standard_notes (fn (this, other) => other <> 0 andalso this <> other) kind local_facts; fun locale_target_declaration locale syntax decl lthy = lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_declaration locale syntax (Morphism.transform (Local_Theory.standard_morphism lthy ctxt) decl)); fun locale_target_const locale phi_pred prmode ((b, mx), rhs) = locale_target_declaration locale true (const_decl phi_pred prmode ((b, mx), rhs)) (** locale operations **) fun locale_declaration locale {syntax, pervasive} decl = pervasive ? background_declaration decl #> locale_target_declaration locale syntax decl #> standard_declaration (fn (_, other) => other <> 0) decl; fun locale_const locale prmode ((b, mx), rhs) = locale_target_const locale (K true) prmode ((b, mx), rhs) #> standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), rhs); fun locale_dependency loc registration lthy = lthy |> Local_Theory.raw_theory (Locale.add_dependency loc registration) |> standard_registration (K true) registration; (** locale abbreviations **) fun locale_target_abbrev locale prmode (b, mx) global_rhs params = background_abbrev (b, global_rhs) (snd params) #-> (fn (lhs, _) => locale_const locale prmode ((b, mx), lhs)); fun locale_abbrev locale = abbrev (locale_target_abbrev locale); end; diff --git a/src/Pure/Isar/method.ML b/src/Pure/Isar/method.ML --- a/src/Pure/Isar/method.ML +++ b/src/Pure/Isar/method.ML @@ -1,821 +1,821 @@ (* Title: Pure/Isar/method.ML Author: Markus Wenzel, TU Muenchen Isar proof methods. *) signature METHOD = sig type method = thm list -> context_tactic val CONTEXT_METHOD: (thm list -> context_tactic) -> method val METHOD: (thm list -> tactic) -> method val fail: method val succeed: method val insert_tac: Proof.context -> thm list -> int -> tactic val insert: thm list -> method val SIMPLE_METHOD: tactic -> method val SIMPLE_METHOD': (int -> tactic) -> method val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method val goal_cases_tac: string list -> context_tactic val cheating: bool -> method val intro: Proof.context -> thm list -> method val elim: Proof.context -> thm list -> method val unfold: thm list -> Proof.context -> method val fold: thm list -> Proof.context -> method val atomize: bool -> Proof.context -> method val this: Proof.context -> method val fact: thm list -> Proof.context -> method val assm_tac: Proof.context -> int -> tactic val all_assm_tac: Proof.context -> tactic val assumption: Proof.context -> method val rule_trace: bool Config.T val trace: Proof.context -> thm list -> unit val rule_tac: Proof.context -> thm list -> thm list -> int -> tactic val some_rule_tac: Proof.context -> thm list -> thm list -> int -> tactic val intros_tac: Proof.context -> thm list -> thm list -> tactic val try_intros_tac: Proof.context -> thm list -> thm list -> tactic val rule: Proof.context -> thm list -> method val erule: Proof.context -> int -> thm list -> method val drule: Proof.context -> int -> thm list -> method val frule: Proof.context -> int -> thm list -> method val set_tactic: (morphism -> thm list -> tactic) -> Context.generic -> Context.generic val clean_facts: thm list -> thm list val set_facts: thm list -> Proof.context -> Proof.context val get_facts: Proof.context -> thm list type combinator_info val no_combinator_info: combinator_info datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int datatype text = Source of Token.src | Basic of Proof.context -> method | Combinator of combinator_info * combinator * text list val map_source: (Token.src -> Token.src) -> text -> text val primitive_text: (Proof.context -> thm -> thm) -> text val succeed_text: text val standard_text: text val this_text: text val done_text: text val sorry_text: bool -> text val finish_text: text option * bool -> text val print_methods: bool -> Proof.context -> unit val check_name: Proof.context -> xstring * Position.T -> string val check_src: Proof.context -> Token.src -> Token.src val check_text: Proof.context -> text -> text val checked_text: text -> bool val method_syntax: (Proof.context -> method) context_parser -> Token.src -> Proof.context -> method val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory val local_setup: binding -> (Proof.context -> method) context_parser -> string -> local_theory -> local_theory val method_setup: bstring * Position.T -> Input.source -> string -> local_theory -> local_theory val method: Proof.context -> Token.src -> Proof.context -> method val method_closure: Proof.context -> Token.src -> Token.src val closure: bool Config.T val method_cmd: Proof.context -> Token.src -> Proof.context -> method val detect_closure_state: thm -> bool val STATIC: (unit -> unit) -> context_tactic val RUNTIME: context_tactic -> context_tactic val sleep: Time.time -> context_tactic val evaluate: text -> Proof.context -> method val evaluate_runtime: text -> Proof.context -> method type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T} val modifier: attribute -> Position.T -> modifier val old_section_parser: bool Config.T val sections: modifier parser list -> unit context_parser type text_range = text * Position.range val text: text_range option -> text option val position: text_range option -> Position.T val reports_of: text_range -> Position.report list val report: text_range -> unit val parser: int -> text_range parser val parse: text_range parser val read: Proof.context -> Token.src -> text val read_closure: Proof.context -> Token.src -> text * Token.src val read_closure_input: Proof.context -> Input.source -> text * Token.src val text_closure: text context_parser end; structure Method: METHOD = struct (** proof methods **) (* type method *) type method = thm list -> context_tactic; fun CONTEXT_METHOD tac : method = fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac) #> Seq.maps_results (tac facts); fun METHOD tac : method = fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac THEN tac facts); val fail = METHOD (K no_tac); val succeed = METHOD (K all_tac); (* insert facts *) fun insert_tac _ [] _ = all_tac | insert_tac ctxt facts i = - EVERY (map (fn r => resolve_tac ctxt [Drule.forall_intr_vars r COMP_INCR revcut_rl] i) facts); + EVERY (map (fn r => resolve_tac ctxt [Thm.forall_intr_vars r COMP_INCR revcut_rl] i) facts); fun insert thms = CONTEXT_METHOD (fn _ => fn (ctxt, st) => st |> ALLGOALS (insert_tac ctxt thms) |> TACTIC_CONTEXT ctxt); fun SIMPLE_METHOD tac = CONTEXT_METHOD (fn facts => fn (ctxt, st) => st |> (ALLGOALS (insert_tac ctxt facts) THEN tac) |> TACTIC_CONTEXT ctxt); fun SIMPLE_METHOD'' quant tac = CONTEXT_METHOD (fn facts => fn (ctxt, st) => st |> quant (insert_tac ctxt facts THEN' tac) |> TACTIC_CONTEXT ctxt); val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL; (* goals as cases *) fun goal_cases_tac case_names : context_tactic = fn (ctxt, st) => let val cases = (if null case_names then map string_of_int (1 upto Thm.nprems_of st) else case_names) |> map (rpair [] o rpair []) |> Rule_Cases.make_common ctxt (Thm.prop_of (Rule_Cases.internalize_params st)); in CONTEXT_CASES cases all_tac (ctxt, st) end; (* cheating *) fun cheating int = CONTEXT_METHOD (fn _ => fn (ctxt, st) => if int orelse Config.get ctxt quick_and_dirty then TACTIC_CONTEXT ctxt (ALLGOALS (Skip_Proof.cheat_tac ctxt) st) else error "Cheating requires quick_and_dirty mode!"); (* unfold intro/elim rules *) fun intro ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (match_tac ctxt ths)); fun elim ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt ths)); (* unfold/fold definitions *) fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.unfold_tac ctxt ths)); fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.fold_tac ctxt ths)); (* atomize rule statements *) fun atomize false ctxt = SIMPLE_METHOD' (CHANGED_PROP o Object_Logic.atomize_prems_tac ctxt) | atomize true ctxt = Context_Tactic.CONTEXT_TACTIC o K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt)); (* this -- resolve facts directly *) fun this ctxt = METHOD (EVERY o map (HEADGOAL o resolve_tac ctxt o single)); (* fact -- composition by facts from context *) fun fact [] ctxt = SIMPLE_METHOD' (Proof_Context.some_fact_tac ctxt) | fact rules ctxt = SIMPLE_METHOD' (Proof_Context.fact_tac ctxt rules); (* assumption *) local fun cond_rtac ctxt cond rule = SUBGOAL (fn (prop, i) => if cond (Logic.strip_assums_concl prop) then resolve_tac ctxt [rule] i else no_tac); in fun assm_tac ctxt = assume_tac ctxt APPEND' Goal.assume_rule_tac ctxt APPEND' cond_rtac ctxt (can Logic.dest_equals) Drule.reflexive_thm APPEND' cond_rtac ctxt (can Logic.dest_term) Drule.termI; fun all_assm_tac ctxt = let fun tac i st = if i > Thm.nprems_of st then all_tac st else ((assm_tac ctxt i THEN tac i) ORELSE tac (i + 1)) st; in tac 1 end; fun assumption ctxt = METHOD (HEADGOAL o (fn [] => assm_tac ctxt | [fact] => solve_tac ctxt [fact] | _ => K no_tac)); fun finish immed ctxt = METHOD (K ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac ctxt)); end; (* rule etc. -- single-step refinements *) val rule_trace = Attrib.setup_config_bool \<^binding>\rule_trace\ (fn _ => false); fun trace ctxt rules = if Config.get ctxt rule_trace andalso not (null rules) then Pretty.big_list "rules:" (map (Thm.pretty_thm_item ctxt) rules) |> Pretty.string_of |> tracing else (); local fun gen_rule_tac tac ctxt rules facts = (fn i => fn st => if null facts then tac ctxt rules i st else Seq.maps (fn rule => (tac ctxt o single) rule i st) (Drule.multi_resolves (SOME ctxt) facts rules)) THEN_ALL_NEW Goal.norm_hhf_tac ctxt; fun gen_arule_tac tac ctxt j rules facts = EVERY' (gen_rule_tac tac ctxt rules facts :: replicate j (assume_tac ctxt)); fun gen_some_rule_tac tac ctxt arg_rules facts = SUBGOAL (fn (goal, i) => let val rules = if not (null arg_rules) then arg_rules else flat (Context_Rules.find_rules ctxt false facts goal); in trace ctxt rules; tac ctxt rules facts i end); fun meth tac x y = METHOD (HEADGOAL o tac x y); fun meth' tac x y z = METHOD (HEADGOAL o tac x y z); in val rule_tac = gen_rule_tac resolve_tac; val rule = meth rule_tac; val some_rule_tac = gen_some_rule_tac rule_tac; val some_rule = meth some_rule_tac; val erule = meth' (gen_arule_tac eresolve_tac); val drule = meth' (gen_arule_tac dresolve_tac); val frule = meth' (gen_arule_tac forward_tac); end; (* intros_tac -- pervasive search spanned by intro rules *) fun gen_intros_tac goals ctxt intros facts = goals (insert_tac ctxt facts THEN' REPEAT_ALL_NEW (resolve_tac ctxt intros)) THEN Tactic.distinct_subgoals_tac; val intros_tac = gen_intros_tac ALLGOALS; val try_intros_tac = gen_intros_tac TRYALL; (** method syntax **) (* context data *) structure Data = Generic_Data ( type T = {methods: ((Token.src -> Proof.context -> method) * string) Name_Space.table, ml_tactic: (morphism -> thm list -> tactic) option, facts: thm list option}; val empty : T = {methods = Name_Space.empty_table Markup.methodN, ml_tactic = NONE, facts = NONE}; val extend = I; fun merge ({methods = methods1, ml_tactic = ml_tactic1, facts = facts1}, {methods = methods2, ml_tactic = ml_tactic2, facts = facts2}) : T = {methods = Name_Space.merge_tables (methods1, methods2), ml_tactic = merge_options (ml_tactic1, ml_tactic2), facts = merge_options (facts1, facts2)}; ); fun map_data f = Data.map (fn {methods, ml_tactic, facts} => let val (methods', ml_tactic', facts') = f (methods, ml_tactic, facts) in {methods = methods', ml_tactic = ml_tactic', facts = facts'} end); val get_methods = #methods o Data.get; val ops_methods = {get_data = get_methods, put_data = fn methods => map_data (fn (_, ml_tactic, facts) => (methods, ml_tactic, facts))}; (* ML tactic *) fun set_tactic ml_tactic = map_data (fn (methods, _, facts) => (methods, SOME ml_tactic, facts)); fun the_tactic context = (case #ml_tactic (Data.get context) of SOME tac => tac | NONE => raise Fail "Undefined ML tactic"); val parse_tactic = Scan.state :|-- (fn context => Scan.lift (Args.text_declaration (fn source => let val tac = context |> ML_Context.expression (Input.pos_of source) (ML_Lex.read "Context.>> (Method.set_tactic (fn morphism: Morphism.morphism => fn facts: thm list => (" @ ML_Lex.read_source source @ ML_Lex.read ")))") |> the_tactic; in fn phi => set_tactic (fn _ => Context.setmp_generic_context (SOME context) (tac phi)) end)) >> (fn decl => Morphism.form (the_tactic (Morphism.form decl context)))); (* method facts *) val clean_facts = filter_out Thm.is_dummy; fun set_facts facts = (Context.proof_map o map_data) (fn (methods, ml_tactic, _) => (methods, ml_tactic, SOME (clean_facts facts))); val get_facts_generic = these o #facts o Data.get; val get_facts = get_facts_generic o Context.Proof; val _ = Theory.setup (Global_Theory.add_thms_dynamic (Binding.make ("method_facts", \<^here>), get_facts_generic)); (* method text *) datatype combinator_info = Combinator_Info of {keywords: Position.T list}; fun combinator_info keywords = Combinator_Info {keywords = keywords}; val no_combinator_info = combinator_info []; datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int; datatype text = Source of Token.src | Basic of Proof.context -> method | Combinator of combinator_info * combinator * text list; fun map_source f (Source src) = Source (f src) | map_source _ (Basic meth) = Basic meth | map_source f (Combinator (info, comb, txts)) = Combinator (info, comb, map (map_source f) txts); fun primitive_text r = Basic (SIMPLE_METHOD o PRIMITIVE o r); val succeed_text = Basic (K succeed); val standard_text = Source (Token.make_src ("standard", Position.none) []); val this_text = Basic this; val done_text = Basic (K (SIMPLE_METHOD all_tac)); fun sorry_text int = Basic (fn _ => cheating int); fun finish_text (NONE, immed) = Basic (finish immed) | finish_text (SOME txt, immed) = Combinator (no_combinator_info, Then, [txt, Basic (finish immed)]); (* method definitions *) fun print_methods verbose ctxt = let val meths = get_methods (Context.Proof ctxt); fun prt_meth (name, (_, "")) = Pretty.mark_str name | prt_meth (name, (_, comment)) = Pretty.block (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); in [Pretty.big_list "methods:" (map prt_meth (Name_Space.markup_table verbose ctxt meths))] |> Pretty.writeln_chunks end; (* define *) fun define_global binding meth comment = Entity.define_global ops_methods binding (meth, comment); fun define binding meth comment = Entity.define ops_methods binding (meth, comment); (* check *) fun check_name ctxt = let val context = Context.Proof ctxt in #1 o Name_Space.check context (get_methods context) end; fun check_src ctxt = #1 o Token.check_src ctxt (get_methods o Context.Proof); fun check_text ctxt (Source src) = Source (check_src ctxt src) | check_text _ (Basic m) = Basic m | check_text ctxt (Combinator (x, y, body)) = Combinator (x, y, map (check_text ctxt) body); fun checked_text (Source src) = Token.checked_src src | checked_text (Basic _) = true | checked_text (Combinator (_, _, body)) = forall checked_text body; (* method setup *) fun method_syntax scan src ctxt : method = let val (m, ctxt') = Token.syntax scan src ctxt in m ctxt' end; fun setup binding scan comment = define_global binding (method_syntax scan) comment #> snd; fun local_setup binding scan comment = define binding (method_syntax scan) comment #> snd; fun method_setup binding source comment = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Method.local_setup (" ^ ML_Syntax.make_binding binding ^ ") (") @ ML_Lex.read_source source @ ML_Lex.read (")" ^ ML_Syntax.print_string comment ^ ")")) |> Context.proof_map; (* prepare methods *) fun method ctxt = let val table = get_methods (Context.Proof ctxt) in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end; fun method_closure ctxt src = let val src' = map Token.init_assignable src; val ctxt' = Context_Position.not_really ctxt; val _ = Seq.pull (method ctxt' src' ctxt' [] (ctxt', Goal.protect 0 Drule.dummy_thm)); in map Token.closure src' end; val closure = Config.declare_bool ("Method.closure", \<^here>) (K true); fun method_cmd ctxt = check_src ctxt #> Config.get ctxt closure ? method_closure ctxt #> method ctxt; (* static vs. runtime state *) fun detect_closure_state st = (case try Logic.dest_term (Thm.concl_of (perhaps (try Goal.conclude) st)) of NONE => false | SOME t => Term.is_dummy_pattern t); fun STATIC test : context_tactic = fn (ctxt, st) => if detect_closure_state st then (test (); Seq.single (Seq.Result (ctxt, st))) else Seq.empty; fun RUNTIME (tac: context_tactic) (ctxt, st) = if detect_closure_state st then Seq.empty else tac (ctxt, st); fun sleep t = RUNTIME (fn ctxt_st => (OS.Process.sleep t; Seq.single (Seq.Result ctxt_st))); (* evaluate method text *) local val op THEN = Seq.THEN; fun BYPASS_CONTEXT (tac: tactic) = fn result => (case result of Seq.Error _ => Seq.single result | Seq.Result (ctxt, st) => tac st |> TACTIC_CONTEXT ctxt); val preparation = BYPASS_CONTEXT (ALLGOALS Goal.conjunction_tac); fun RESTRICT_GOAL i n method = BYPASS_CONTEXT (PRIMITIVE (Goal.restrict i n)) THEN method THEN BYPASS_CONTEXT (PRIMITIVE (Goal.unrestrict i)); fun SELECT_GOAL method i = RESTRICT_GOAL i 1 method; fun (method1 THEN_ALL_NEW method2) i (result : context_state Seq.result) = (case result of Seq.Error _ => Seq.single result | Seq.Result (_, st) => result |> method1 i |> Seq.maps (fn result' => (case result' of Seq.Error _ => Seq.single result' | Seq.Result (_, st') => result' |> Seq.INTERVAL method2 i (i + Thm.nprems_of st' - Thm.nprems_of st)))) fun COMBINATOR1 comb [meth] = comb meth | COMBINATOR1 _ _ = raise Fail "Method combinator requires exactly one argument"; fun combinator Then = Seq.EVERY | combinator Then_All_New = (fn [] => Seq.single | methods => preparation THEN (foldl1 (op THEN_ALL_NEW) (map SELECT_GOAL methods) 1)) | combinator Orelse = Seq.FIRST | combinator Try = COMBINATOR1 Seq.TRY | combinator Repeat1 = COMBINATOR1 Seq.REPEAT1 | combinator (Select_Goals n) = COMBINATOR1 (fn method => preparation THEN RESTRICT_GOAL 1 n method); in fun evaluate text ctxt0 facts = let val ctxt = set_facts facts ctxt0; fun eval0 m = Seq.single #> Seq.maps_results (m facts); fun eval (Basic m) = eval0 (m ctxt) | eval (Source src) = eval0 (method_cmd ctxt src ctxt) | eval (Combinator (_, c, txts)) = combinator c (map eval txts); in eval text o Seq.Result end; end; fun evaluate_runtime text ctxt = let val text' = text |> (map_source o map o Token.map_facts) (fn SOME name => (case Proof_Context.lookup_fact ctxt name of SOME {dynamic = true, thms} => K thms | _ => I) | NONE => I); val ctxt' = Config.put closure false ctxt; in fn facts => RUNTIME (fn st => evaluate text' ctxt' facts st) end; (** concrete syntax **) (* type modifier *) type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T}; fun modifier attribute pos : modifier = {init = I, attribute = attribute, pos = pos}; (* sections *) val old_section_parser = Config.declare_bool ("Method.old_section_parser", \<^here>) (K false); local fun thms ss = Scan.repeats (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm); fun app {init, attribute, pos = _} ths context = fold_map (Thm.apply_attribute attribute) ths (Context.map_proof init context); fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|-- (fn (m, ths) => Scan.succeed (swap (app m ths context)))); in fun old_sections ss = Scan.repeat (section ss) >> K (); end; local fun sect (modifier : modifier parser) = Scan.depend (fn context => Scan.ahead Parse.not_eof -- Scan.trace modifier -- Scan.repeat (Scan.unless modifier Parse.thm) >> (fn ((tok0, ({init, attribute, pos}, modifier_toks)), xthms) => let val decl = (case Token.get_value tok0 of SOME (Token.Declaration decl) => decl | _ => let val ctxt = Context.proof_of context; val prep_att = Attrib.check_src ctxt #> map (Token.assign NONE); val thms = map (fn (a, bs) => (Proof_Context.get_fact ctxt a, map prep_att bs)) xthms; val facts = Attrib.partial_evaluation ctxt [((Binding.name "dummy", []), thms)] |> map (fn (_, bs) => ((Binding.empty, [Attrib.internal (K attribute)]), bs)); fun decl phi = Context.mapping I init #> Attrib.generic_notes "" (Attrib.transform_facts phi facts) #> snd; val modifier_report = (#1 (Token.range_of modifier_toks), Position.entity_markup Markup.method_modifierN ("", pos)); val _ = Context_Position.reports ctxt (modifier_report :: Token.reports_of_value tok0); val _ = Token.assign (SOME (Token.Declaration decl)) tok0; in decl end); in (Morphism.form decl context, decl) end)); in fun sections ss = Args.context :|-- (fn ctxt => if Config.get ctxt old_section_parser then old_sections ss else Scan.repeat (sect (Scan.first ss)) >> K ()); end; (* extra rule methods *) fun xrule_meth meth = Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >> (fn (n, ths) => fn ctxt => meth ctxt n ths); (* text range *) type text_range = text * Position.range; fun text NONE = NONE | text (SOME (txt, _)) = SOME txt; fun position NONE = Position.none | position (SOME (_, (pos, _))) = pos; (* reports *) local fun keyword_positions (Source _) = [] | keyword_positions (Basic _) = [] | keyword_positions (Combinator (Combinator_Info {keywords}, _, texts)) = keywords @ maps keyword_positions texts; in fun reports_of ((text, (pos, _)): text_range) = (pos, Markup.language_method) :: maps (fn p => map (pair p) (Markup.keyword3 :: Completion.suppress_abbrevs "")) (keyword_positions text); fun report text_range = if Context_Position.pide_reports () then Position.reports (reports_of text_range) else (); end; (* parser *) local fun is_symid_meth s = s <> "|" andalso s <> "?" andalso s <> "+" andalso Token.ident_or_symbolic s; in fun parser pri = let val meth_name = Parse.token Parse.name; fun meth5 x = (meth_name >> (Source o single) || Scan.ahead Parse.cartouche |-- Parse.not_eof >> (fn tok => Source (Token.make_src ("cartouche", Position.none) [tok])) || Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x and meth4 x = (meth5 -- Parse.position (Parse.$$$ "?") >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Try, [m])) || meth5 -- Parse.position (Parse.$$$ "+") >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Repeat1, [m])) || meth5 -- (Parse.position (Parse.$$$ "[") -- Scan.optional Parse.nat 1 -- Parse.position (Parse.$$$ "]")) >> (fn (m, (((_, pos1), n), (_, pos2))) => Combinator (combinator_info [pos1, pos2], Select_Goals n, [m])) || meth5) x and meth3 x = (meth_name ::: Parse.args1 is_symid_meth >> Source || meth4) x and meth2 x = (Parse.enum1_positions "," meth3 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then, ms))) x and meth1 x = (Parse.enum1_positions ";" meth2 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then_All_New, ms))) x and meth0 x = (Parse.enum1_positions "|" meth1 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Orelse, ms))) x; val meth = nth [meth0, meth1, meth2, meth3, meth4, meth5] pri handle General.Subscript => raise Fail ("Bad method parser priority " ^ string_of_int pri); in Scan.trace meth >> (fn (m, toks) => (m, Token.range_of toks)) end; val parse = parser 4; end; (* read method text *) fun read ctxt src = (case Scan.read Token.stopper (Parse.!!! (parser 0 --| Scan.ahead Parse.eof)) src of SOME (text, range) => if checked_text text then text else (report (text, range); check_text ctxt text) | NONE => error ("Failed to parse method" ^ Position.here (#1 (Token.range_of src)))); fun read_closure ctxt src0 = let val src1 = map Token.init_assignable src0; val text = read ctxt src1 |> map_source (method_closure ctxt); val src2 = map Token.closure src1; in (text, src2) end; fun read_closure_input ctxt input = let val syms = Input.source_explode input in (case Token.read_body (Thy_Header.get_keywords' ctxt) (Scan.many Token.not_eof) syms of SOME res => read_closure ctxt res | NONE => error ("Bad input" ^ Position.here (Input.pos_of input))) end; val text_closure = Args.context -- Scan.lift (Parse.token Parse.text) >> (fn (ctxt, tok) => (case Token.get_value tok of SOME (Token.Source src) => read ctxt src | _ => let val (text, src) = read_closure_input ctxt (Token.input_of tok); val _ = Token.assign (SOME (Token.Source src)) tok; in text end)); (* theory setup *) val _ = Theory.setup (setup \<^binding>\fail\ (Scan.succeed (K fail)) "force failure" #> setup \<^binding>\succeed\ (Scan.succeed (K succeed)) "succeed" #> setup \<^binding>\sleep\ (Scan.lift Parse.real >> (fn s => fn _ => fn _ => sleep (seconds s))) "succeed after delay (in seconds)" #> setup \<^binding>\-\ (Scan.succeed (K (SIMPLE_METHOD all_tac))) "insert current facts, nothing else" #> setup \<^binding>\goal_cases\ (Scan.lift (Scan.repeat Args.name_token) >> (fn names => fn _ => CONTEXT_METHOD (fn _ => fn (ctxt, st) => (case drop (Thm.nprems_of st) names of [] => NONE | bad => if detect_closure_state st then NONE else SOME (fn () => ("Excessive case name(s): " ^ commas_quote (map Token.content_of bad) ^ Position.here (#1 (Token.range_of bad))))) |> (fn SOME msg => Seq.single (Seq.Error msg) | NONE => goal_cases_tac (map Token.content_of names) (ctxt, st))))) "bind cases for goals" #> setup \<^binding>\subproofs\ (text_closure >> (Context_Tactic.SUBPROOFS ooo evaluate_runtime)) "apply proof method to subproofs with closed derivation" #> setup \<^binding>\insert\ (Attrib.thms >> (K o insert)) "insert theorems, ignoring facts" #> setup \<^binding>\intro\ (Attrib.thms >> (fn ths => fn ctxt => intro ctxt ths)) "repeatedly apply introduction rules" #> setup \<^binding>\elim\ (Attrib.thms >> (fn ths => fn ctxt => elim ctxt ths)) "repeatedly apply elimination rules" #> setup \<^binding>\unfold\ (Attrib.thms >> unfold_meth) "unfold definitions" #> setup \<^binding>\fold\ (Attrib.thms >> fold_meth) "fold definitions" #> setup \<^binding>\atomize\ (Scan.lift (Args.mode "full") >> atomize) "present local premises as object-level statements" #> setup \<^binding>\rule\ (Attrib.thms >> (fn ths => fn ctxt => some_rule ctxt ths)) "apply some intro/elim rule" #> setup \<^binding>\erule\ (xrule_meth erule) "apply rule in elimination manner (improper)" #> setup \<^binding>\drule\ (xrule_meth drule) "apply rule in destruct manner (improper)" #> setup \<^binding>\frule\ (xrule_meth frule) "apply rule in forward manner (improper)" #> setup \<^binding>\this\ (Scan.succeed this) "apply current facts as rules" #> setup \<^binding>\fact\ (Attrib.thms >> fact) "composition by facts from context" #> setup \<^binding>\assumption\ (Scan.succeed assumption) "proof by assumption, preferring facts" #> setup \<^binding>\rename_tac\ (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >> (fn (quant, xs) => K (SIMPLE_METHOD'' quant (rename_tac xs)))) "rename parameters of goal" #> setup \<^binding>\rotate_tac\ (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >> (fn (quant, i) => K (SIMPLE_METHOD'' quant (rotate_tac i)))) "rotate assumptions of goal" #> setup \<^binding>\tactic\ (parse_tactic >> (K o METHOD)) "ML tactic as proof method" #> setup \<^binding>\raw_tactic\ (parse_tactic >> (fn tac => fn _ => Context_Tactic.CONTEXT_TACTIC o tac)) "ML tactic as raw proof method" #> setup \<^binding>\use\ (Attrib.thms -- (Scan.lift (Parse.$$$ "in") |-- text_closure) >> (fn (thms, text) => fn ctxt => fn _ => evaluate_runtime text ctxt thms)) "indicate method facts and context for method expression"); (*final declarations of this structure!*) val unfold = unfold_meth; val fold = fold_meth; end; val CONTEXT_METHOD = Method.CONTEXT_METHOD; val METHOD = Method.METHOD; val SIMPLE_METHOD = Method.SIMPLE_METHOD; val SIMPLE_METHOD' = Method.SIMPLE_METHOD'; val SIMPLE_METHOD'' = Method.SIMPLE_METHOD''; diff --git a/src/Pure/Isar/obtain.ML b/src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML +++ b/src/Pure/Isar/obtain.ML @@ -1,419 +1,427 @@ (* Title: Pure/Isar/obtain.ML Author: Markus Wenzel, TU Muenchen Generalized existence and cases rules within Isar proof text. *) signature OBTAIN = sig val obtain_thesis: Proof.context -> ((string * typ) * term) * Proof.context val obtains_attributes: ('typ, 'term) Element.obtain list -> attribute list val obtains_attribs: ('typ, 'term) Element.obtain list -> Token.src list val read_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list val cert_obtains: Proof.context -> term -> Element.obtains_i -> (binding * term) list val parse_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list val consider: Element.obtains_i -> bool -> Proof.state -> Proof.state val consider_cmd: Element.obtains -> bool -> Proof.state -> Proof.state val obtain: binding -> (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> (term * term list) list list -> (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state val obtain_cmd: binding -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> (string * string list) list list -> (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state val result: (Proof.context -> tactic) -> thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context val guess: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state val guess_cmd: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state end; structure Obtain: OBTAIN = struct (** specification elements **) (* obtain_export *) (* [x, A x] : B -------- B *) fun eliminate_term ctxt xs tm = let val vs = map (dest_Free o Thm.term_of) xs; val bads = Term.fold_aterms (fn t as Free v => if member (op =) vs v then insert (op aconv) t else I | _ => I) tm []; val _ = null bads orelse error ("Result contains obtained parameters: " ^ space_implode " " (map (Syntax.string_of_term ctxt) bads)); in tm end; fun eliminate ctxt rule xs As thm = let val _ = eliminate_term ctxt xs (Thm.full_prop_of thm); val _ = Object_Logic.is_judgment ctxt (Thm.concl_of thm) orelse error "Conclusion in obtained context must be object-logic judgment"; val ((_, [thm']), ctxt') = Variable.import true [thm] ctxt; val prems = Drule.strip_imp_prems (Thm.cprop_of thm'); in ((Drule.implies_elim_list thm' (map Thm.assume prems) |> Drule.implies_intr_list (map (Drule.norm_hhf_cterm ctxt') As) |> Drule.forall_intr_list xs) COMP rule) |> Drule.implies_intr_list prems |> singleton (Variable.export ctxt' ctxt) end; fun obtain_export ctxt rule xs _ As = (eliminate ctxt rule xs As, eliminate_term ctxt xs); (* result declaration *) fun case_names (obtains: ('typ, 'term) Element.obtain list) = obtains |> map_index (fn (i, (b, _)) => if Binding.is_empty b then string_of_int (i + 1) else Name_Space.base_name b); fun obtains_attributes obtains = [Rule_Cases.consumes (~ (length obtains)), Rule_Cases.case_names (case_names obtains)]; fun obtains_attribs obtains = [Attrib.consumes (~ (length obtains)), Attrib.case_names (case_names obtains)]; (* obtain thesis *) fun obtain_thesis ctxt = let val ([x], ctxt') = Proof_Context.add_fixes [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] ctxt; val t = Object_Logic.fixed_judgment ctxt x; val v = dest_Free (Object_Logic.drop_judgment ctxt t); in ((v, t), ctxt') end; (* obtain clauses *) local val mk_all_external = Logic.all_constraint o Variable.default_type; fun mk_all_internal ctxt (y, z) t = let + val frees = + Term_Subst.Frees.build (t |> Term.fold_aterms + (fn Free v => Term_Subst.Frees.add (v, ()) | _ => I)); val T = - (case AList.lookup (op =) (Term.add_frees t []) z of + (case Term_Subst.Frees.get_first (fn ((x, T), ()) => if x = z then SOME T else NONE) frees of SOME T => T | NONE => the_default dummyT (Variable.default_type ctxt z)); in Logic.all_const T $ Term.lambda_name (y, Free (z, T)) t end; fun prepare_clause prep_var parse_prop mk_all ctxt thesis raw_vars raw_props = let val ((xs', vars), ctxt') = ctxt |> fold_map prep_var raw_vars |-> (fn vars => Proof_Context.add_fixes vars ##>> pair vars); val xs = map (Variable.check_name o #1) vars; in Logic.list_implies (map (parse_prop ctxt') raw_props, thesis) |> fold_rev (mk_all ctxt') (xs ~~ xs') end; fun prepare_obtains prep_clause check_terms ctxt thesis (raw_obtains: ('typ, 'term) Element.obtain list) = let val clauses = raw_obtains |> map (fn (_, (raw_vars, raw_props)) => prep_clause ctxt thesis raw_vars raw_props) |> check_terms ctxt; in map fst raw_obtains ~~ clauses end; val parse_clause = prepare_clause Proof_Context.read_var Syntax.parse_prop mk_all_external; val cert_clause = prepare_clause Proof_Context.cert_var (K I) mk_all_internal; in val read_obtains = prepare_obtains parse_clause Syntax.check_terms; val cert_obtains = prepare_obtains cert_clause (K I); val parse_obtains = prepare_obtains parse_clause (K I); end; (** consider: generalized elimination and cases rule **) (* consider (a) x where "A x" | (b) y where "B y" | ... \ have thesis if a [intro?]: "\x. A x \ thesis" and b [intro?]: "\y. B y \ thesis" and ... for thesis apply (insert that) *) local fun gen_consider prep_obtains raw_obtains int state = let val _ = Proof.assert_forward_or_chain state; val ctxt = Proof.context_of state; val ((_, thesis), thesis_ctxt) = obtain_thesis ctxt; val obtains = prep_obtains thesis_ctxt thesis raw_obtains; val atts = Rule_Cases.cases_open :: obtains_attributes raw_obtains; in state |> Proof.have true NONE (K I) [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] (map (fn (a, A) => ((a, [Context_Rules.intro_query NONE]), [(A, [])])) obtains) [((Binding.empty, atts), [(thesis, [])])] int |-> Proof.refine_insert end; in val consider = gen_consider cert_obtains; val consider_cmd = gen_consider read_obtains; end; (** obtain: augmented context based on generalized existence rule **) (* obtain (a) x where "A x" \ have thesis if a [intro?]: "\x. A x \ thesis" for thesis apply (insert that) fix x assm <> "A x" *) local fun gen_obtain prep_stmt prep_att that_binding raw_decls raw_fixes raw_prems raw_concls int state = let val _ = Proof.assert_forward_or_chain state; val ((_, thesis), thesis_ctxt) = obtain_thesis (Proof.context_of state); val ({vars, propss, binds, result_binds, ...}, params_ctxt) = prep_stmt (raw_decls @ raw_fixes) (raw_prems @ map #2 raw_concls) thesis_ctxt; val (decls, fixes) = chop (length raw_decls) vars ||> map #2; val (premss, conclss) = chop (length raw_prems) propss; val propss' = (map o map) (Logic.close_prop fixes (flat premss)) conclss; val that_prop = Logic.list_rename_params (map (#1 o #2) decls) (fold_rev (Logic.all o #2 o #2) decls (Logic.list_implies (flat propss', thesis))); val cparams = map (Thm.cterm_of params_ctxt o #2 o #2) decls; val asms = map (fn ((b, raw_atts), _) => (b, map (prep_att params_ctxt) raw_atts)) raw_concls ~~ map (map (rpair [])) propss'; fun after_qed (result_ctxt, results) state' = let val [rule] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in state' |> Proof.fix (map #1 decls) |> Proof.map_context (fold (Variable.bind_term o apsnd (Logic.close_term fixes)) binds) |> Proof.assm (obtain_export params_ctxt rule cparams) [] [] asms end; in state |> Proof.have true NONE after_qed [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] [((that_binding, [Context_Rules.intro_query NONE]), [(that_prop, [])])] [(Binding.empty_atts, [(thesis, [])])] int |-> Proof.refine_insert |> Proof.map_context (fold Variable.bind_term result_binds) end; in val obtain = gen_obtain Proof_Context.cert_stmt (K I); val obtain_cmd = gen_obtain Proof_Context.read_stmt Attrib.attribute_cmd; end; (** tactical result **) fun check_result ctxt thesis th = (case Thm.prems_of th of [prem] => if Thm.concl_of th aconv thesis andalso Logic.strip_assums_concl prem aconv thesis then th else error ("Guessed a different clause:\n" ^ Thm.string_of_thm ctxt th) | [] => error "Goal solved -- nothing guessed" | _ => error ("Guess split into several cases:\n" ^ Thm.string_of_thm ctxt th)); fun result tac facts ctxt = let val ((thesis_var, thesis), thesis_ctxt) = obtain_thesis ctxt; val st = Goal.init (Thm.cterm_of ctxt thesis); val rule = (case SINGLE (Method.insert_tac thesis_ctxt facts 1 THEN tac thesis_ctxt) st of NONE => raise THM ("Obtain.result: tactic failed", 0, facts) | SOME th => check_result thesis_ctxt thesis (Raw_Simplifier.norm_hhf thesis_ctxt (Goal.conclude th))); val closed_rule = Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) rule; val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt; val obtain_rule = Thm.forall_elim (Thm.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule'; val ((params, stmt), fix_ctxt) = Variable.focus_cterm NONE (Thm.cprem_of obtain_rule 1) ctxt'; val (prems, ctxt'') = Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params)) (Drule.strip_imp_prems stmt) fix_ctxt; in ((params, prems), ctxt'') end; (** guess: obtain based on tactical result **) (* guess x \ { fix thesis have "PROP ?guess" apply magic \ \turn goal into \thesis \ #thesis\\ apply_end magic \ \turn final \(\x. P x \ thesis) \ #thesis\ into\ \ \\#((\x. A x \ thesis) \ thesis)\ which is a finished goal state\ } fix x assm <> "A x" *) local fun unify_params vars thesis_var raw_rule ctxt = let val thy = Proof_Context.theory_of ctxt; val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt); fun err msg th = error (msg ^ ":\n" ^ Thm.string_of_thm ctxt th); val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1; val rule = Thm.incr_indexes (maxidx + 1) raw_rule; val params = Rule_Cases.strip_params (Logic.nth_prem (1, Thm.prop_of rule)); val m = length vars; val n = length params; val _ = m <= n orelse err "More variables than parameters in obtained rule" rule; fun unify ((x, T), (y, U)) (tyenv, max) = Sign.typ_unify thy (T, U) (tyenv, max) handle Type.TUNIFY => err ("Failed to unify variable " ^ string_of_term (Free (x, Envir.norm_type tyenv T)) ^ " against parameter " ^ string_of_term (Syntax_Trans.mark_bound_abs (y, Envir.norm_type tyenv U)) ^ " in") rule; val (tyenv, _) = fold unify (map #1 vars ~~ take m params) (Vartab.empty, Int.max (maxidx, Thm.maxidx_of rule)); val norm_type = Envir.norm_type tyenv; val xs = map (apsnd norm_type o fst) vars; val ys = map (apsnd norm_type) (drop m params); val ys' = map Name.internal (Name.variant_list (map fst xs) (map fst ys)) ~~ map #2 ys; val terms = map (Drule.mk_term o Thm.cterm_of ctxt o Free) (xs @ ys'); val instT = - fold (Term.add_tvarsT o #2) params [] - |> map (fn v => (v, Thm.ctyp_of ctxt (norm_type (TVar v)))); + Term_Subst.TVars.build + (params |> fold (#2 #> Term.fold_atyps (fn T => fn instT => + (case T of + TVar v => + if Term_Subst.TVars.defined instT v then instT + else Term_Subst.TVars.update (v, Thm.ctyp_of ctxt (norm_type T)) instT + | _ => instT)))); val closed_rule = rule |> Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) - |> Thm.instantiate (instT, []); + |> Thm.instantiate (Term_Subst.TVars.dest instT, []); val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt; val vars' = map (dest_Free o Thm.term_of o Drule.dest_term) terms' ~~ (map snd vars @ replicate (length ys) NoSyn); val rule'' = Thm.forall_elim (Thm.cterm_of ctxt' (Logic.varify_global (Free thesis_var))) rule'; in ((vars', rule''), ctxt') end; fun inferred_type (binding, _, mx) ctxt = let val x = Variable.check_name binding; val ((_, T), ctxt') = Proof_Context.inferred_param x ctxt in ((x, T, mx), ctxt') end; fun polymorphic ctxt vars = let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars)) in map2 (fn (x, _, mx) => fn T => ((x, T), mx)) vars Ts end; fun gen_guess prep_var raw_vars int state = let val _ = Proof.assert_forward_or_chain state; val ctxt = Proof.context_of state; val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else []; val (thesis_var, thesis) = #1 (obtain_thesis ctxt); val vars = ctxt |> fold_map prep_var raw_vars |-> fold_map inferred_type |> fst |> polymorphic ctxt; fun guess_context raw_rule state' = let val ((parms, rule), ctxt') = unify_params vars thesis_var raw_rule (Proof.context_of state'); val (xs, _) = Variable.add_fixes (map (#1 o #1) parms) ctxt'; val ps = xs ~~ map (#2 o #1) parms; val ts = map Free ps; val asms = Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule)) |> map (fn asm => (Term.betapplys (fold_rev Term.abs ps asm, ts), [])); val _ = not (null asms) orelse error "Trivial result -- nothing guessed"; in state' |> Proof.map_context (K ctxt') |> Proof.fix (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms) |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm (obtain_export fix_ctxt rule (map (Thm.cterm_of ctxt) ts)) [] [] [(Binding.empty_atts, asms)]) |> Proof.map_context (fold Variable.unbind_term Auto_Bind.no_facts) end; val goal = Var (("guess", 0), propT); val pos = Position.thread_data (); fun print_result ctxt' (k, [(s, [_, th])]) = Proof_Display.print_results int pos ctxt' (k, [(s, [th])]); val before_qed = Method.primitive_text (fn ctxt => Goal.conclude #> Raw_Simplifier.norm_hhf ctxt #> (fn th => Goal.protect 0 (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th))); fun after_qed (result_ctxt, results) state' = let val [_, res] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in state' |> Proof.end_block |> guess_context (check_result ctxt thesis res) end; in state |> Proof.enter_forward |> Proof.begin_block |> Proof.fix [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] |> Proof.chain_facts chain_facts |> Proof.internal_goal print_result Proof_Context.mode_schematic true "guess" (SOME before_qed) after_qed [] [] [(Binding.empty_atts, [(Logic.mk_term goal, []), (goal, [])])] |> snd |> Proof.refine_singleton (Method.primitive_text (fn _ => fn _ => Goal.init (Thm.cterm_of ctxt thesis))) end; in val guess = gen_guess Proof_Context.cert_var; val guess_cmd = gen_guess Proof_Context.read_var; end; end; diff --git a/src/Pure/Isar/overloading.ML b/src/Pure/Isar/overloading.ML --- a/src/Pure/Isar/overloading.ML +++ b/src/Pure/Isar/overloading.ML @@ -1,226 +1,226 @@ (* Title: Pure/Isar/overloading.ML Author: Florian Haftmann, TU Muenchen Overloaded definitions without any discipline. *) signature OVERLOADING = sig type improvable_syntax val activate_improvable_syntax: Proof.context -> Proof.context val map_improvable_syntax: (improvable_syntax -> improvable_syntax) -> Proof.context -> Proof.context val set_primary_constraints: Proof.context -> Proof.context val show_reverted_improvements: bool Config.T; val overloading: (string * (string * typ) * bool) list -> theory -> local_theory val overloading_cmd: (string * string * bool) list -> theory -> local_theory val theory_map: (string * (string * typ) * bool) list -> (local_theory -> local_theory) -> theory -> theory val theory_map_result: (string * (string * typ) * bool) list -> (morphism -> 'a -> 'b) -> (local_theory -> 'a * local_theory) -> theory -> 'b * theory end; structure Overloading: OVERLOADING = struct (* generic check/uncheck combinators for improvable constants *) type improvable_syntax = { primary_constraints: (string * typ) list, secondary_constraints: (string * typ) list, improve: string * typ -> (typ * typ) option, subst: string * typ -> (typ * term) option, no_subst_in_abbrev_mode: bool, unchecks: (term * term) list } structure Improvable_Syntax = Proof_Data ( type T = {syntax: improvable_syntax, secondary_pass: bool} fun init _ = {syntax = { primary_constraints = [], secondary_constraints = [], improve = K NONE, subst = K NONE, no_subst_in_abbrev_mode = false, unchecks = [] }, secondary_pass = false}: T; ); fun map_improvable_syntax f = Improvable_Syntax.map (fn {syntax, secondary_pass} => {syntax = f syntax, secondary_pass = secondary_pass}); val mark_passed = Improvable_Syntax.map (fn {syntax, secondary_pass} => {syntax = syntax, secondary_pass = true}); fun improve_term_check ts ctxt = let val thy = Proof_Context.theory_of ctxt; val {syntax = {secondary_constraints, improve, subst, no_subst_in_abbrev_mode, ...}, secondary_pass} = Improvable_Syntax.get ctxt; val no_subst = Proof_Context.abbrev_mode ctxt andalso no_subst_in_abbrev_mode; fun accumulate_improvements (Const (c, ty)) = (case improve (c, ty) of SOME ty_ty' => Sign.typ_match thy ty_ty' | _ => I) | accumulate_improvements _ = I; fun apply_subst t = Envir.expand_term (fn Const (c, ty) => (case subst (c, ty) of SOME (ty', t') => if Sign.typ_instance thy (ty, ty') then SOME (ty', apply_subst t') else NONE | NONE => NONE) | _ => NONE) t; - val improvements = (fold o fold_aterms) accumulate_improvements ts Vartab.empty; + val improvements = Vartab.build ((fold o fold_aterms) accumulate_improvements ts); val ts' = ts |> (map o map_types) (Envir.subst_type improvements) |> not no_subst ? map apply_subst; in if secondary_pass orelse no_subst then if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) else SOME (ts', ctxt |> fold (Proof_Context.add_const_constraint o apsnd SOME) secondary_constraints |> mark_passed) end; fun rewrite_liberal thy unchecks t = (case try (Pattern.rewrite_term_top thy unchecks []) t of NONE => NONE | SOME t' => if t aconv t' then NONE else SOME t'); val show_reverted_improvements = Attrib.setup_config_bool \<^binding>\show_reverted_improvements\ (K true); fun improve_term_uncheck ts ctxt = let val thy = Proof_Context.theory_of ctxt; val {syntax = {unchecks, ...}, ...} = Improvable_Syntax.get ctxt; val revert = Config.get ctxt show_reverted_improvements; val ts' = map (rewrite_liberal thy unchecks) ts; in if revert andalso exists is_some ts' then SOME (map2 the_default ts ts', ctxt) else NONE end; fun set_primary_constraints ctxt = let val {syntax = {primary_constraints, ...}, ...} = Improvable_Syntax.get ctxt; in fold (Proof_Context.add_const_constraint o apsnd SOME) primary_constraints ctxt end; val activate_improvable_syntax = Context.proof_map (Syntax_Phases.term_check' 0 "improvement" improve_term_check #> Syntax_Phases.term_uncheck' 0 "improvement" improve_term_uncheck) #> set_primary_constraints; (* overloading target *) structure Data = Proof_Data ( type T = ((string * typ) * (string * bool)) list; fun init _ = []; ); val get_overloading = Data.get o Local_Theory.target_of; val map_overloading = Local_Theory.target o Data.map; fun operation lthy b = get_overloading lthy |> get_first (fn ((c, _), (v, checked)) => if Binding.name_of b = v then SOME (c, (v, checked)) else NONE); fun synchronize_syntax ctxt = let val overloading = Data.get ctxt; fun subst (c, ty) = (case AList.lookup (op =) overloading (c, ty) of SOME (v, _) => SOME (ty, Free (v, ty)) | NONE => NONE); val unchecks = map (fn (c_ty as (_, ty), (v, _)) => (Free (v, ty), Const c_ty)) overloading; in ctxt |> map_improvable_syntax (K {primary_constraints = [], secondary_constraints = [], improve = K NONE, subst = subst, no_subst_in_abbrev_mode = false, unchecks = unchecks}) end; fun define_overloaded (c, U) (v, checked) (b_def, rhs) = Local_Theory.background_theory_result (Thm.add_def_global (not checked) true (Thm.def_binding_optional (Binding.name v) b_def, Logic.mk_equals (Const (c, Term.fastype_of rhs), rhs))) ##> map_overloading (filter_out (fn (_, (v', _)) => v' = v)) ##> Local_Theory.map_contexts (K synchronize_syntax) #-> (fn (_, def) => pair (Const (c, U), def)); fun foundation (((b, U), mx), (b_def, rhs)) params lthy = (case operation lthy b of SOME (c, (v, checked)) => if Mixfix.is_empty mx then lthy |> define_overloaded (c, U) (v, checked) (b_def, rhs) else error ("Illegal mixfix syntax for overloaded constant " ^ quote c) | NONE => lthy |> Generic_Target.theory_target_foundation (((b, U), mx), (b_def, rhs)) params); fun pretty lthy = let val overloading = get_overloading lthy; fun pr_operation ((c, ty), (v, _)) = Pretty.block (Pretty.breaks [Pretty.str v, Pretty.str "\", Proof_Context.pretty_const lthy c, Pretty.str "::", Syntax.pretty_typ lthy ty]); in [Pretty.block (Pretty.fbreaks (Pretty.keyword1 "overloading" :: map pr_operation overloading))] end; fun conclude lthy = let val overloading = get_overloading lthy; val _ = if null overloading then () else error ("Missing definition(s) for parameter(s) " ^ commas_quote (map (Syntax.string_of_term lthy o Const o fst) overloading)); in lthy end; fun gen_overloading prep_const raw_overloading_spec thy = let val ctxt = Proof_Context.init_global thy; val _ = if null raw_overloading_spec then error "At least one parameter must be given" else (); val overloading_spec = raw_overloading_spec |> map (fn (v, const, checked) => (Term.dest_Const (prep_const ctxt const), (v, checked))); in thy |> Local_Theory.init {background_naming = Sign.naming_of thy, setup = Proof_Context.init_global #> Data.put overloading_spec #> fold (fn ((_, ty), (v, _)) => Variable.declare_names (Free (v, ty))) overloading_spec #> activate_improvable_syntax #> synchronize_syntax, conclude = conclude} {define = Generic_Target.define foundation, notes = Generic_Target.notes Generic_Target.theory_target_notes, abbrev = Generic_Target.abbrev Generic_Target.theory_target_abbrev, declaration = K Generic_Target.theory_declaration, theory_registration = Generic_Target.theory_registration, locale_dependency = fn _ => error "Not possible in overloading target", pretty = pretty} end; val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const); val overloading_cmd = gen_overloading Syntax.read_term; fun theory_map overloading_spec g = overloading overloading_spec #> g #> Local_Theory.exit_global; fun theory_map_result overloading_spec f g = overloading overloading_spec #> g #> Local_Theory.exit_result_global f; end; diff --git a/src/Pure/Isar/proof_context.ML b/src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML +++ b/src/Pure/Isar/proof_context.ML @@ -1,1657 +1,1657 @@ (* Title: Pure/Isar/proof_context.ML Author: Markus Wenzel, TU Muenchen The key concept of Isar proof contexts: elevates primitive local reasoning Gamma |- phi to a structured concept, with generic context elements. See also structure Variable and Assumption. *) signature PROOF_CONTEXT = sig val theory_of: Proof.context -> theory val init_global: theory -> Proof.context val get_global: theory -> string -> Proof.context type mode val mode_default: mode val mode_pattern: mode val mode_schematic: mode val mode_abbrev: mode val set_mode: mode -> Proof.context -> Proof.context val get_mode: Proof.context -> mode val restore_mode: Proof.context -> Proof.context -> Proof.context val abbrev_mode: Proof.context -> bool val syn_of: Proof.context -> Syntax.syntax val tsig_of: Proof.context -> Type.tsig val set_defsort: sort -> Proof.context -> Proof.context val default_sort: Proof.context -> indexname -> sort val arity_sorts: Proof.context -> string -> sort -> sort list val consts_of: Proof.context -> Consts.T val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context val map_naming: (Name_Space.naming -> Name_Space.naming) -> Proof.context -> Proof.context val naming_of: Proof.context -> Name_Space.naming val restore_naming: Proof.context -> Proof.context -> Proof.context val full_name: Proof.context -> binding -> string val get_scope: Proof.context -> Binding.scope option val new_scope: Proof.context -> Binding.scope * Proof.context val private_scope: Binding.scope -> Proof.context -> Proof.context val private: Position.T -> Proof.context -> Proof.context val qualified_scope: Binding.scope -> Proof.context -> Proof.context val qualified: Position.T -> Proof.context -> Proof.context val concealed: Proof.context -> Proof.context val class_space: Proof.context -> Name_Space.T val type_space: Proof.context -> Name_Space.T val const_space: Proof.context -> Name_Space.T val defs_context: Proof.context -> Defs.context val intern_class: Proof.context -> xstring -> string val intern_type: Proof.context -> xstring -> string val intern_const: Proof.context -> xstring -> string val extern_class: Proof.context -> string -> xstring val markup_class: Proof.context -> string -> string val pretty_class: Proof.context -> string -> Pretty.T val extern_type: Proof.context -> string -> xstring val markup_type: Proof.context -> string -> string val pretty_type: Proof.context -> string -> Pretty.T val extern_const: Proof.context -> string -> xstring val markup_const: Proof.context -> string -> string val pretty_const: Proof.context -> string -> Pretty.T val transfer: theory -> Proof.context -> Proof.context val transfer_facts: theory -> Proof.context -> Proof.context val background_theory: (theory -> theory) -> Proof.context -> Proof.context val background_theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context val facts_of: Proof.context -> Facts.T val facts_of_fact: Proof.context -> string -> Facts.T val markup_extern_fact: Proof.context -> string -> Markup.T list * xstring val augment: term -> Proof.context -> Proof.context val print_name: Proof.context -> string -> string val pretty_name: Proof.context -> string -> Pretty.T val pretty_term_abbrev: Proof.context -> term -> Pretty.T val pretty_fact: Proof.context -> string * thm list -> Pretty.T val check_class: Proof.context -> xstring * Position.T -> class * Position.report list val read_class: Proof.context -> string -> class val read_typ: Proof.context -> string -> typ val read_typ_syntax: Proof.context -> string -> typ val read_typ_abbrev: Proof.context -> string -> typ val cert_typ: Proof.context -> typ -> typ val cert_typ_syntax: Proof.context -> typ -> typ val cert_typ_abbrev: Proof.context -> typ -> typ val infer_type: Proof.context -> string * typ -> typ val inferred_param: string -> Proof.context -> (string * typ) * Proof.context val inferred_fixes: Proof.context -> (string * typ) list * Proof.context val check_type_name: {proper: bool, strict: bool} -> Proof.context -> xstring * Position.T -> typ * Position.report list val read_type_name: {proper: bool, strict: bool} -> Proof.context -> string -> typ val consts_completion_message: Proof.context -> xstring * Position.T list -> string val check_const: {proper: bool, strict: bool} -> Proof.context -> xstring * Position.T list -> term * Position.report list val read_const: {proper: bool, strict: bool} -> Proof.context -> string -> term val read_arity: Proof.context -> xstring * string list * string -> arity val cert_arity: Proof.context -> arity -> arity val allow_dummies: Proof.context -> Proof.context val prepare_sortsT: Proof.context -> typ list -> string list * typ list val prepare_sorts: Proof.context -> term list -> string list * term list val check_tfree: Proof.context -> string * sort -> string * sort val read_term_pattern: Proof.context -> string -> term val read_term_schematic: Proof.context -> string -> term val read_term_abbrev: Proof.context -> string -> term val show_abbrevs: bool Config.T val expand_abbrevs: Proof.context -> term -> term val cert_term: Proof.context -> term -> term val cert_prop: Proof.context -> term -> term val def_type: Proof.context -> indexname -> typ option val standard_typ_check: Proof.context -> typ list -> typ list val standard_term_check_finish: Proof.context -> term list -> term list val standard_term_uncheck: Proof.context -> term list -> term list val goal_export: Proof.context -> Proof.context -> thm list -> thm list val export: Proof.context -> Proof.context -> thm list -> thm list val export_morphism: Proof.context -> Proof.context -> morphism val norm_export_morphism: Proof.context -> Proof.context -> morphism val auto_bind_goal: term list -> Proof.context -> Proof.context val auto_bind_facts: term list -> Proof.context -> Proof.context val simult_matches: Proof.context -> term * term list -> (indexname * term) list val maybe_bind_term: indexname * term option -> Proof.context -> Proof.context val bind_term: indexname * term -> Proof.context -> Proof.context val cert_propp: Proof.context -> (term * term list) list list -> (term list list * (indexname * term) list) val read_propp: Proof.context -> (string * string list) list list -> (term list list * (indexname * term) list) val fact_tac: Proof.context -> thm list -> int -> tactic val some_fact_tac: Proof.context -> int -> tactic val lookup_fact: Proof.context -> string -> {dynamic: bool, thms: thm list} option val dynamic_facts_dummy: bool Config.T val get_fact_generic: Context.generic -> Facts.ref -> string option * thm list val get_fact: Proof.context -> Facts.ref -> thm list val get_fact_single: Proof.context -> Facts.ref -> thm val get_thms: Proof.context -> xstring -> thm list val get_thm: Proof.context -> xstring -> thm val is_stmt: Proof.context -> bool val set_stmt: bool -> Proof.context -> Proof.context val restore_stmt: Proof.context -> Proof.context -> Proof.context val add_thms_dynamic: binding * (Context.generic -> thm list) -> Proof.context -> string * Proof.context val add_thms_lazy: string -> (binding * thm list lazy) -> Proof.context -> Proof.context val note_thms: string -> Thm.binding * (thm list * attribute list) list -> Proof.context -> (string * thm list) * Proof.context val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list -> Proof.context -> (string * thm list) list * Proof.context val put_thms: bool -> string * thm list option -> Proof.context -> Proof.context val alias_fact: binding -> string -> Proof.context -> Proof.context val read_var: binding * string option * mixfix -> Proof.context -> (binding * typ option * mixfix) * Proof.context val cert_var: binding * typ option * mixfix -> Proof.context -> (binding * typ option * mixfix) * Proof.context val add_fixes: (binding * typ option * mixfix) list -> Proof.context -> string list * Proof.context val add_fixes_cmd: (binding * string option * mixfix) list -> Proof.context -> string list * Proof.context val add_assms: Assumption.export -> (Thm.binding * (term * term list) list) list -> Proof.context -> (string * thm list) list * Proof.context val add_assms_cmd: Assumption.export -> (Thm.binding * (string * string list) list) list -> Proof.context -> (string * thm list) list * Proof.context val dest_cases: Proof.context option -> Proof.context -> (string * Rule_Cases.T) list val update_cases: (string * Rule_Cases.T option) list -> Proof.context -> Proof.context val apply_case: Rule_Cases.T -> Proof.context -> (string * term list) list * Proof.context val check_case: Proof.context -> bool -> string * Position.T -> binding option list -> Rule_Cases.T val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> Proof.context -> Proof.context val notation: bool -> Syntax.mode -> (term * mixfix) list -> Proof.context -> Proof.context val generic_type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> morphism -> Context.generic -> Context.generic val generic_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism -> Context.generic -> Context.generic val type_alias: binding -> string -> Proof.context -> Proof.context val const_alias: binding -> string -> Proof.context -> Proof.context val add_const_constraint: string * typ option -> Proof.context -> Proof.context val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context val revert_abbrev: string -> string -> Proof.context -> Proof.context val generic_add_abbrev: string -> binding * term -> Context.generic -> (term * term) * Context.generic val generic_revert_abbrev: string -> string -> Context.generic -> Context.generic type stmt = {vars: ((binding * typ option * mixfix) * (string * term)) list, propss: term list list, binds: (indexname * term) list, result_binds: (indexname * term) list} val cert_stmt: (binding * typ option * mixfix) list -> (term * term list) list list -> Proof.context -> stmt * Proof.context val read_stmt: (binding * string option * mixfix) list -> (string * string list) list list -> Proof.context -> stmt * Proof.context type statement = {fixes: (string * term) list, assumes: term list list, shows: term list list, result_binds: (indexname * term option) list, text: term, result_text: term} val cert_statement: (binding * typ option * mixfix) list -> (term * term list) list list -> (term * term list) list list -> Proof.context -> statement * Proof.context val read_statement: (binding * string option * mixfix) list -> (string * string list) list list -> (string * string list) list list -> Proof.context -> statement * Proof.context val print_syntax: Proof.context -> unit val print_abbrevs: bool -> Proof.context -> unit val pretty_term_bindings: Proof.context -> Pretty.T list val pretty_local_facts: bool -> Proof.context -> Pretty.T list val print_local_facts: bool -> Proof.context -> unit val pretty_cases: Proof.context -> Pretty.T list val print_cases_proof: Proof.context -> Proof.context -> string val debug: bool Config.T val verbose: bool Config.T val pretty_ctxt: Proof.context -> Pretty.T list val pretty_context: Proof.context -> Pretty.T list end; structure Proof_Context: PROOF_CONTEXT = struct val theory_of = Proof_Context.theory_of; val init_global = Proof_Context.init_global; val get_global = Proof_Context.get_global; (** inner syntax mode **) datatype mode = Mode of {pattern: bool, (*pattern binding schematic variables*) schematic: bool, (*term referencing loose schematic variables*) abbrev: bool}; (*abbrev mode -- no normalization*) fun make_mode (pattern, schematic, abbrev) = Mode {pattern = pattern, schematic = schematic, abbrev = abbrev}; val mode_default = make_mode (false, false, false); val mode_pattern = make_mode (true, false, false); val mode_schematic = make_mode (false, true, false); val mode_abbrev = make_mode (false, false, true); (** Isar proof context information **) type cases = Rule_Cases.T Name_Space.table; val empty_cases: cases = Name_Space.empty_table Markup.caseN; datatype data = Data of {mode: mode, (*inner syntax mode*) syntax: Local_Syntax.T, (*local syntax*) tsig: Type.tsig * Type.tsig, (*local/global type signature -- local name space / defsort only*) consts: Consts.T * Consts.T, (*local/global consts -- local name space / abbrevs only*) facts: Facts.T, (*local facts, based on initial global facts*) cases: cases}; (*named case contexts*) fun make_data (mode, syntax, tsig, consts, facts, cases) = Data {mode = mode, syntax = syntax, tsig = tsig, consts = consts, facts = facts, cases = cases}; structure Data = Proof_Data ( type T = data; fun init thy = make_data (mode_default, Local_Syntax.init thy, (Type.change_ignore (Sign.tsig_of thy), Sign.tsig_of thy), (Consts.change_ignore (Sign.consts_of thy), Sign.consts_of thy), Global_Theory.facts_of thy, empty_cases); ); fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep); fun map_data_result f ctxt = let val Data {mode, syntax, tsig, consts, facts, cases} = Data.get ctxt; val (res, data') = f (mode, syntax, tsig, consts, facts, cases) ||> make_data; in (res, Data.put data' ctxt) end; fun map_data f = snd o map_data_result (pair () o f); fun set_mode mode = map_data (fn (_, syntax, tsig, consts, facts, cases) => (mode, syntax, tsig, consts, facts, cases)); fun map_syntax f = map_data (fn (mode, syntax, tsig, consts, facts, cases) => (mode, f syntax, tsig, consts, facts, cases)); fun map_syntax_idents f ctxt = let val (opt_idents', syntax') = f (#syntax (rep_data ctxt)) in ctxt |> map_syntax (K syntax') |> (case opt_idents' of NONE => I | SOME idents' => Syntax_Trans.put_idents idents') end; fun map_tsig f = map_data (fn (mode, syntax, tsig, consts, facts, cases) => (mode, syntax, f tsig, consts, facts, cases)); fun map_consts f = map_data (fn (mode, syntax, tsig, consts, facts, cases) => (mode, syntax, tsig, f consts, facts, cases)); fun map_facts_result f = map_data_result (fn (mode, syntax, tsig, consts, facts, cases) => let val (res, facts') = f facts in (res, (mode, syntax, tsig, consts, facts', cases)) end); fun map_facts f = snd o map_facts_result (pair () o f); fun map_cases f = map_data (fn (mode, syntax, tsig, consts, facts, cases) => (mode, syntax, tsig, consts, facts, f cases)); val get_mode = #mode o rep_data; val restore_mode = set_mode o get_mode; val abbrev_mode = get_mode #> (fn Mode {abbrev, ...} => abbrev); val syntax_of = #syntax o rep_data; val syn_of = Local_Syntax.syn_of o syntax_of; val set_syntax_mode = map_syntax o Local_Syntax.set_mode; val restore_syntax_mode = map_syntax o Local_Syntax.restore_mode o syntax_of; val tsig_of = #1 o #tsig o rep_data; val set_defsort = map_tsig o apfst o Type.set_defsort; fun default_sort ctxt = the_default (Type.defaultS (tsig_of ctxt)) o Variable.def_sort ctxt; fun arity_sorts ctxt = Type.arity_sorts (Context.Proof ctxt) (tsig_of ctxt); val consts_of = #1 o #consts o rep_data; val cases_of = #cases o rep_data; (* naming *) val naming_of = Name_Space.naming_of o Context.Proof; val map_naming = Context.proof_map o Name_Space.map_naming; val restore_naming = map_naming o K o naming_of; val full_name = Name_Space.full_name o naming_of; val get_scope = Name_Space.get_scope o naming_of; fun new_scope ctxt = let val (scope, naming') = Name_Space.new_scope (naming_of ctxt); val ctxt' = map_naming (K naming') ctxt; in (scope, ctxt') end; val private_scope = map_naming o Name_Space.private_scope; val private = map_naming o Name_Space.private; val qualified_scope = map_naming o Name_Space.qualified_scope; val qualified = map_naming o Name_Space.qualified; val concealed = map_naming Name_Space.concealed; (* name spaces *) val class_space = Type.class_space o tsig_of; val type_space = Type.type_space o tsig_of; val const_space = Consts.space_of o consts_of; fun defs_context ctxt = (ctxt, (const_space ctxt, type_space ctxt)); val intern_class = Name_Space.intern o class_space; val intern_type = Name_Space.intern o type_space; val intern_const = Name_Space.intern o const_space; fun extern_class ctxt = Name_Space.extern ctxt (class_space ctxt); fun extern_type ctxt = Name_Space.extern ctxt (type_space ctxt); fun extern_const ctxt = Name_Space.extern ctxt (const_space ctxt); fun markup_class ctxt c = Name_Space.markup_extern ctxt (class_space ctxt) c |-> Markup.markup; fun markup_type ctxt c = Name_Space.markup_extern ctxt (type_space ctxt) c |-> Markup.markup; fun markup_const ctxt c = Name_Space.markup_extern ctxt (const_space ctxt) c |-> Markup.markup; fun pretty_class ctxt c = Name_Space.markup_extern ctxt (class_space ctxt) c |> Pretty.mark_str; fun pretty_type ctxt c = Name_Space.markup_extern ctxt (type_space ctxt) c |> Pretty.mark_str; fun pretty_const ctxt c = Name_Space.markup_extern ctxt (const_space ctxt) c |> Pretty.mark_str; (* theory transfer *) fun transfer_syntax thy ctxt = ctxt |> map_syntax (Local_Syntax.rebuild thy) |> map_tsig (fn tsig as (local_tsig, global_tsig) => let val thy_tsig = Sign.tsig_of thy in if Type.eq_tsig (thy_tsig, global_tsig) then tsig else (Type.merge_tsig (Context.Proof ctxt) (local_tsig, thy_tsig), thy_tsig) (*historic merge order*) end) |> map_consts (fn consts as (local_consts, global_consts) => let val thy_consts = Sign.consts_of thy in if Consts.eq_consts (thy_consts, global_consts) then consts else (Consts.merge (local_consts, thy_consts), thy_consts) (*historic merge order*) end); fun transfer thy = Context.raw_transfer thy #> transfer_syntax thy; fun transfer_facts thy = map_facts (fn local_facts => Facts.merge (Global_Theory.facts_of thy, local_facts)); fun background_theory f ctxt = transfer (f (theory_of ctxt)) ctxt; fun background_theory_result f ctxt = let val (res, thy') = f (theory_of ctxt) in (res, ctxt |> transfer thy') end; (* hybrid facts *) val facts_of = #facts o rep_data; fun facts_of_fact ctxt name = let val local_facts = facts_of ctxt; val global_facts = Global_Theory.facts_of (theory_of ctxt); in if Facts.defined local_facts name then local_facts else global_facts end; fun markup_extern_fact ctxt name = let val facts = facts_of_fact ctxt name; val (markup, xname) = Facts.markup_extern ctxt facts name; val markups = if Facts.is_dynamic facts name then [markup, Markup.dynamic_fact name] else [markup]; in (markups, xname) end; (* augment context by implicit term declarations *) fun augment t ctxt = ctxt |> Variable.add_fixes_implicit t |> Variable.declare_term t |> Soft_Type_System.augment t; (** pretty printing **) val print_name = Token.print_name o Thy_Header.get_keywords'; val pretty_name = Pretty.str oo print_name; fun pretty_term_abbrev ctxt = Syntax.pretty_term (set_mode mode_abbrev ctxt); fun pretty_fact_name ctxt a = Pretty.block [Pretty.marks_str (markup_extern_fact ctxt a), Pretty.str ":"]; fun pretty_fact ctxt = let val pretty_thm = Thm.pretty_thm ctxt; val pretty_thms = map (Thm.pretty_thm_item ctxt); in fn ("", [th]) => pretty_thm th | ("", ths) => Pretty.blk (0, Pretty.fbreaks (pretty_thms ths)) | (a, [th]) => Pretty.block [pretty_fact_name ctxt a, Pretty.brk 1, pretty_thm th] | (a, ths) => Pretty.block (Pretty.fbreaks (pretty_fact_name ctxt a :: pretty_thms ths)) end; (** prepare types **) (* classes *) fun check_class ctxt (xname, pos) = let val tsig = tsig_of ctxt; val class_space = Type.class_space tsig; val name = Type.cert_class tsig (Name_Space.intern class_space xname) handle TYPE (msg, _, _) => error (msg ^ Position.here pos ^ Completion.markup_report [Name_Space.completion (Context.Proof ctxt) class_space (K true) (xname, pos)]); val reports = if Context_Position.is_reported ctxt pos then [(pos, Name_Space.markup class_space name)] else []; in (name, reports) end; fun read_class ctxt text = let val source = Syntax.read_input text; val (c, reports) = check_class ctxt (Input.source_content source); val _ = Context_Position.reports ctxt reports; in c end; (* types *) fun read_typ_mode mode ctxt s = Syntax.read_typ (Type.set_mode mode ctxt) s; val read_typ = read_typ_mode Type.mode_default; val read_typ_syntax = read_typ_mode Type.mode_syntax; val read_typ_abbrev = read_typ_mode Type.mode_abbrev; fun cert_typ_mode mode ctxt T = Type.cert_typ_mode mode (tsig_of ctxt) T handle TYPE (msg, _, _) => error msg; val cert_typ = cert_typ_mode Type.mode_default; val cert_typ_syntax = cert_typ_mode Type.mode_syntax; val cert_typ_abbrev = cert_typ_mode Type.mode_abbrev; (** prepare terms and propositions **) (* inferred types of parameters *) fun infer_type ctxt x = Term.fastype_of (singleton (Syntax.check_terms (set_mode mode_schematic ctxt)) (Free x)); fun inferred_param x ctxt = let val p = (x, infer_type ctxt (x, dummyT)) in (p, ctxt |> Variable.declare_term (Free p)) end; fun inferred_fixes ctxt = fold_map inferred_param (map #2 (Variable.dest_fixes ctxt)) ctxt; (* type names *) fun check_type_name {proper, strict} ctxt (c, pos) = if Lexicon.is_tid c then if proper then error ("Not a type constructor: " ^ quote c ^ Position.here pos) else let val reports = if Context_Position.is_reported ctxt pos then [(pos, Markup.tfree)] else []; in (TFree (c, default_sort ctxt (c, ~1)), reports) end else let val ((d, reports), decl) = Type.check_decl (Context.Proof ctxt) (tsig_of ctxt) (c, pos); fun err () = error ("Bad type name: " ^ quote d ^ Position.here pos); val args = (case decl of Type.LogicalType n => n | Type.Abbreviation (vs, _, _) => if strict then err () else length vs | Type.Nonterminal => if strict then err () else 0); in (Type (d, replicate args dummyT), reports) end; fun read_type_name flags ctxt text = let val source = Syntax.read_input text; val (T, reports) = check_type_name flags ctxt (Input.source_content source); val _ = Context_Position.reports ctxt reports; in T end; (* constant names *) fun consts_completion_message ctxt (c, ps) = ps |> map (fn pos => Name_Space.completion (Context.Proof ctxt) (Consts.space_of (consts_of ctxt)) (K true) (c, pos)) |> Completion.markup_report; fun check_const {proper, strict} ctxt (c, ps) = let val _ = Name.reject_internal (c, ps) handle ERROR msg => error (msg ^ consts_completion_message ctxt (c, ps)); fun err msg = error (msg ^ Position.here_list ps); val consts = consts_of ctxt; val fixed = if proper then NONE else Variable.lookup_fixed ctxt c; val (t, reports) = (case (fixed, Variable.lookup_const ctxt c) of (SOME x, NONE) => let val reports = ps |> filter (Context_Position.is_reported ctxt) |> map (fn pos => (pos, Markup.name x (if Name.is_skolem x then Markup.skolem else Markup.free))); in (Free (x, infer_type ctxt (x, dummyT)), reports) end | (_, SOME d) => let val T = Consts.type_scheme consts d handle TYPE (msg, _, _) => err msg; val reports = ps |> filter (Context_Position.is_reported ctxt) |> map (fn pos => (pos, Name_Space.markup (Consts.space_of consts) d)); in (Const (d, T), reports) end | _ => Consts.check_const (Context.Proof ctxt) consts (c, ps)); val _ = (case (strict, t) of (true, Const (d, _)) => (ignore (Consts.the_const consts d) handle TYPE (msg, _, _) => err msg) | _ => ()); in (t, reports) end; fun read_const flags ctxt text = let val source = Syntax.read_input text; val (c, pos) = Input.source_content source; val (t, reports) = check_const flags ctxt (c, [pos]); val _ = Context_Position.reports ctxt reports; in t end; (* type arities *) local fun prep_arity prep_tycon prep_sort ctxt (t, Ss, S) = let val arity = (prep_tycon ctxt t, map (prep_sort ctxt) Ss, prep_sort ctxt S) in Type.add_arity (Context.Proof ctxt) arity (tsig_of ctxt); arity end; in val read_arity = prep_arity ((#1 o dest_Type) oo read_type_name {proper = true, strict = true}) Syntax.read_sort; val cert_arity = prep_arity (K I) (Type.cert_sort o tsig_of); end; (* read_term *) fun read_term_mode mode ctxt = Syntax.read_term (set_mode mode ctxt); val read_term_pattern = read_term_mode mode_pattern; val read_term_schematic = read_term_mode mode_schematic; val read_term_abbrev = read_term_mode mode_abbrev; (* local abbreviations *) local fun certify_consts ctxt = Consts.certify (Context.Proof ctxt) (tsig_of ctxt) (not (abbrev_mode ctxt)) (consts_of ctxt); fun expand_binds ctxt = let val Mode {pattern, schematic, ...} = get_mode ctxt; fun reject_schematic (t as Var _) = error ("Unbound schematic variable: " ^ Syntax.string_of_term ctxt t) | reject_schematic (Abs (_, _, t)) = reject_schematic t | reject_schematic (t $ u) = (reject_schematic t; reject_schematic u) | reject_schematic _ = (); in if pattern then I else Variable.expand_binds ctxt #> (if schematic then I else tap reject_schematic) end; in fun expand_abbrevs ctxt = certify_consts ctxt #> expand_binds ctxt; end; val show_abbrevs = Config.declare_bool ("show_abbrevs", \<^here>) (K true); fun contract_abbrevs ctxt t = let val thy = theory_of ctxt; val consts = consts_of ctxt; val Mode {abbrev, ...} = get_mode ctxt; val retrieve = Consts.retrieve_abbrevs consts (print_mode_value () @ [""]); fun match_abbrev u = Option.map #1 (get_first (Pattern.match_rew thy u) (retrieve u)); in if abbrev orelse not (Config.get ctxt show_abbrevs) orelse not (can Term.type_of t) then t else Pattern.rewrite_term_top thy [] [match_abbrev] t end; (* patterns *) fun prepare_patternT ctxt T = let val Mode {pattern, schematic, ...} = get_mode ctxt; val _ = pattern orelse schematic orelse T |> Term.exists_subtype (fn T as TVar (xi, _) => not (Type_Infer.is_param xi) andalso error ("Illegal schematic type variable: " ^ Syntax.string_of_typ ctxt T) | _ => false) in T end; local val dummies = Config.declare_bool ("Proof_Context.dummies", \<^here>) (K false); fun check_dummies ctxt t = if Config.get ctxt dummies then t else Term.no_dummy_patterns t handle TERM _ => error "Illegal dummy pattern(s) in term"; fun prepare_dummies ts = #1 (fold_map Term.replace_dummy_patterns ts 1); in val allow_dummies = Config.put dummies true; fun prepare_patterns ctxt = let val Mode {pattern, ...} = get_mode ctxt in Type_Infer.fixate ctxt pattern #> pattern ? Variable.polymorphic ctxt #> (map o Term.map_types) (prepare_patternT ctxt) #> (if pattern then prepare_dummies else map (check_dummies ctxt)) end; end; (* sort constraints *) local fun prepare_sorts_env ctxt tys = let val tsig = tsig_of ctxt; val defaultS = Type.defaultS tsig; val dummy_var = ("'_dummy_", ~1); fun constraint (xi, raw_S) env = let val (ps, S) = Term_Position.decode_positionS raw_S in if xi = dummy_var orelse S = dummyS then env else Vartab.insert (op =) (xi, Type.minimize_sort tsig S) env handle Vartab.DUP _ => error ("Inconsistent sort constraints for type variable " ^ quote (Term.string_of_vname' xi) ^ Position.here_list ps) end; val env = - (fold o fold_atyps) + Vartab.build (tys |> (fold o fold_atyps) (fn TFree (x, S) => constraint ((x, ~1), S) | TVar v => constraint v - | _ => I) tys Vartab.empty; + | _ => I)); fun get_sort xi raw_S = if xi = dummy_var then Type.minimize_sort tsig (#2 (Term_Position.decode_positionS raw_S)) else (case (Vartab.lookup env xi, Variable.def_sort ctxt xi) of (NONE, NONE) => defaultS | (NONE, SOME S) => S | (SOME S, NONE) => S | (SOME S, SOME S') => if Type.eq_sort tsig (S, S') then S' else error ("Sort constraint " ^ Syntax.string_of_sort ctxt S ^ " inconsistent with default " ^ Syntax.string_of_sort ctxt S' ^ " for type variable " ^ quote (Term.string_of_vname' xi))); fun add_report S pos reports = if Position.is_reported pos andalso not (AList.defined (op =) reports pos) then (pos, Position.reported_text pos Markup.sorting (Syntax.string_of_sort ctxt S)) :: reports else reports; fun get_sort_reports xi raw_S = let val ps = #1 (Term_Position.decode_positionS raw_S); val S = get_sort xi raw_S handle ERROR msg => error (msg ^ Position.here_list ps); in fold (add_report S) ps end; val reports = (fold o fold_atyps) (fn T => if Term_Position.is_positionT T then I else (case T of TFree (x, raw_S) => get_sort_reports (x, ~1) raw_S | TVar (xi, raw_S) => get_sort_reports xi raw_S | _ => I)) tys []; in (map #2 reports, get_sort) end; fun replace_sortsT get_sort = map_atyps (fn T => if Term_Position.is_positionT T then T else (case T of TFree (x, raw_S) => TFree (x, get_sort (x, ~1) raw_S) | TVar (xi, raw_S) => TVar (xi, get_sort xi raw_S) | _ => T)); in fun prepare_sortsT ctxt tys = let val (sorting_report, get_sort) = prepare_sorts_env ctxt tys in (sorting_report, map (replace_sortsT get_sort) tys) end; fun prepare_sorts ctxt tms = let val tys = rev ((fold o fold_types) cons tms []); val (sorting_report, get_sort) = prepare_sorts_env ctxt tys; in (sorting_report, (map o map_types) (replace_sortsT get_sort) tms) end; fun check_tfree ctxt v = let val (sorting_report, [TFree a]) = prepare_sortsT ctxt [TFree v]; val _ = if Context_Position.reports_enabled ctxt then Output.report sorting_report else (); in a end; end; (* certify terms *) local fun gen_cert prop ctxt t = t |> expand_abbrevs ctxt |> (fn t' => #1 (Sign.certify' prop (Context.Proof ctxt) false (consts_of ctxt) (theory_of ctxt) t') handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg); in val cert_term = gen_cert false; val cert_prop = gen_cert true; end; (* check/uncheck *) fun def_type ctxt = let val Mode {pattern, ...} = get_mode ctxt in Variable.def_type ctxt pattern end; fun standard_typ_check ctxt = map (cert_typ_mode (Type.get_mode ctxt) ctxt #> prepare_patternT ctxt); val standard_term_check_finish = prepare_patterns; fun standard_term_uncheck ctxt = map (contract_abbrevs ctxt); (** export results **) fun common_export is_goal inner outer = map (Assumption.export is_goal inner outer) #> Variable.export inner outer; val goal_export = common_export true; val export = common_export false; fun export_morphism inner outer = Assumption.export_morphism inner outer $> Variable.export_morphism inner outer; fun norm_export_morphism inner outer = export_morphism inner outer $> Morphism.thm_morphism "Proof_Context.norm_export" (Goal.norm_result outer); (** term bindings **) (* auto bindings *) fun auto_bind f props ctxt = fold Variable.maybe_bind_term (f ctxt props) ctxt; val auto_bind_goal = auto_bind Auto_Bind.goal; val auto_bind_facts = auto_bind Auto_Bind.facts; (* match bindings *) fun simult_matches ctxt (t, pats) = (case Seq.pull (Unify.matchers (Context.Proof ctxt) (map (rpair t) pats)) of NONE => error "Pattern match failed!" | SOME (env, _) => Vartab.fold (fn (v, (_, t)) => cons (v, t)) (Envir.term_env env) []); fun maybe_bind_term (xi, t) ctxt = ctxt |> Variable.maybe_bind_term (xi, Option.map (cert_term (set_mode mode_default ctxt)) t); val bind_term = maybe_bind_term o apsnd SOME; (* propositions with patterns *) local fun prep_propp prep_props ctxt raw_args = let val props = prep_props ctxt (maps (map fst) raw_args); val props_ctxt = fold Variable.declare_term props ctxt; val patss = maps (map (prep_props (set_mode mode_pattern props_ctxt) o snd)) raw_args; val propps = unflat raw_args (props ~~ patss); val binds = (maps o maps) (simult_matches props_ctxt) propps; in (map (map fst) propps, binds) end; in val cert_propp = prep_propp (map o cert_prop); val read_propp = prep_propp Syntax.read_props; end; (** theorems **) (* fact_tac *) local fun comp_hhf_tac ctxt th i st = PRIMSEQ (Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = true} (false, Drule.lift_all ctxt (Thm.cprem_of st i) th, 0) i) st; fun comp_incr_tac _ [] _ = no_tac | comp_incr_tac ctxt (th :: ths) i = (fn st => comp_hhf_tac ctxt (Drule.incr_indexes st th) i st) APPEND comp_incr_tac ctxt ths i; val vacuous_facts = [Drule.termI]; in fun potential_facts ctxt prop = let val body = Term.strip_all_body prop; val vacuous = filter (fn th => Term.could_unify (body, Thm.concl_of th)) vacuous_facts |> map (rpair Position.none); in Facts.could_unify (facts_of ctxt) body @ vacuous end; fun fact_tac ctxt facts = Goal.norm_hhf_tac ctxt THEN' comp_incr_tac ctxt facts; fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) => fact_tac ctxt (map #1 (potential_facts ctxt goal)) i); end; (* lookup facts *) fun lookup_fact ctxt name = let val context = Context.Proof ctxt; val thy = Proof_Context.theory_of ctxt; in (case Facts.lookup context (facts_of ctxt) name of NONE => Facts.lookup context (Global_Theory.facts_of thy) name | some => some) end; (* retrieve facts *) val dynamic_facts_dummy = Config.declare_bool ("dynamic_facts_dummy_", \<^here>) (K false); local fun retrieve_global context = Facts.retrieve context (Global_Theory.facts_of (Context.theory_of context)); fun retrieve_generic (context as Context.Proof ctxt) arg = (Facts.retrieve context (facts_of ctxt) arg handle ERROR local_msg => (retrieve_global context arg handle ERROR _ => error local_msg)) | retrieve_generic context arg = retrieve_global context arg; fun retrieve pick context (Facts.Fact s) = let val ctxt = Context.the_proof context; val pos = Syntax.read_input_pos s; val prop = Syntax.read_prop (ctxt |> set_mode mode_default |> allow_dummies) s |> singleton (Variable.polymorphic ctxt); fun err ps msg = error (msg ^ Position.here_list (pos :: ps) ^ ":\n" ^ Syntax.string_of_term ctxt prop); val (prop', _) = Term.replace_dummy_patterns prop (Variable.maxidx_of ctxt + 1); fun prove th = Goal.prove ctxt [] [] prop' (K (ALLGOALS (fact_tac ctxt [th]))); val results = map_filter (try (apfst prove)) (potential_facts ctxt prop'); val (thm, thm_pos) = (case distinct (eq_fst Thm.eq_thm_prop) results of [res] => res | [] => err [] "Failed to retrieve literal fact" | dups => err (distinct (op =) (map #2 dups)) "Ambiguous specification of literal fact"); val markup = Position.entity_markup Markup.literal_factN ("", thm_pos); val _ = Context_Position.report_generic context pos markup; in pick true ("", thm_pos) [thm] end | retrieve pick context (Facts.Named ((xname, pos), sel)) = let val thy = Context.theory_of context; fun immediate thms = {name = xname, dynamic = false, thms = map (Thm.transfer thy) thms}; val {name, dynamic, thms} = (case xname of "" => immediate [Drule.dummy_thm] | "_" => immediate [Drule.asm_rl] | "nothing" => immediate [] | _ => retrieve_generic context (xname, pos)); val thms' = if dynamic andalso Config.get_generic context dynamic_facts_dummy then [Drule.free_dummy_thm] else Facts.select (Facts.Named ((name, pos), sel)) thms; in pick (dynamic andalso is_none sel) (name, pos) thms' end; in val get_fact_generic = retrieve (fn dynamic => fn (name, _) => fn thms => (if dynamic then SOME name else NONE, thms)); val get_fact = retrieve (K (K I)) o Context.Proof; val get_fact_single = retrieve (K Facts.the_single) o Context.Proof; fun get_thms ctxt = get_fact ctxt o Facts.named; fun get_thm ctxt = get_fact_single ctxt o Facts.named; end; (* inner statement mode *) val inner_stmt = Config.declare_bool ("inner_stmt", \<^here>) (K false); fun is_stmt ctxt = Config.get ctxt inner_stmt; val set_stmt = Config.put inner_stmt; val restore_stmt = set_stmt o is_stmt; (* facts *) fun add_thms_dynamic arg ctxt = ctxt |> map_facts_result (Facts.add_dynamic (Context.Proof ctxt) arg); local fun add_facts {index} arg ctxt = ctxt |> map_facts_result (Facts.add_static (Context.Proof ctxt) {strict = false, index = index} arg); fun update_facts flags (b, SOME ths) ctxt = ctxt |> add_facts flags (b, Lazy.value ths) |> #2 | update_facts _ (b, NONE) ctxt = ctxt |> map_facts (Facts.del (full_name ctxt b)); fun bind_name ctxt b = (full_name ctxt b, Binding.default_pos_of b); in fun add_thms_lazy kind (b, ths) ctxt = let val name_pos = bind_name ctxt b; val ths' = Global_Theory.check_thms_lazy ths |> Lazy.map_finished (Global_Theory.name_thms Global_Theory.unofficial1 name_pos #> map (Thm.kind_rule kind)); val (_, ctxt') = add_facts {index = is_stmt ctxt} (b, ths') ctxt; in ctxt' end; fun note_thms kind ((b, more_atts), facts) ctxt = let val (name, pos) = bind_name ctxt b; val facts' = facts |> Global_Theory.burrow_fact (Global_Theory.name_thms Global_Theory.unofficial1 (name, pos)); fun app (ths, atts) = fold_map (Thm.proof_attributes (surround (Thm.kind kind) (atts @ more_atts))) ths; val (res, ctxt') = fold_map app facts' ctxt; val thms = Global_Theory.name_thms Global_Theory.unofficial2 (name, pos) (flat res); val (_, ctxt'') = ctxt' |> add_facts {index = is_stmt ctxt} (b, Lazy.value thms); in ((name, thms), ctxt'') end; val note_thmss = fold_map o note_thms; fun put_thms index thms ctxt = ctxt |> map_naming (K Name_Space.local_naming) |> Context_Position.set_visible false |> update_facts {index = index} (apfst Binding.name thms) |> Context_Position.restore_visible ctxt |> restore_naming ctxt; end; fun alias_fact b c ctxt = map_facts (Facts.alias (naming_of ctxt) b c) ctxt; (** basic logical entities **) (* variables *) fun declare_var (x, opt_T, mx) ctxt = let val T = (case opt_T of SOME T => T | NONE => Mixfix.default_constraint mx) in (T, ctxt |> Variable.declare_constraints (Free (x, T))) end; fun add_syntax vars ctxt = map_syntax_idents (Local_Syntax.add_syntax ctxt (map (pair Local_Syntax.Fixed) vars)) ctxt; fun check_var internal b = let val x = Variable.check_name b; val check = if internal then Name.reject_skolem else Name.reject_internal; val _ = if can check (x, []) andalso Symbol_Pos.is_identifier x then () else error ("Bad name: " ^ Binding.print b); in x end; local fun check_mixfix ctxt (b, T, mx) = let val ([x], ctxt') = Variable.add_fixes_binding [Binding.reset_pos b] ctxt; val mx' = Mixfix.reset_pos mx; val _ = add_syntax [(x, T, if Context_Position.reports_enabled ctxt then mx else mx')] ctxt'; in mx' end; fun prep_var prep_typ internal (b, raw_T, mx) ctxt = let val x = check_var internal b; fun cond_tvars T = if internal then T else Type.no_tvars T handle TYPE (msg, _, _) => error msg; val opt_T = Option.map (cond_tvars o prep_typ ctxt) raw_T; val (T, ctxt') = ctxt |> declare_var (x, opt_T, mx); val mx' = if Mixfix.is_empty mx then mx else check_mixfix ctxt' (b, T, mx); in ((b, SOME T, mx'), ctxt') end; in val read_var = prep_var Syntax.read_typ false; val cert_var = prep_var cert_typ true; end; (* notation *) local fun type_syntax (Type (c, args), mx) = SOME (Local_Syntax.Type, (Lexicon.mark_type c, Mixfix.make_type (length args), mx)) | type_syntax _ = NONE; fun const_syntax _ (Free (x, T), mx) = SOME (Local_Syntax.Fixed, (x, T, mx)) | const_syntax ctxt (Const (c, _), mx) = (case try (Consts.type_scheme (consts_of ctxt)) c of SOME T => SOME (Local_Syntax.Const, (Lexicon.mark_const c, T, mx)) | NONE => NONE) | const_syntax _ _ = NONE; fun gen_notation syntax add mode args ctxt = ctxt |> map_syntax_idents (Local_Syntax.update_modesyntax ctxt add mode (map_filter (syntax ctxt) args)); in val type_notation = gen_notation (K type_syntax); val notation = gen_notation const_syntax; fun generic_type_notation add mode args phi = let val args' = args |> map_filter (fn (T, mx) => let val T' = Morphism.typ phi T; val similar = (case (T, T') of (Type (c, _), Type (c', _)) => c = c' | _ => false); in if similar then SOME (T', mx) else NONE end); in Context.mapping (Sign.type_notation add mode args') (type_notation add mode args') end; fun generic_notation add mode args phi = let val args' = args |> map_filter (fn (t, mx) => let val t' = Morphism.term phi t in if Term.aconv_untyped (t, t') then SOME (t', mx) else NONE end); in Context.mapping (Sign.notation add mode args') (notation add mode args') end; end; (* aliases *) fun type_alias b c ctxt = (map_tsig o apfst) (Type.type_alias (naming_of ctxt) b c) ctxt; fun const_alias b c ctxt = (map_consts o apfst) (Consts.alias (naming_of ctxt) b c) ctxt; (* local constants *) fun add_const_constraint (c, opt_T) ctxt = let fun prepT raw_T = let val T = cert_typ ctxt raw_T in cert_term ctxt (Const (c, T)); T end; in ctxt |> (map_consts o apfst) (Consts.constrain (c, Option.map prepT opt_T)) end; fun add_abbrev mode (b, raw_t) ctxt = let val t0 = cert_term (ctxt |> set_mode mode_abbrev) raw_t handle ERROR msg => cat_error msg ("in constant abbreviation " ^ Binding.print b); val [t] = Variable.exportT_terms (Variable.declare_term t0 ctxt) ctxt [t0]; val ((lhs, rhs), consts') = consts_of ctxt |> Consts.abbreviate (Context.Proof ctxt) (tsig_of ctxt) mode (b, t); in ctxt |> (map_consts o apfst) (K consts') |> Variable.declare_term rhs |> pair (lhs, rhs) end; fun revert_abbrev mode c = (map_consts o apfst) (Consts.revert_abbrev mode c); fun generic_add_abbrev mode arg = Context.mapping_result (Sign.add_abbrev mode arg) (add_abbrev mode arg); fun generic_revert_abbrev mode arg = Context.mapping (Sign.revert_abbrev mode arg) (revert_abbrev mode arg); (* fixes *) local fun gen_fixes prep_var raw_vars ctxt = let val (vars, _) = fold_map prep_var raw_vars ctxt; val (xs, ctxt') = Variable.add_fixes_binding (map #1 vars) ctxt; val _ = Context_Position.reports ctxt' (flat (map2 (fn x => fn pos => [(pos, Variable.markup ctxt' x), (pos, Variable.markup_entity_def ctxt' x)]) xs (map (Binding.pos_of o #1) vars))); val vars' = map2 (fn x => fn (_, opt_T, mx) => (x, opt_T, mx)) xs vars; val (Ts, ctxt'') = fold_map declare_var vars' ctxt'; val vars'' = map2 (fn T => fn (x, _, mx) => (x, T, mx)) Ts vars'; in (xs, add_syntax vars'' ctxt'') end; in val add_fixes = gen_fixes cert_var; val add_fixes_cmd = gen_fixes read_var; end; (** assumptions **) local fun gen_assms prep_propp exp args ctxt = let val (propss, binds) = prep_propp ctxt (map snd args); val props = flat propss; in ctxt |> fold Variable.declare_term props |> tap (Variable.warn_extra_tfrees ctxt) |> fold_burrow (Assumption.add_assms exp o map (Thm.cterm_of ctxt)) propss |-> (fn premss => auto_bind_facts props #> fold Variable.bind_term binds #> note_thmss "" (map fst args ~~ map (map (fn th => ([th], []))) premss)) end; in val add_assms = gen_assms cert_propp; val add_assms_cmd = gen_assms read_propp; end; (** cases **) fun dest_cases prev_ctxt ctxt = let val serial_of = #serial oo (Name_Space.the_entry o Name_Space.space_of_table); val ignored = (case prev_ctxt of NONE => Inttab.empty | SOME ctxt0 => let val cases0 = cases_of ctxt0 in - Name_Space.fold_table (fn (a, _) => Inttab.insert_set (serial_of cases0 a)) - cases0 Inttab.empty + Inttab.build (cases0 |> Name_Space.fold_table (fn (a, _) => + Inttab.insert_set (serial_of cases0 a))) end); val cases = cases_of ctxt; in Name_Space.fold_table (fn (a, c) => let val i = serial_of cases a in not (Inttab.defined ignored i) ? cons (i, (a, c)) end) cases [] |> sort (int_ord o apply2 #1) |> map #2 end; local fun drop_schematic (b as (xi, SOME t)) = if Term.exists_subterm is_Var t then (xi, NONE) else b | drop_schematic b = b; fun update_case _ ("", _) cases = cases | update_case _ (name, NONE) cases = Name_Space.del_table name cases | update_case context (name, SOME c) cases = #2 (Name_Space.define context false (Binding.name name, c) cases); fun fix (b, T) ctxt = let val ([x], ctxt') = add_fixes [(b, SOME T, NoSyn)] ctxt in (Free (x, T), ctxt') end; in fun update_cases args ctxt = let val context = Context.Proof ctxt |> Name_Space.map_naming (K Name_Space.global_naming); in map_cases (fold (update_case context) args) ctxt end; fun case_result c ctxt = let val Rule_Cases.Case {fixes, ...} = c; val (ts, ctxt') = ctxt |> fold_map fix fixes; val Rule_Cases.Case {assumes, binds, cases, ...} = Rule_Cases.apply ts c; in ctxt' |> fold (maybe_bind_term o drop_schematic) binds |> update_cases (map (apsnd SOME) cases) |> pair (assumes, (binds, cases)) end; val apply_case = apfst fst oo case_result; fun check_case ctxt internal (name, pos) param_specs = let val (_, Rule_Cases.Case {fixes, assumes, binds, cases}) = Name_Space.check (Context.Proof ctxt) (cases_of ctxt) (name, pos); val _ = List.app (fn NONE => () | SOME b => ignore (check_var internal b)) param_specs; fun replace (opt_x :: xs) ((y, T) :: ys) = (the_default y opt_x, T) :: replace xs ys | replace [] ys = ys | replace (_ :: _) [] = error ("Too many parameters for case " ^ quote name ^ Position.here pos); val fixes' = replace param_specs fixes; val binds' = map drop_schematic binds; in if null (fold (Term.add_tvarsT o snd) fixes []) andalso null (fold (fold Term.add_vars o snd) assumes []) then Rule_Cases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases} else error ("Illegal schematic variable(s) in case " ^ quote name ^ Position.here pos) end; end; (* structured statements *) type stmt = {vars: ((binding * typ option * mixfix) * (string * term)) list, propss: term list list, binds: (indexname * term) list, result_binds: (indexname * term) list}; type statement = {fixes: (string * term) list, assumes: term list list, shows: term list list, result_binds: (indexname * term option) list, text: term, result_text: term}; local fun export_binds ctxt' ctxt params binds = let val rhss = map (the_list o Option.map (Logic.close_term params) o snd) binds |> burrow (Variable.export_terms ctxt' ctxt) |> map (try the_single); in map fst binds ~~ rhss end; fun prep_stmt prep_var prep_propp raw_vars raw_propps ctxt = let val (vars, vars_ctxt) = fold_map prep_var raw_vars ctxt; val xs = map (Variable.check_name o #1) vars; val (xs', fixes_ctxt) = add_fixes vars vars_ctxt; val (propss, binds) = prep_propp fixes_ctxt raw_propps; val (ps, params_ctxt) = fixes_ctxt |> (fold o fold) Variable.declare_term propss |> fold_map inferred_param xs'; val params = xs ~~ map Free ps; val vars' = map2 (fn (b, _, mx) => fn (_, T) => (b, SOME T, mx)) vars ps; val binds' = binds |> map (apsnd SOME) |> export_binds params_ctxt ctxt params |> map (apsnd the); val _ = Variable.warn_extra_tfrees fixes_ctxt params_ctxt; val result : stmt = {vars = vars' ~~ params, propss = propss, binds = binds, result_binds = binds'}; in (result, params_ctxt) end; fun prep_statement prep_var prep_propp raw_fixes raw_assumes raw_shows ctxt = let val ((fixes, (assumes, shows), binds), ctxt') = ctxt |> prep_stmt prep_var prep_propp raw_fixes (raw_assumes @ raw_shows) |-> (fn {vars, propss, binds, ...} => fold Variable.bind_term binds #> pair (map #2 vars, chop (length raw_assumes) propss, binds)); val binds' = (Auto_Bind.facts ctxt' (flat shows) @ (case try List.last (flat shows) of NONE => [] | SOME prop => map (apsnd (SOME o Auto_Bind.abs_params prop)) binds)) |> export_binds ctxt' ctxt fixes; val text = Logic.close_prop fixes (flat assumes) (Logic.mk_conjunction_list (flat shows)); val text' = singleton (Variable.export_terms ctxt' ctxt) text; val result : statement = {fixes = fixes, assumes = assumes, shows = shows, result_binds = binds', text = text, result_text = text'}; in (result, ctxt') end; in val cert_stmt = prep_stmt cert_var cert_propp; val read_stmt = prep_stmt read_var read_propp; val cert_statement = prep_statement cert_var cert_propp; val read_statement = prep_statement read_var read_propp; end; (** print context information **) (* local syntax *) val print_syntax = Syntax.print_syntax o syn_of; (* abbreviations *) fun pretty_abbrevs verbose show_globals ctxt = let val space = const_space ctxt; val (constants, global_constants) = apply2 (#constants o Consts.dest) (#consts (rep_data ctxt)); val globals = Symtab.make global_constants; fun add_abbr (_, (_, NONE)) = I | add_abbr (c, (T, SOME t)) = if not show_globals andalso Symtab.defined globals c then I else cons (c, Logic.mk_equals (Const (c, T), t)); val abbrevs = Name_Space.markup_entries verbose ctxt space (fold add_abbr constants []); in if null abbrevs then [] else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)] end; fun print_abbrevs verbose = Pretty.writeln_chunks o pretty_abbrevs verbose true; (* term bindings *) fun pretty_term_bindings ctxt = let val binds = Variable.binds_of ctxt; fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t)); in if Vartab.is_empty binds then [] else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))] end; (* local facts *) fun pretty_local_facts verbose ctxt = let val facts = facts_of ctxt; val props = map #1 (Facts.props facts); val local_facts = (if null props then [] else [("", props)]) @ Facts.dest_static verbose [Global_Theory.facts_of (theory_of ctxt)] facts; in if null local_facts then [] else [Pretty.big_list "local facts:" (map #1 (sort_by (#1 o #2) (map (`(pretty_fact ctxt)) local_facts)))] end; fun print_local_facts verbose ctxt = Pretty.writeln_chunks (pretty_local_facts verbose ctxt); (* named local contexts *) local fun pretty_case (name, (fixes, ((asms, (lets, cs)), ctxt))) = let val prt_name = pretty_name ctxt; val prt_term = Syntax.pretty_term ctxt; fun prt_let (xi, t) = Pretty.block [Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1, Pretty.quote (prt_term t)]; fun prt_asm (a, ts) = Pretty.block (Pretty.breaks ((if a = "" then [] else [prt_name a, Pretty.str ":"]) @ map (Pretty.quote o prt_term) ts)); fun prt_sect _ _ _ [] = [] | prt_sect head sep prt xs = [Pretty.block (Pretty.breaks (head :: flat (separate sep (map (single o prt) xs))))]; in Pretty.block (prt_name name :: Pretty.str ":" :: Pretty.fbrk :: Pretty.fbreaks (prt_sect (Pretty.keyword1 "fix") [] (prt_name o Binding.name_of o fst) fixes @ prt_sect (Pretty.keyword1 "let") [Pretty.keyword2 "and"] prt_let (map_filter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @ (if forall (null o #2) asms then [] else prt_sect (Pretty.keyword1 "assume") [Pretty.keyword2 "and"] prt_asm asms) @ prt_sect (Pretty.str "subcases:") [] (prt_name o fst) cs)) end; in fun pretty_cases ctxt = let val cases = dest_cases NONE ctxt |> map (fn (name, c as Rule_Cases.Case {fixes, ...}) => (name, (fixes, case_result c ctxt))); in if null cases then [] else [Pretty.big_list "cases:" (map pretty_case cases)] end; end; fun print_cases_proof ctxt0 ctxt = let fun trim_name x = if Name.is_internal x then Name.clean x else "_"; val trim_names = map trim_name #> drop_suffix (equal "_"); fun print_case name xs = (case trim_names xs of [] => print_name ctxt name | xs' => enclose "(" ")" (space_implode " " (map (print_name ctxt) (name :: xs')))); fun is_case x t = x = Rule_Cases.case_conclN andalso not (Term.exists_subterm Term.is_Var t); fun indentation depth = prefix (Symbol.spaces (2 * depth)); fun print_proof depth (name, Rule_Cases.Case {fixes, binds, cases, ...}) = let val indent = indentation depth; val head = indent ("case " ^ print_case name (map (Binding.name_of o #1) fixes)); val tail = if null cases then let val concl = if exists (fn ((x, _), SOME t) => is_case x t | _ => false) binds then Rule_Cases.case_conclN else Auto_Bind.thesisN in indent ("then show ?" ^ concl ^ " sorry") end else print_proofs depth cases; in head ^ "\n" ^ tail end and print_proofs 0 [] = "" | print_proofs depth cases = let val indent = indentation depth; val body = map (print_proof (depth + 1)) cases |> separate (indent "next") in if depth = 0 then body @ [indent "qed"] else if length cases = 1 then body else indent "{" :: body @ [indent "}"] end |> cat_lines; in (case print_proofs 0 (dest_cases (SOME ctxt0) ctxt) of "" => "" | s => "Proof outline with cases:\n" ^ Active.sendback_markup_command s) end; (* core context *) val debug = Config.declare_bool ("Proof_Context.debug", \<^here>) (K false); val verbose = Config.declare_bool ("Proof_Context.verbose", \<^here>) (K false); fun pretty_ctxt ctxt = if not (Config.get ctxt debug) then [] else let val prt_term = Syntax.pretty_term ctxt; (*structures*) val {structs, ...} = Syntax_Trans.get_idents ctxt; val prt_structs = if null structs then [] else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 :: Pretty.commas (map Pretty.str structs))]; (*fixes*) fun prt_fix (x, x') = if x = x' then Pretty.str x else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')]; val fixes = filter_out ((Name.is_internal orf member (op =) structs) o #1) (Variable.dest_fixes ctxt); val prt_fixes = if null fixes then [] else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 :: Pretty.commas (map prt_fix fixes))]; (*assumptions*) val prt_assms = (case Assumption.all_prems_of ctxt of [] => [] | prems => [Pretty.big_list "assumptions:" [pretty_fact ctxt ("", prems)]]); in prt_structs @ prt_fixes @ prt_assms end; (* main context *) fun pretty_context ctxt = let val verbose = Config.get ctxt verbose; fun verb f x = if verbose then f (x ()) else []; val prt_term = Syntax.pretty_term ctxt; val prt_typ = Syntax.pretty_typ ctxt; val prt_sort = Syntax.pretty_sort ctxt; (*theory*) val pretty_thy = Pretty.block [Pretty.str "theory:", Pretty.brk 1, Context.pretty_thy (theory_of ctxt)]; (*defaults*) fun prt_atom prt prtT (x, X) = Pretty.block [prt x, Pretty.str " ::", Pretty.brk 1, prtT X]; fun prt_var (x, ~1) = prt_term (Syntax.free x) | prt_var xi = prt_term (Syntax.var xi); fun prt_varT (x, ~1) = prt_typ (TFree (x, [])) | prt_varT xi = prt_typ (TVar (xi, [])); val prt_defT = prt_atom prt_var prt_typ; val prt_defS = prt_atom prt_varT prt_sort; val (types, sorts) = Variable.constraints_of ctxt; in verb single (K pretty_thy) @ pretty_ctxt ctxt @ verb (pretty_abbrevs true false) (K ctxt) @ verb pretty_term_bindings (K ctxt) @ verb (pretty_local_facts true) (K ctxt) @ verb pretty_cases (K ctxt) @ verb single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @ verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts))) end; end; val show_abbrevs = Proof_Context.show_abbrevs; diff --git a/src/Pure/Isar/specification.ML b/src/Pure/Isar/specification.ML --- a/src/Pure/Isar/specification.ML +++ b/src/Pure/Isar/specification.ML @@ -1,478 +1,478 @@ (* Title: Pure/Isar/specification.ML Author: Makarius Derived local theory specifications --- with type-inference and toplevel polymorphism. *) signature SPECIFICATION = sig val read_props: string list -> (binding * string option * mixfix) list -> Proof.context -> term list * Proof.context val check_spec_open: (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> term list -> term -> Proof.context -> ((binding * typ option * mixfix) list * string list * (string -> Position.T list) * term) * Proof.context val read_spec_open: (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> string list -> string -> Proof.context -> ((binding * typ option * mixfix) list * string list * (string -> Position.T list) * term) * Proof.context type multi_specs = ((Attrib.binding * term) * term list * (binding * typ option * mixfix) list) list type multi_specs_cmd = ((Attrib.binding * string) * string list * (binding * string option * mixfix) list) list val check_multi_specs: (binding * typ option * mixfix) list -> multi_specs -> Proof.context -> (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context val read_multi_specs: (binding * string option * mixfix) list -> multi_specs_cmd -> Proof.context -> (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context val axiomatization: (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> term list -> (Attrib.binding * term) list -> theory -> (term list * thm list) * theory val axiomatization_cmd: (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> string list -> (Attrib.binding * string) list -> theory -> (term list * thm list) * theory val axiom: Attrib.binding * term -> theory -> thm * theory val definition: (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term list -> Attrib.binding * term -> local_theory -> (term * (string * thm)) * local_theory val definition': (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term list -> Attrib.binding * term -> bool -> local_theory -> (term * (string * thm)) * local_theory val definition_cmd: (binding * string option * mixfix) option -> (binding * string option * mixfix) list -> string list -> Attrib.binding * string -> bool -> local_theory -> (term * (string * thm)) * local_theory val abbreviation: Syntax.mode -> (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term -> bool -> local_theory -> local_theory val abbreviation_cmd: Syntax.mode -> (binding * string option * mixfix) option -> (binding * string option * mixfix) list -> string -> bool -> local_theory -> local_theory val alias: binding * string -> local_theory -> local_theory val alias_cmd: binding * (xstring * Position.T) -> local_theory -> local_theory val type_alias: binding * string -> local_theory -> local_theory val type_alias_cmd: binding * (xstring * Position.T) -> local_theory -> local_theory val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val theorems: string -> (Attrib.binding * Attrib.thms) list -> (binding * typ option * mixfix) list -> bool -> local_theory -> (string * thm list) list * local_theory val theorems_cmd: string -> (Attrib.binding * (Facts.ref * Token.src list) list) list -> (binding * string option * mixfix) list -> bool -> local_theory -> (string * thm list) list * local_theory val theorem: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> string list -> Element.context_i list -> Element.statement_i -> bool -> local_theory -> Proof.state val theorem_cmd: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> (xstring * Position.T) list -> Element.context list -> Element.statement -> bool -> local_theory -> Proof.state val schematic_theorem: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> string list -> Element.context_i list -> Element.statement_i -> bool -> local_theory -> Proof.state val schematic_theorem_cmd: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> (xstring * Position.T) list -> Element.context list -> Element.statement -> bool -> local_theory -> Proof.state end; structure Specification: SPECIFICATION = struct (* prepare propositions *) fun read_props raw_props raw_fixes ctxt = let val (_, ctxt1) = ctxt |> Proof_Context.add_fixes_cmd raw_fixes; val props1 = map (Syntax.parse_prop ctxt1) raw_props; val (props2, ctxt2) = ctxt1 |> fold_map Variable.fix_dummy_patterns props1; val props3 = Syntax.check_props ctxt2 props2; val ctxt3 = ctxt2 |> fold Variable.declare_term props3; in (props3, ctxt3) end; (* prepare specification *) fun get_positions ctxt x = let fun get Cs (Const ("_type_constraint_", C) $ t) = get (C :: Cs) t | get Cs (Free (y, T)) = if x = y then map_filter Term_Position.decode_positionT (T :: map (Type.constraint_type ctxt) Cs) else [] | get _ (t $ u) = get [] t @ get [] u | get _ (Abs (_, _, t)) = get [] t | get _ _ = []; in get [] end; local fun prep_decls prep_var raw_vars ctxt = let val (vars, ctxt') = fold_map prep_var raw_vars ctxt; val (xs, ctxt'') = ctxt' |> Context_Position.set_visible false |> Proof_Context.add_fixes vars ||> Context_Position.restore_visible ctxt'; val _ = Context_Position.reports ctxt'' (map (Binding.pos_of o #1) vars ~~ map (Variable.markup_entity_def ctxt'' ##> Properties.remove Markup.kindN) xs); in ((vars, xs), ctxt'') end; fun close_form ctxt ys prems concl = let val xs = rev (fold (Variable.add_free_names ctxt) (prems @ [concl]) (rev ys)); val pos_props = Logic.strip_imp_concl concl :: Logic.strip_imp_prems concl @ prems; fun get_pos x = maps (get_positions ctxt x) pos_props; val _ = Context_Position.reports ctxt (maps (Syntax_Phases.reports_of_scope o get_pos) xs); in Logic.close_prop_constraint (Variable.default_type ctxt) (xs ~~ xs) prems concl end; fun dummy_frees ctxt xs tss = let val names = Variable.names_of ((fold o fold) Variable.declare_term tss ctxt) |> fold Name.declare xs; val (tss', _) = (fold_map o fold_map) Term.free_dummy_patterns tss names; in tss' end; fun prep_spec_open prep_var parse_prop raw_vars raw_params raw_prems raw_concl ctxt = let val ((vars, xs), vars_ctxt) = prep_decls prep_var raw_vars ctxt; val (ys, params_ctxt) = vars_ctxt |> fold_map prep_var raw_params |-> Proof_Context.add_fixes; val props = map (parse_prop params_ctxt) (raw_concl :: raw_prems) |> singleton (dummy_frees params_ctxt (xs @ ys)); val concl :: prems = Syntax.check_props params_ctxt props; val spec = Logic.list_implies (prems, concl); val spec_ctxt = Variable.declare_term spec params_ctxt; fun get_pos x = maps (get_positions spec_ctxt x) props; in ((vars, xs, get_pos, spec), spec_ctxt) end; fun prep_specs prep_var parse_prop prep_att raw_vars raw_specss ctxt = let val ((vars, xs), vars_ctxt) = prep_decls prep_var raw_vars ctxt; val propss0 = raw_specss |> map (fn ((_, raw_concl), raw_prems, raw_params) => let val (ys, ctxt') = vars_ctxt |> fold_map prep_var raw_params |-> Proof_Context.add_fixes in (ys, map (pair ctxt') (raw_concl :: raw_prems)) end); val props = burrow (grouped 10 Par_List.map_independent (uncurry parse_prop)) (map #2 propss0) |> dummy_frees vars_ctxt xs |> map2 (fn (ys, _) => fn concl :: prems => close_form vars_ctxt ys prems concl) propss0; val specs = Syntax.check_props vars_ctxt props; val specs_ctxt = vars_ctxt |> fold Variable.declare_term specs; val ps = specs_ctxt |> fold_map Proof_Context.inferred_param xs |> fst; val params = map2 (fn (b, _, mx) => fn (_, T) => ((b, T), mx)) vars ps; val name_atts: Attrib.binding list = map (fn ((name, atts), _) => (name, map (prep_att ctxt) atts)) (map #1 raw_specss); in ((params, name_atts ~~ specs), specs_ctxt) end; in val check_spec_open = prep_spec_open Proof_Context.cert_var (K I); val read_spec_open = prep_spec_open Proof_Context.read_var Syntax.parse_prop; type multi_specs = ((Attrib.binding * term) * term list * (binding * typ option * mixfix) list) list; type multi_specs_cmd = ((Attrib.binding * string) * string list * (binding * string option * mixfix) list) list; fun check_multi_specs xs specs = prep_specs Proof_Context.cert_var (K I) (K I) xs specs; fun read_multi_specs xs specs = prep_specs Proof_Context.read_var Syntax.parse_prop Attrib.check_src xs specs; end; (* axiomatization -- within global theory *) fun gen_axioms prep_stmt prep_att raw_decls raw_fixes raw_prems raw_concls thy = let (*specification*) val ({vars, propss = [prems, concls], ...}, vars_ctxt) = Proof_Context.init_global thy |> prep_stmt (raw_decls @ raw_fixes) ((map o map) (rpair []) [raw_prems, map snd raw_concls]); val (decls, fixes) = chop (length raw_decls) vars; val frees = rev ((fold o fold) (Variable.add_frees vars_ctxt) [prems, concls] []) |> map (fn (x, T) => (x, Free (x, T))); val close = Logic.close_prop (map #2 fixes @ frees) prems; val specs = map ((apsnd o map) (prep_att vars_ctxt) o fst) raw_concls ~~ map close concls; val spec_name = Binding.conglomerate (if null decls then map (#1 o #1) specs else map (#1 o #1) decls); (*consts*) val (consts, consts_thy) = thy |> fold_map (fn ((b, _, mx), (_, t)) => Theory.specify_const ((b, Term.type_of t), mx)) decls; val subst = Term.subst_atomic (map (#2 o #2) decls ~~ consts); (*axioms*) val (axioms, axioms_thy) = (specs, consts_thy) |-> fold_map (fn ((b, atts), prop) => Thm.add_axiom_global (b, subst prop) #>> (fn (_, th) => ((b, atts), [([th], [])]))); (*facts*) val (facts, facts_lthy) = axioms_thy |> Named_Target.theory_init |> Spec_Rules.add spec_name Spec_Rules.Unknown consts (maps (maps #1 o #2) axioms) |> Local_Theory.notes axioms; in ((consts, map (the_single o #2) facts), Local_Theory.exit_global facts_lthy) end; val axiomatization = gen_axioms Proof_Context.cert_stmt (K I); val axiomatization_cmd = gen_axioms Proof_Context.read_stmt Attrib.check_src; fun axiom (b, ax) = axiomatization [] [] [] [(b, ax)] #>> (hd o snd); (* definition *) fun gen_def prep_spec prep_att raw_var raw_params raw_prems ((a, raw_atts), raw_spec) int lthy = let val atts = map (prep_att lthy) raw_atts; val ((vars, xs, get_pos, spec), _) = lthy |> prep_spec (the_list raw_var) raw_params raw_prems raw_spec; val (((x, T), rhs), prove) = Local_Defs.derived_def lthy get_pos {conditional = true} spec; val _ = Name.reject_internal (x, []); val (b, mx) = (case (vars, xs) of ([], []) => (Binding.make (x, (case get_pos x of [] => Position.none | p :: _ => p)), NoSyn) | ([(b, _, mx)], [y]) => if x = y then (b, mx) else error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b))); val const_name = Local_Theory.full_name lthy b; val name = Thm.def_binding_optional b a; val ((lhs, (_, raw_th)), lthy2) = lthy |> Local_Theory.define_internal ((b, mx), ((Binding.suffix_name "_raw" name, []), rhs)); val th = prove lthy2 raw_th; val lthy3 = lthy2 |> Spec_Rules.add name Spec_Rules.equational [lhs] [th]; val ([(def_name, [th'])], lthy4) = lthy3 |> Local_Theory.notes [((name, atts), [([th], [])])]; val lthy5 = lthy4 |> Code.declare_default_eqns [(th', true)]; val lhs' = Morphism.term (Local_Theory.target_morphism lthy5) lhs; val _ = Proof_Display.print_consts int (Position.thread_data ()) lthy5 - (member (op =) (Term.add_frees lhs' [])) [(x, T)]; + (Term_Subst.Frees.defined (Term_Subst.Frees.build (Term_Subst.add_frees lhs'))) [(x, T)]; in ((lhs, (def_name, th')), lthy5) end; val definition' = gen_def check_spec_open (K I); fun definition xs ys As B = definition' xs ys As B false; val definition_cmd = gen_def read_spec_open Attrib.check_src; (* abbreviation *) fun gen_abbrev prep_spec mode raw_var raw_params raw_spec int lthy = let val lthy1 = lthy |> Proof_Context.set_syntax_mode mode; val ((vars, xs, get_pos, spec), _) = lthy |> Proof_Context.set_mode Proof_Context.mode_abbrev |> prep_spec (the_list raw_var) raw_params [] raw_spec; val ((x, T), rhs) = Local_Defs.abs_def (#2 (Local_Defs.cert_def lthy1 get_pos spec)); val _ = Name.reject_internal (x, []); val (b, mx) = (case (vars, xs) of ([], []) => (Binding.make (x, (case get_pos x of [] => Position.none | p :: _ => p)), NoSyn) | ([(b, _, mx)], [y]) => if x = y then (b, mx) else error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b))); val lthy2 = lthy1 |> Local_Theory.abbrev mode ((b, mx), rhs) |> snd |> Proof_Context.restore_syntax_mode lthy; val _ = Proof_Display.print_consts int (Position.thread_data ()) lthy2 (K false) [(x, T)]; in lthy2 end; val abbreviation = gen_abbrev check_spec_open; val abbreviation_cmd = gen_abbrev read_spec_open; (* alias *) fun gen_alias decl check (b, arg) lthy = let val (c, reports) = check {proper = true, strict = false} lthy arg; val _ = Context_Position.reports lthy reports; in decl b c lthy end; val alias = gen_alias Local_Theory.const_alias (K (K (fn c => (c, [])))); val alias_cmd = gen_alias Local_Theory.const_alias (fn flags => fn ctxt => fn (c, pos) => apfst (#1 o dest_Const) (Proof_Context.check_const flags ctxt (c, [pos]))); val type_alias = gen_alias Local_Theory.type_alias (K (K (fn c => (c, [])))); val type_alias_cmd = gen_alias Local_Theory.type_alias (apfst (#1 o dest_Type) ooo Proof_Context.check_type_name); (* notation *) local fun gen_type_notation prep_type add mode args lthy = lthy |> Local_Theory.type_notation add mode (map (apfst (prep_type lthy)) args); fun gen_notation prep_const add mode args lthy = lthy |> Local_Theory.notation add mode (map (apfst (prep_const lthy)) args); in val type_notation = gen_type_notation (K I); val type_notation_cmd = gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false}); val notation = gen_notation (K I); val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false}); end; (* fact statements *) local fun gen_theorems prep_fact prep_att add_fixes kind raw_facts raw_fixes int lthy = let val facts = raw_facts |> map (fn ((name, atts), bs) => ((name, map (prep_att lthy) atts), bs |> map (fn (b, more_atts) => (prep_fact lthy b, map (prep_att lthy) more_atts)))); val (_, ctxt') = add_fixes raw_fixes lthy; val facts' = facts |> Attrib.partial_evaluation ctxt' |> Attrib.transform_facts (Proof_Context.export_morphism ctxt' lthy); val (res, lthy') = lthy |> Local_Theory.notes_kind kind facts'; val _ = Proof_Display.print_results int (Position.thread_data ()) lthy' ((kind, ""), res); in (res, lthy') end; in val theorems = gen_theorems (K I) (K I) Proof_Context.add_fixes; val theorems_cmd = gen_theorems Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd; end; (* complex goal statements *) local fun prep_statement prep_att prep_stmt raw_elems raw_stmt ctxt = let val (stmt, elems_ctxt) = prep_stmt raw_elems raw_stmt ctxt; val prems = Assumption.local_prems_of elems_ctxt ctxt; val stmt_ctxt = fold (fold (Proof_Context.augment o fst) o snd) stmt elems_ctxt; in (case raw_stmt of Element.Shows _ => let val stmt' = Attrib.map_specs (map prep_att) stmt in (([], prems, stmt', NONE), stmt_ctxt) end | Element.Obtains raw_obtains => let val asms_ctxt = stmt_ctxt |> fold (fn ((name, _), asm) => snd o Proof_Context.add_assms Assumption.assume_export [((name, [Context_Rules.intro_query NONE]), asm)]) stmt; val that = Assumption.local_prems_of asms_ctxt stmt_ctxt; val ([(_, that')], that_ctxt) = asms_ctxt |> Proof_Context.set_stmt true |> Proof_Context.note_thmss "" [((Binding.name Auto_Bind.thatN, []), [(that, [])])] ||> Proof_Context.restore_stmt asms_ctxt; val stmt' = [(Binding.empty_atts, [(#2 (#1 (Obtain.obtain_thesis ctxt)), [])])]; in ((Obtain.obtains_attribs raw_obtains, prems, stmt', SOME that'), that_ctxt) end) end; fun gen_theorem schematic bundle_includes prep_att prep_stmt long kind before_qed after_qed (name, raw_atts) raw_includes raw_elems raw_concl int lthy = let val _ = Local_Theory.assert lthy; val elems = raw_elems |> map (Element.map_ctxt_attrib (prep_att lthy)); val ((more_atts, prems, stmt, facts), goal_ctxt) = lthy |> bundle_includes raw_includes |> prep_statement (prep_att lthy) prep_stmt elems raw_concl; val atts = more_atts @ map (prep_att lthy) raw_atts; val pos = Position.thread_data (); fun after_qed' results goal_ctxt' = let val results' = burrow (map (Goal.norm_result lthy) o Proof_Context.export goal_ctxt' lthy) results; val (res, lthy') = if forall (Binding.is_empty_atts o fst) stmt then (map (pair "") results', lthy) else Local_Theory.notes_kind kind (map2 (fn (b, _) => fn ths => (b, [(ths, [])])) stmt results') lthy; val lthy'' = if Binding.is_empty_atts (name, atts) then (Proof_Display.print_results int pos lthy' ((kind, ""), res); lthy') else let val ([(res_name, _)], lthy'') = Local_Theory.notes_kind kind [((name, atts), [(maps #2 res, [])])] lthy'; val _ = Proof_Display.print_results int pos lthy' ((kind, res_name), res); in lthy'' end; in after_qed results' lthy'' end; val prems_name = if long then Auto_Bind.assmsN else Auto_Bind.thatN; in goal_ctxt |> not (null prems) ? (Proof_Context.note_thmss "" [((Binding.name prems_name, []), [(prems, [])])] #> snd) |> Proof.theorem before_qed after_qed' (map snd stmt) |> (case facts of NONE => I | SOME ths => Proof.refine_insert ths) |> tap (fn state => not schematic andalso Proof.schematic_goal state andalso error "Illegal schematic goal statement") end; in val theorem = gen_theorem false Bundle.includes (K I) Expression.cert_statement; val theorem_cmd = gen_theorem false Bundle.includes_cmd Attrib.check_src Expression.read_statement; val schematic_theorem = gen_theorem true Bundle.includes (K I) Expression.cert_statement; val schematic_theorem_cmd = gen_theorem true Bundle.includes_cmd Attrib.check_src Expression.read_statement; end; end; diff --git a/src/Pure/Isar/subgoal.ML b/src/Pure/Isar/subgoal.ML --- a/src/Pure/Isar/subgoal.ML +++ b/src/Pure/Isar/subgoal.ML @@ -1,257 +1,257 @@ (* Title: Pure/Isar/subgoal.ML Author: Makarius Author: Daniel Matichuk, NICTA/UNSW Tactical operations with explicit subgoal focus, based on canonical proof decomposition. The "visible" part of the text within the context is fixed, the remaining goal may be schematic. Isar subgoal command for proof structure within unstructured proof scripts. *) signature SUBGOAL = sig type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list, concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list} val focus_params: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_params_fixed: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_prems: Proof.context -> int -> binding list option -> thm -> focus * thm val focus: Proof.context -> int -> binding list option -> thm -> focus * thm val retrofit: Proof.context -> Proof.context -> (string * cterm) list -> cterm list -> int -> thm -> thm -> thm Seq.seq val FOCUS_PARAMS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_PARAMS_FIXED: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_PREMS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS: (focus -> tactic) -> Proof.context -> int -> tactic val SUBPROOF: (focus -> tactic) -> Proof.context -> int -> tactic val subgoal: Attrib.binding -> Attrib.binding option -> bool * (string option * Position.T) list -> Proof.state -> focus * Proof.state val subgoal_cmd: Attrib.binding -> Attrib.binding option -> bool * (string option * Position.T) list -> Proof.state -> focus * Proof.state end; structure Subgoal: SUBGOAL = struct (* focus *) type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list, concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list}; fun gen_focus (do_prems, do_concl) ctxt i bindings raw_st = let val st = raw_st |> Thm.solve_constraints |> Thm.transfer' ctxt |> Raw_Simplifier.norm_hhf_protect ctxt; val ((schematic_types, [st']), ctxt1) = Variable.importT [st] ctxt; val ((params, goal), ctxt2) = Variable.focus_cterm bindings (Thm.cprem_of st' i) ctxt1; val (asms, concl) = if do_prems then (Drule.strip_imp_prems goal, Drule.strip_imp_concl goal) else ([], goal); val text = asms @ (if do_concl then [concl] else []); val ((_, inst), ctxt3) = Variable.import_inst true (map Thm.term_of text) ctxt2; val schematic_terms = Term_Subst.Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt3 t)) inst []; - val schematics = (schematic_types, schematic_terms); + val schematics = (Term_Subst.TVars.dest schematic_types, schematic_terms); val asms' = map (Thm.instantiate_cterm schematics) asms; val concl' = Thm.instantiate_cterm schematics concl; val (prems, context) = Assumption.add_assumes asms' ctxt3; in ({context = context, params = params, prems = prems, asms = asms', concl = concl', schematics = schematics}, Goal.init concl') end; val focus_params = gen_focus (false, false); val focus_params_fixed = gen_focus (false, true); val focus_prems = gen_focus (true, false); val focus = gen_focus (true, true); (* lift and retrofit *) (* B [?'b, ?y] ---------------- B ['b, y params] *) fun lift_import idx params th ctxt = let val ((_, [th']), ctxt') = Variable.importT [th] ctxt; val Ts = map Thm.typ_of_cterm params; val ts = map Thm.term_of params; val prop = Thm.full_prop_of th'; - val concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []; - val vars = rev (Term.add_vars prop []); + val concl_vars = Term_Subst.Vars.build (Term_Subst.add_vars (Logic.strip_imp_concl prop)); + val vars = build_rev (Term.add_vars prop); val (ys, ctxt'') = Variable.variant_fixes (map (Name.clean o #1 o #1) vars) ctxt'; fun var_inst v y = let val ((x, i), T) = v; val (U, args) = - if member (op =) concl_vars v then (T, []) + if Term_Subst.Vars.defined concl_vars v then (T, []) else (Ts ---> T, ts); val u = Free (y, U); in ((Var v, list_comb (u, args)), (u, Var ((x, i + idx), U))) end; val (inst1, inst2) = split_list (map (apply2 (apply2 (Thm.cterm_of ctxt))) (map2 var_inst vars ys)); val th'' = Thm.instantiate ([], map (apfst (Term.dest_Var o Thm.term_of)) inst1) th'; in ((inst2, th''), ctxt'') end; (* [x, A x] : B x \ C ------------------ [\x. A x \ B x] : C *) fun lift_subgoals ctxt params asms th = let fun lift ct = fold_rev (Thm.all_name ctxt) params (Drule.list_implies (asms, ct)); val unlift = fold (Thm.elim_implies o Thm.assume) asms o Drule.forall_elim_list (map #2 params) o Thm.assume; val subgoals = map lift (Drule.strip_imp_prems (Thm.cprop_of th)); val th' = fold (Thm.elim_implies o unlift) subgoals th; in (subgoals, th') end; fun retrofit ctxt1 ctxt0 params asms i st1 st0 = let val idx = Thm.maxidx_of st0 + 1; val ps = map #2 params; val ((subgoal_inst, st2), ctxt2) = lift_import idx ps st1 ctxt1; val (subgoals, st3) = lift_subgoals ctxt2 params asms st2; val result = st3 |> Goal.conclude |> Drule.implies_intr_list asms |> Drule.forall_intr_list ps |> Drule.implies_intr_list subgoals |> fold_rev (Thm.forall_intr o #1) subgoal_inst |> fold (Thm.forall_elim o #2) subgoal_inst |> Thm.adjust_maxidx_thm idx |> singleton (Variable.export ctxt2 ctxt0); in Thm.bicompose (SOME ctxt0) {flatten = true, match = false, incremented = false} (false, result, Thm.nprems_of st1) i st0 end; (* tacticals *) fun GEN_FOCUS flags tac ctxt i st = if Thm.nprems_of st < i then Seq.empty else let val (args as {context = ctxt', params, asms, ...}, st') = gen_focus flags (ctxt |> Variable.set_bound_focus true) i NONE st; in Seq.lifts (retrofit ctxt' ctxt params asms i) (tac args st') st end; val FOCUS_PARAMS = GEN_FOCUS (false, false); val FOCUS_PARAMS_FIXED = GEN_FOCUS (false, true); val FOCUS_PREMS = GEN_FOCUS (true, false); val FOCUS = GEN_FOCUS (true, true); fun SUBPROOF tac ctxt = FOCUS (Seq.map (Goal.check_finished ctxt) oo tac) ctxt; (* Isar subgoal command *) local fun param_bindings ctxt (param_suffix, raw_param_specs) st = let val _ = if Thm.no_prems st then error "No subgoals!" else (); val subgoal = #1 (Logic.dest_implies (Thm.prop_of st)); val subgoal_params = map (apfst (Name.internal o Name.clean)) (Term.strip_all_vars subgoal) |> Term.variant_frees subgoal |> map #1; val n = length subgoal_params; val m = length raw_param_specs; val _ = m <= n orelse error ("Excessive subgoal parameter specification" ^ Position.here_list (map snd (drop n raw_param_specs))); val param_specs = raw_param_specs |> map (fn (NONE, _) => NONE | (SOME x, pos) => let val b = #1 (#1 (Proof_Context.cert_var (Binding.make (x, pos), NONE, NoSyn) ctxt)); val _ = Variable.check_name b; in SOME b end) |> param_suffix ? append (replicate (n - m) NONE); fun bindings (SOME x :: xs) (_ :: ys) = x :: bindings xs ys | bindings (NONE :: xs) (y :: ys) = Binding.name y :: bindings xs ys | bindings _ ys = map Binding.name ys; in bindings param_specs subgoal_params end; fun gen_subgoal prep_atts raw_result_binding raw_prems_binding param_specs state = let val _ = Proof.assert_backward state; val state1 = state |> Proof.refine_insert []; val {context = ctxt, facts = facts, goal = st} = Proof.raw_goal state1; val result_binding = apsnd (map (prep_atts ctxt)) raw_result_binding; val (prems_binding, do_prems) = (case raw_prems_binding of SOME (b, raw_atts) => ((b, map (prep_atts ctxt) raw_atts), true) | NONE => (Binding.empty_atts, false)); val (subgoal_focus, _) = (if do_prems then focus else focus_params_fixed) ctxt 1 (SOME (param_bindings ctxt param_specs st)) st; fun after_qed (ctxt'', [[result]]) = Proof.end_block #> (fn state' => let val ctxt' = Proof.context_of state'; val results' = Proof_Context.export ctxt'' ctxt' (Conjunction.elim_conjunctions result); in state' |> Proof.refine_primitive (fn _ => fn _ => retrofit ctxt'' ctxt' (#params subgoal_focus) (#asms subgoal_focus) 1 (Goal.protect 0 result) st |> Seq.hd) |> Proof.map_context (#2 o Proof_Context.note_thmss "" [(result_binding, [(results', [])])]) end) #> Proof.reset_facts #> Proof.enter_backward; in state1 |> Proof.enter_forward |> Proof.using_facts [] |> Proof.begin_block |> Proof.map_context (fn _ => #context subgoal_focus |> Proof_Context.note_thmss "" [(prems_binding, [(#prems subgoal_focus, [])])] |> #2) |> Proof.internal_goal (K (K ())) (Proof_Context.get_mode ctxt) true "subgoal" NONE after_qed [] [] [(Binding.empty_atts, [(Thm.term_of (#concl subgoal_focus), [])])] |> #2 |> Proof.using_facts facts |> pair subgoal_focus end; in val subgoal = gen_subgoal Attrib.attribute; val subgoal_cmd = gen_subgoal Attrib.attribute_cmd; end; end; val SUBPROOF = Subgoal.SUBPROOF; diff --git a/src/Pure/PIDE/document.ML b/src/Pure/PIDE/document.ML --- a/src/Pure/PIDE/document.ML +++ b/src/Pure/PIDE/document.ML @@ -1,921 +1,921 @@ (* Title: Pure/PIDE/document.ML Author: Makarius Document as collection of named nodes, each consisting of an editable list of commands, associated with asynchronous execution process. *) signature DOCUMENT = sig val timing: bool Unsynchronized.ref type node_header = {master: string, header: Thy_Header.header, errors: string list} type overlay = Document_ID.command * (string * string list) datatype node_edit = Edits of (Document_ID.command option * Document_ID.command option) list | Deps of node_header | Perspective of bool * Document_ID.command list * overlay list type edit = string * node_edit type state val init_state: state val define_blob: string -> string -> state -> state type blob_digest = {file_node: string, src_path: Path.T, digest: string option} Exn.result type command = {command_id: Document_ID.command, name: string, parents: string list, blobs_digests: blob_digest list, blobs_index: int, tokens: ((int * int) * string) list} val define_command: command -> state -> state val command_exec: state -> string -> Document_ID.command -> Command.exec option val remove_versions: Document_ID.version list -> state -> state val start_execution: state -> state val update: Document_ID.version -> Document_ID.version -> edit list -> string list -> state -> string list * Document_ID.exec list * (Document_ID.command * Document_ID.exec list) list * state val state: unit -> state val change_state: (state -> state) -> unit end; structure Document: DOCUMENT = struct val timing = Unsynchronized.ref false; fun timeit msg e = cond_timeit (! timing) msg e; (** document structure **) fun err_dup kind id = error ("Duplicate " ^ kind ^ ": " ^ Document_ID.print id); fun err_undef kind id = error ("Undefined " ^ kind ^ ": " ^ Document_ID.print id); type node_header = {master: string, header: Thy_Header.header, errors: string list}; type perspective = {required: bool, (*required node*) visible: Inttab.set, (*visible commands*) visible_last: Document_ID.command option, (*last visible command*) overlays: (string * string list) list Inttab.table}; (*command id -> print functions with args*) structure Entries = Linear_Set(type key = Document_ID.command val ord = int_ord); abstype node = Node of {header: node_header, (*master directory, theory header, errors*) keywords: Keyword.keywords option, (*outer syntax keywords*) perspective: perspective, (*command perspective*) entries: Command.exec option Entries.T, (*command entries with executions*) result: (Document_ID.command * Command.eval) option, (*result of last execution*) consolidated: unit lazy} (*consolidated status of eval forks*) and version = Version of node String_Graph.T (*development graph wrt. static imports*) with fun make_node (header, keywords, perspective, entries, result, consolidated) = Node {header = header, keywords = keywords, perspective = perspective, entries = entries, result = result, consolidated = consolidated}; fun map_node f (Node {header, keywords, perspective, entries, result, consolidated}) = make_node (f (header, keywords, perspective, entries, result, consolidated)); fun make_perspective (required, command_ids, overlays) : perspective = {required = required, visible = Inttab.make_set command_ids, visible_last = try List.last command_ids, overlays = Inttab.make_list overlays}; val no_header: node_header = {master = "", header = Thy_Header.make ("", Position.none) [] [], errors = []}; val no_perspective = make_perspective (false, [], []); val empty_node = make_node (no_header, NONE, no_perspective, Entries.empty, NONE, Lazy.value ()); fun is_no_perspective ({required, visible, visible_last, overlays}: perspective) = not required andalso Inttab.is_empty visible andalso is_none visible_last andalso Inttab.is_empty overlays; fun is_empty_node (Node {header, keywords, perspective, entries, result, consolidated}) = header = no_header andalso is_none keywords andalso is_no_perspective perspective andalso Entries.is_empty entries andalso is_none result andalso Lazy.is_finished consolidated; (* basic components *) fun master_directory (Node {header = {master, ...}, ...}) = (case try Url.explode master of SOME (Url.File path) => path | _ => Path.current); fun set_header master header errors = map_node (fn (_, keywords, perspective, entries, result, consolidated) => ({master = master, header = header, errors = errors}, keywords, perspective, entries, result, consolidated)); fun get_header (Node {header, ...}) = header; fun set_keywords keywords = map_node (fn (header, _, perspective, entries, result, consolidated) => (header, keywords, perspective, entries, result, consolidated)); fun get_keywords (Node {keywords, ...}) = keywords; fun read_header node span = let val {header, errors, ...} = get_header node; val _ = if null errors then () else cat_lines errors |> (case Position.id_of (Position.thread_data ()) of NONE => I | SOME id => Protocol_Message.command_positions_yxml id) |> error; val {name = (name, _), imports, ...} = header; val {name = (_, pos), imports = imports', keywords} = Thy_Header.read_tokens Position.none span; val imports'' = (map #1 imports ~~ map #2 imports') handle ListPair.UnequalLengths => imports; in Thy_Header.make (name, pos) imports'' keywords end; fun get_perspective (Node {perspective, ...}) = perspective; fun set_perspective args = map_node (fn (header, keywords, _, entries, result, consolidated) => (header, keywords, make_perspective args, entries, result, consolidated)); val required_node = #required o get_perspective; val visible_command = Inttab.defined o #visible o get_perspective; val visible_last = #visible_last o get_perspective; val visible_node = is_some o visible_last val overlays = Inttab.lookup_list o #overlays o get_perspective; fun map_entries f = map_node (fn (header, keywords, perspective, entries, result, consolidated) => (header, keywords, perspective, f entries, result, consolidated)); fun get_entries (Node {entries, ...}) = entries; fun iterate_entries f = Entries.iterate NONE f o get_entries; fun iterate_entries_after start f (Node {entries, ...}) = (case Entries.get_after entries start of NONE => I | SOME id => Entries.iterate (SOME id) f entries); fun get_result (Node {result, ...}) = result; fun set_result result = map_node (fn (header, keywords, perspective, entries, _, consolidated) => (header, keywords, perspective, entries, result, consolidated)); fun pending_result node = (case get_result node of SOME (_, eval) => not (Command.eval_finished eval) | NONE => false); fun finished_result node = (case get_result node of SOME (_, eval) => Command.eval_finished eval | NONE => false); fun finished_result_theory node = if finished_result node then let val (result_id, eval) = the (get_result node); val st = Command.eval_result_state eval; in SOME (result_id, Toplevel.end_theory Position.none st) handle ERROR _ => NONE end else NONE; val reset_consolidated = map_node (fn (header, keywords, perspective, entries, result, _) => (header, keywords, perspective, entries, result, Lazy.lazy I)); fun commit_consolidated (Node {consolidated, ...}) = (Lazy.force consolidated; Output.status [Markup.markup_only Markup.consolidated]); fun consolidated_node (Node {consolidated, ...}) = Lazy.is_finished consolidated; fun could_consolidate node = not (consolidated_node node) andalso is_some (finished_result_theory node); fun get_node nodes name = String_Graph.get_node nodes name handle String_Graph.UNDEF _ => empty_node; fun default_node name = String_Graph.default_node (name, empty_node); fun update_node name f = default_node name #> String_Graph.map_node name f; (* node edits and associated executions *) type overlay = Document_ID.command * (string * string list); datatype node_edit = Edits of (Document_ID.command option * Document_ID.command option) list | Deps of node_header | Perspective of bool * Document_ID.command list * overlay list; type edit = string * node_edit; val after_entry = Entries.get_after o get_entries; fun lookup_entry node id = (case Entries.lookup (get_entries node) id of NONE => NONE | SOME (exec, _) => exec); fun the_entry node id = (case Entries.lookup (get_entries node) id of NONE => err_undef "command entry" id | SOME (exec, _) => exec); fun assign_entry (command_id, exec) node = if is_none (Entries.lookup (get_entries node) command_id) then node else map_entries (Entries.update (command_id, exec)) node; fun reset_after id entries = (case Entries.get_after entries id of NONE => entries | SOME next => Entries.update (next, NONE) entries); val edit_node = map_entries o fold (fn (id, SOME id2) => Entries.insert_after id (id2, NONE) | (id, NONE) => Entries.delete_after id #> reset_after id); (* version operations *) val empty_version = Version String_Graph.empty; fun nodes_of (Version nodes) = nodes; val node_of = get_node o nodes_of; fun cycle_msg names = "Cyclic dependency of " ^ space_implode " via " (map quote names); fun edit_nodes (name, node_edit) (Version nodes) = Version (case node_edit of Edits edits => update_node name (edit_node edits) nodes | Deps {master, header, errors} => let val imports = map fst (#imports header); val nodes1 = nodes |> default_node name |> fold default_node imports; val nodes2 = nodes1 |> String_Graph.Keys.fold (fn dep => String_Graph.del_edge (dep, name)) (String_Graph.imm_preds nodes1 name); val (nodes3, errors1) = (String_Graph.add_deps_acyclic (name, imports) nodes2, errors) handle String_Graph.CYCLES cs => (nodes2, errors @ map cycle_msg cs); in String_Graph.map_node name (set_header master header errors1) nodes3 end | Perspective perspective => update_node name (set_perspective perspective) nodes); fun update_keywords name nodes = nodes |> String_Graph.map_node name (fn node => if is_empty_node node then node else let val {master, header, errors} = get_header node; val imports_keywords = map_filter (get_keywords o get_node nodes o #1) (#imports header); val keywords = Library.foldl Keyword.merge_keywords (Session.get_keywords (), imports_keywords); val (keywords', errors') = (Keyword.add_keywords (#keywords header) keywords, errors) handle ERROR msg => (keywords, if member (op =) errors msg then errors else errors @ [msg]); in node |> set_header master header errors' |> set_keywords (SOME keywords') end); fun edit_keywords edits (Version nodes) = Version (fold update_keywords (String_Graph.all_succs nodes (map_filter (fn (a, Deps _) => SOME a | _ => NONE) edits)) nodes); fun suppressed_node (nodes: node String_Graph.T) (name, node) = String_Graph.is_maximal nodes name andalso is_empty_node node; fun put_node (name, node) (Version nodes) = let val nodes1 = update_node name (K node) nodes; val nodes2 = if suppressed_node nodes1 (name, node) then String_Graph.del_node name nodes1 else nodes1; in Version nodes2 end; end; (** main state -- document structure and execution process **) type blob_digest = {file_node: string, src_path: Path.T, digest: string option} Exn.result; type execution = {version_id: Document_ID.version, (*static version id*) execution_id: Document_ID.execution, (*dynamic execution id*) delay_request: unit future, (*pending event timer request*) parallel_prints: Command.exec list}; (*parallel prints for early execution*) val no_execution: execution = {version_id = Document_ID.none, execution_id = Document_ID.none, delay_request = Future.value (), parallel_prints = []}; fun new_execution version_id delay_request parallel_prints : execution = {version_id = version_id, execution_id = Execution.start (), delay_request = delay_request, parallel_prints = parallel_prints}; abstype state = State of {versions: version Inttab.table, (*version id -> document content*) blobs: (SHA1.digest * string list) Symtab.table, (*raw digest -> digest, lines*) commands: (string * blob_digest list * int * Token.T list lazy) Inttab.table, (*command id -> name, inlined files, token index of files, command span*) execution: execution} (*current execution process*) with fun make_state (versions, blobs, commands, execution) = State {versions = versions, blobs = blobs, commands = commands, execution = execution}; fun map_state f (State {versions, blobs, commands, execution}) = make_state (f (versions, blobs, commands, execution)); val init_state = make_state (Inttab.make [(Document_ID.none, empty_version)], Symtab.empty, Inttab.empty, no_execution); (* document versions *) fun parallel_prints_node node = iterate_entries (fn (_, opt_exec) => fn rev_result => (case opt_exec of SOME (eval, prints) => (case filter Command.parallel_print prints of [] => SOME rev_result | prints' => SOME ((eval, prints') :: rev_result)) | NONE => NONE)) node [] |> rev; fun define_version version_id version assigned_nodes = map_state (fn (versions, blobs, commands, {delay_request, ...}) => let val version' = fold put_node assigned_nodes version; val versions' = Inttab.update_new (version_id, version') versions handle Inttab.DUP dup => err_dup "document version" dup; val parallel_prints = maps (parallel_prints_node o #2) assigned_nodes; val execution' = new_execution version_id delay_request parallel_prints; in (versions', blobs, commands, execution') end); fun the_version (State {versions, ...}) version_id = (case Inttab.lookup versions version_id of NONE => err_undef "document version" version_id | SOME version => version); fun delete_version version_id versions = Inttab.delete version_id versions handle Inttab.UNDEF _ => err_undef "document version" version_id; (* inlined files *) fun define_blob digest text = map_state (fn (versions, blobs, commands, execution) => let val blobs' = Symtab.update (digest, (SHA1.fake digest, split_lines text)) blobs in (versions, blobs', commands, execution) end); fun the_blob (State {blobs, ...}) digest = (case Symtab.lookup blobs digest of NONE => error ("Undefined blob: " ^ digest) | SOME content => content); fun resolve_blob state (blob_digest: blob_digest) = blob_digest |> Exn.map_res (fn {file_node, src_path, digest} => {file_node = file_node, src_path = src_path, content = Option.map (the_blob state) digest}); fun blob_reports pos (blob_digest: blob_digest) = (case blob_digest of Exn.Res {file_node, ...} => [(pos, Markup.path file_node)] | _ => []); (* commands *) type command = {command_id: Document_ID.command, name: string, parents: string list, blobs_digests: blob_digest list, blobs_index: int, tokens: ((int * int) * string) list}; fun define_command {command_id, name, parents, blobs_digests, blobs_index, tokens} = map_state (fn (versions, blobs, commands, execution) => let val id = Document_ID.print command_id; val span = Lazy.lazy_name "Document.define_command" (fn () => Position.setmp_thread_data (Position.id_only id) (fn () => let val (tokens, _) = fold_map Token.make tokens (Position.id id); val imports = if name = Thy_Header.theoryN then (#imports (Thy_Header.read_tokens Position.none tokens) handle ERROR _ => []) else []; val _ = if length parents = length imports then (parents, imports) |> ListPair.app (fn (parent, (_, pos)) => let val markup = (case Thy_Info.lookup_theory parent of SOME thy => Theory.get_markup thy | NONE => (case try Url.explode parent of NONE => Markup.path parent | SOME (Url.File path) => Markup.path (Path.implode path) | SOME _ => Markup.path parent)) in Position.report pos markup end) else (); val _ = if blobs_index < 0 then (*inlined errors*) map_filter Exn.get_exn blobs_digests |> List.app (Output.error_message o Runtime.exn_message) else (*auxiliary files*) let val pos = Token.pos_of (nth tokens blobs_index) in Position.reports (maps (blob_reports pos) blobs_digests) end; in tokens end) ()); val commands' = Inttab.update_new (command_id, (name, blobs_digests, blobs_index, span)) commands handle Inttab.DUP dup => err_dup "command" dup; in (versions, blobs, commands', execution) end); fun the_command (State {commands, ...}) command_id = (case Inttab.lookup commands command_id of NONE => err_undef "command" command_id | SOME command => command); val the_command_name = #1 oo the_command; (* execution *) fun get_execution (State {execution, ...}) = execution; fun get_execution_version state = the_version state (#version_id (get_execution state)); fun command_exec state node_name command_id = let val version = get_execution_version state; val node = get_node (nodes_of version) node_name; in the_entry node command_id end; end; (* remove_versions *) fun remove_versions version_ids state = state |> map_state (fn (versions, _, _, execution) => let val _ = member (op =) version_ids (#version_id execution) andalso error ("Attempt to remove execution version " ^ Document_ID.print (#version_id execution)); val versions' = fold delete_version version_ids versions; val commands' = - (versions', Inttab.empty) |-> + Inttab.build (versions' |> Inttab.fold (fn (_, version) => nodes_of version |> String_Graph.fold (fn (_, (node, _)) => node |> iterate_entries (fn ((_, command_id), _) => - SOME o Inttab.insert (K true) (command_id, the_command state command_id)))); + SOME o Inttab.insert (K true) (command_id, the_command state command_id))))); val blobs' = - (commands', Symtab.empty) |-> + Symtab.build (commands' |> Inttab.fold (fn (_, (_, blobs, _, _)) => blobs |> - fold (fn Exn.Res {digest = SOME b, ...} => Symtab.update (b, the_blob state b) | _ => I)); + fold (fn Exn.Res {digest = SOME b, ...} => Symtab.update (b, the_blob state b) | _ => I))); in (versions', blobs', commands', execution) end); (* document execution *) fun make_required nodes = let fun all_preds P = String_Graph.fold (fn (a, (node, _)) => P node ? cons a) nodes [] |> String_Graph.all_preds nodes |> Symtab.make_set; val all_visible = all_preds visible_node; val all_required = all_preds required_node; in Symtab.fold (fn (a, ()) => exists (Symtab.defined all_visible) (String_Graph.immediate_succs nodes a) ? Symtab.update (a, ())) all_visible all_required end; fun start_execution state = state |> map_state (fn (versions, blobs, commands, execution) => timeit "Document.start_execution" (fn () => let val {version_id, execution_id, delay_request, parallel_prints} = execution; val delay = seconds (Options.default_real "editor_execution_delay"); val _ = Future.cancel delay_request; val delay_request' = Event_Timer.future {physical = true} (Time.now () + delay); val delay = Future.task_of delay_request'; val parallel_prints' = parallel_prints |> map_filter (Command.exec_parallel_prints execution_id [delay]); fun finished_import (name, (node, _)) = finished_result node orelse is_some (Thy_Info.lookup_theory name); val nodes = nodes_of (the_version state version_id); val required = make_required nodes; val _ = nodes |> String_Graph.schedule (fn deps => fn (name, node) => if Symtab.defined required name orelse visible_node node orelse pending_result node then let val future = (singleton o Future.forks) {name = "theory:" ^ name, group = SOME (Future.new_group NONE), deps = delay :: Execution.active_tasks name @ maps (the_list o #2 o #2) deps, pri = 0, interrupts = false} (fn () => (Execution.worker_task_active true name; if forall finished_import deps then iterate_entries (fn (_, opt_exec) => fn () => (case opt_exec of SOME exec => if Execution.is_running execution_id then SOME (Command.exec execution_id exec) else NONE | NONE => NONE)) node () else (); Execution.worker_task_active false name) handle exn => (Output.system_message (Runtime.exn_message exn); Execution.worker_task_active false name; Exn.reraise exn)); in (node, SOME (Future.task_of future)) end else (node, NONE)); val execution' = {version_id = version_id, execution_id = execution_id, delay_request = delay_request', parallel_prints = parallel_prints'}; in (versions, blobs, commands, execution') end)); (** document update **) (* exec state assignment *) type assign_update = Command.exec option Inttab.table; (*command id -> exec*) val assign_update_empty: assign_update = Inttab.empty; fun assign_update_defined (tab: assign_update) command_id = Inttab.defined tab command_id; fun assign_update_change entry (tab: assign_update) = Inttab.update entry tab; fun assign_update_apply (tab: assign_update) node = Inttab.fold assign_entry tab node; fun assign_update_new upd (tab: assign_update) = Inttab.update_new upd tab handle Inttab.DUP dup => err_dup "exec state assignment" dup; fun assign_update_result (tab: assign_update) = Inttab.fold (fn (command_id, exec) => cons (command_id, Command.exec_ids exec)) tab []; (* update *) local fun init_theory deps node span = let val master_dir = master_directory node; val header = read_header node span; val imports = #imports header; fun maybe_eval_result eval = Command.eval_result_state eval handle Fail _ => Toplevel.init_toplevel (); fun maybe_end_theory pos st = SOME (Toplevel.end_theory pos st) handle ERROR msg => (Output.error_message msg; NONE); val parents_reports = imports |> map_filter (fn (import, pos) => (case Thy_Info.lookup_theory import of NONE => maybe_end_theory pos (case get_result (snd (the (AList.lookup (op =) deps import))) of NONE => Toplevel.init_toplevel () | SOME (_, eval) => maybe_eval_result eval) |> Option.map (fn thy => (thy, (pos, Theory.get_markup thy))) | SOME thy => SOME (thy, (Position.none, Markup.empty)))); val parents = if null parents_reports then [Theory.get_pure ()] else map #1 parents_reports; val _ = List.app (fn (thy, r) => Context_Position.reports_global thy [r]) parents_reports; val thy = Resources.begin_theory master_dir header parents; val _ = Output.status [Markup.markup_only Markup.initialized]; in thy end; fun check_root_theory node = let val master_dir = master_directory node; val header = #header (get_header node); val header_name = #1 (#name header); val parent = if header_name = Sessions.root_name then SOME (Thy_Info.get_theory Sessions.theory_name) else if member (op =) Thy_Header.ml_roots header_name then SOME (Thy_Info.get_theory Thy_Header.ml_bootstrapN) else NONE; in parent |> Option.map (fn thy => Resources.begin_theory master_dir header [thy]) end; fun check_theory full name node = is_some (Thy_Info.lookup_theory name) orelse null (#errors (get_header node)) andalso (not full orelse is_some (get_result node)); fun last_common keywords state node_required node0 node = let fun update_flags prev (visible, initial) = let val visible' = visible andalso prev <> visible_last node; val initial' = initial andalso (case prev of NONE => true | SOME command_id => the_command_name state command_id <> Thy_Header.theoryN); in (visible', initial') end; fun get_common ((prev, command_id), opt_exec) (_, ok, flags, assign_update) = if ok then let val flags' as (visible', _) = update_flags prev flags; val ok' = (case (lookup_entry node0 command_id, opt_exec) of (SOME (eval0, _), SOME (eval, _)) => Command.eval_eq (eval0, eval) andalso (visible' orelse node_required orelse Command.eval_running eval) | _ => false); val assign_update' = assign_update |> ok' ? (case opt_exec of SOME (eval, prints) => let val command_visible = visible_command node command_id; val command_overlays = overlays node command_id; val command_name = the_command_name state command_id; in (case Command.print command_visible command_overlays keywords command_name eval prints of SOME prints' => assign_update_new (command_id, SOME (eval, prints')) | NONE => I) end | NONE => I); in SOME (prev, ok', flags', assign_update') end else NONE; val (common, ok, flags, assign_update') = iterate_entries get_common node (NONE, true, (true, true), assign_update_empty); val (common', flags') = if ok then let val last = Entries.get_after (get_entries node) common in (last, update_flags last flags) end else (common, flags); in (assign_update', common', flags') end; fun illegal_init _ = error "Illegal theory header"; fun new_exec keywords state node proper_init command_id' (assign_update, command_exec, init) = if not proper_init andalso is_none init then NONE else let val command_visible = visible_command node command_id'; val command_overlays = overlays node command_id'; val (command_name, blob_digests, blobs_index, span0) = the_command state command_id'; val blobs = map (resolve_blob state) blob_digests; val span = Lazy.force span0; val eval' = Command.eval keywords (master_directory node) (fn () => the_default illegal_init init span) (blobs, blobs_index) command_id' span (#1 (#2 command_exec)); val prints' = perhaps (Command.print command_visible command_overlays keywords command_name eval') []; val exec' = (eval', prints'); val assign_update' = assign_update_new (command_id', SOME exec') assign_update; val init' = if command_name = Thy_Header.theoryN then NONE else init; in SOME (assign_update', (command_id', exec'), init') end; fun removed_execs node0 (command_id, exec_ids) = subtract (op =) exec_ids (Command.exec_ids (lookup_entry node0 command_id)); fun print_consolidation options the_command_span node_name (assign_update, node) = (case finished_result_theory node of SOME (result_id, thy) => let val active_tasks = (node, false) |-> iterate_entries (fn (_, opt_exec) => fn active => if active then NONE else (case opt_exec of NONE => NONE | SOME (eval, _) => SOME (not (null (Execution.snapshot [Command.eval_exec_id eval]))))); in if not active_tasks then let val consolidation = if Options.bool options "editor_presentation" then let val (_, offsets, rev_segments) = iterate_entries (fn ((prev, _), opt_exec) => fn (offset, offsets, segments) => (case opt_exec of SOME (eval, _) => let val command_id = Command.eval_command_id eval; val span = the_command_span command_id; val st = (case try (#1 o the o the_entry node o the) prev of NONE => Toplevel.init_toplevel () | SOME prev_eval => Command.eval_result_state prev_eval); val exec_id = Command.eval_exec_id eval; val tr = Command.eval_result_command eval; val st' = Command.eval_result_state eval; val offset' = offset + the_default 0 (Command_Span.symbol_length span); val offsets' = offsets |> Inttab.update (command_id, offset) |> Inttab.update (exec_id, offset); val segments' = (span, (st, tr, st')) :: segments; in SOME (offset', offsets', segments') end | NONE => NONE)) node (0, Inttab.empty, []); val adjust = Inttab.lookup offsets; val segments = rev rev_segments |> map (fn (span, (st, tr, st')) => {span = Command_Span.adjust_offsets adjust span, prev_state = st, command = tr, state = st'}); val presentation_context: Thy_Info.presentation_context = {options = options, file_pos = Position.file node_name, adjust_pos = Position.adjust_offsets adjust, segments = segments}; in fn _ => let val _ = Output.status [Markup.markup_only Markup.consolidating]; val res = Exn.capture (Thy_Info.apply_presentation presentation_context) thy; val _ = commit_consolidated node; in Exn.release res end end else fn _ => commit_consolidated node; val result_entry = (case lookup_entry node result_id of NONE => err_undef "result command entry" result_id | SOME (eval, prints) => let val print = eval |> Command.print0 {pri = Task_Queue.urgent_pri + 1, print_fn = K consolidation}; in (result_id, SOME (eval, print :: prints)) end); val assign_update' = assign_update |> assign_update_change result_entry; val node' = node |> assign_entry result_entry; in (assign_update', node') end else (assign_update, node) end | NONE => (assign_update, node)); in fun update old_version_id new_version_id edits consolidate state = Runtime.exn_trace_system (fn () => let val options = Options.default (); val the_command_span = Outer_Syntax.make_span o Lazy.force o #4 o the_command state; val old_version = the_version state old_version_id; val new_version = timeit "Document.edit_nodes" (fn () => old_version |> fold edit_nodes edits |> edit_keywords edits); val consolidate = Symtab.defined (Symtab.make_set consolidate); val nodes = nodes_of new_version; val required = make_required nodes; val required0 = make_required (nodes_of old_version); - val edited = fold (fn (name, _) => Symtab.update (name, ())) edits Symtab.empty; + val edited = Symtab.build (edits |> fold (fn (name, _) => Symtab.update (name, ()))); val updated = timeit "Document.update" (fn () => nodes |> String_Graph.schedule (fn deps => fn (name, node) => (singleton o Future.forks) {name = "Document.update", group = NONE, deps = map (Future.task_of o #2) deps, pri = 1, interrupts = false} (fn () => timeit ("Document.update " ^ name) (fn () => Runtime.exn_trace_system (fn () => let val root_theory = check_root_theory node; val keywords = the_default (Session.get_keywords ()) (get_keywords node); val maybe_consolidate = consolidate name andalso could_consolidate node; val imports = map (apsnd Future.join) deps; val imports_result_changed = exists (#4 o #1 o #2) imports; val node_required = Symtab.defined required name; in if Symtab.defined edited name orelse maybe_consolidate orelse visible_node node orelse imports_result_changed orelse Symtab.defined required0 name <> node_required then let val node0 = node_of old_version name; val init = init_theory imports node; val proper_init = is_some root_theory orelse check_theory false name node andalso forall (fn (name, (_, node)) => check_theory true name node) imports; val (print_execs, common, (still_visible, initial)) = if imports_result_changed then (assign_update_empty, NONE, (true, true)) else last_common keywords state node_required node0 node; val common_command_exec = (case common of SOME id => (id, the_default Command.no_exec (the_entry node id)) | NONE => (Document_ID.none, Command.init_exec root_theory)); val (updated_execs, (command_id', exec'), _) = (print_execs, common_command_exec, if initial then SOME init else NONE) |> (still_visible orelse node_required) ? iterate_entries_after common (fn ((prev, id), _) => fn res => if not node_required andalso prev = visible_last node then NONE else new_exec keywords state node proper_init id res) node; val assign_update = (node0, updated_execs) |-> iterate_entries_after common (fn ((_, command_id0), exec0) => fn res => if is_none exec0 then NONE else if assign_update_defined updated_execs command_id0 then SOME res else SOME (assign_update_new (command_id0, NONE) res)); val last_exec = if command_id' = Document_ID.none then NONE else SOME command_id'; val result = if is_none last_exec orelse is_some (after_entry node last_exec) then NONE else SOME (command_id', #1 exec'); val result_changed = not (eq_option (Command.eval_eq o apply2 #2) (get_result node0, result)); val (assign_update', node') = node |> assign_update_apply assign_update |> set_result result |> result_changed ? reset_consolidated |> pair assign_update |> (not result_changed andalso maybe_consolidate) ? print_consolidation options the_command_span name; val assign_result = assign_update_result assign_update'; val removed = maps (removed_execs node0) assign_result; val _ = List.app Execution.cancel removed; val assigned_node = SOME (name, node'); in ((removed, assign_result, assigned_node, result_changed), node') end else (([], [], NONE, false), node) end)))) |> Future.joins |> map #1); val removed = maps #1 updated; val assign_result = maps #2 updated; val assigned_nodes = map_filter #3 updated; val state' = state |> define_version new_version_id new_version assigned_nodes; in (Symtab.keys edited, removed, assign_result, state') end); end; (** global state **) val global_state = Synchronized.var "Document.global_state" init_state; fun state () = Synchronized.value global_state; val change_state = Synchronized.change global_state; end; diff --git a/src/Pure/PIDE/resources.ML b/src/Pure/PIDE/resources.ML --- a/src/Pure/PIDE/resources.ML +++ b/src/Pure/PIDE/resources.ML @@ -1,441 +1,441 @@ (* Title: Pure/PIDE/resources.ML Author: Makarius Resources for theories and auxiliary files. *) signature RESOURCES = sig val default_qualifier: string val init_session: {session_positions: (string * Properties.T) list, session_directories: (string * string) list, session_chapters: (string * string) list, bibtex_entries: (string * string list) list, command_timings: Properties.T list, scala_functions: (string * (bool * Position.T)) list, global_theories: (string * string) list, loaded_theories: string list} -> unit val init_session_yxml: string -> unit val init_session_file: Path.T -> unit val finish_session_base: unit -> unit val global_theory: string -> string option val loaded_theory: string -> bool val check_session: Proof.context -> string * Position.T -> string val session_chapter: string -> string val last_timing: Toplevel.transition -> Time.time val scala_functions: unit -> string list val check_scala_function: Proof.context -> string * Position.T -> string * bool val master_directory: theory -> Path.T val imports_of: theory -> (string * Position.T) list val begin_theory: Path.T -> Thy_Header.header -> theory list -> theory val thy_path: Path.T -> Path.T val theory_qualifier: string -> string val theory_bibtex_entries: string -> string list val find_theory_file: string -> Path.T option val import_name: string -> Path.T -> string -> {node_name: Path.T, master_dir: Path.T, theory_name: string} val check_thy: Path.T -> string -> {master: Path.T * SHA1.digest, text: string, theory_pos: Position.T, imports: (string * Position.T) list, keywords: Thy_Header.keywords} val parse_files: (Path.T -> Path.T list) -> (theory -> Token.file list) parser val parse_file: (theory -> Token.file) parser val provide: Path.T * SHA1.digest -> theory -> theory val provide_file: Token.file -> theory -> theory val provide_parse_files: (Path.T -> Path.T list) -> (theory -> Token.file list * theory) parser val provide_parse_file: (theory -> Token.file * theory) parser val loaded_files_current: theory -> bool val check_path: Proof.context -> Path.T option -> Input.source -> Path.T val check_file: Proof.context -> Path.T option -> Input.source -> Path.T val check_dir: Proof.context -> Path.T option -> Input.source -> Path.T val check_session_dir: Proof.context -> Path.T option -> Input.source -> Path.T end; structure Resources: RESOURCES = struct (* command timings *) type timings = ((string * Time.time) Inttab.table) Symtab.table; (*file -> offset -> name, time*) val empty_timings: timings = Symtab.empty; fun update_timings props = (case Markup.parse_command_timing_properties props of SOME ({file, offset, name}, time) => Symtab.map_default (file, Inttab.empty) (Inttab.map_default (offset, (name, time)) (fn (_, t) => (name, t + time))) | NONE => I); fun make_timings command_timings = fold update_timings command_timings empty_timings; fun approximative_id name pos = (case (Position.file_of pos, Position.offset_of pos) of (SOME file, SOME offset) => if name = "" then NONE else SOME {file = file, offset = offset, name = name} | _ => NONE); fun get_timings timings tr = (case approximative_id (Toplevel.name_of tr) (Toplevel.pos_of tr) of SOME {file, offset, name} => (case Symtab.lookup timings file of SOME offsets => (case Inttab.lookup offsets offset of SOME (name', time) => if name = name' then SOME time else NONE | NONE => NONE) | NONE => NONE) | NONE => NONE) |> the_default Time.zeroTime; (* session base *) val default_qualifier = "Draft"; type entry = {pos: Position.T, serial: serial}; fun make_entry props : entry = {pos = Position.of_properties props, serial = serial ()}; val empty_session_base = ({session_positions = []: (string * entry) list, session_directories = Symtab.empty: Path.T list Symtab.table, session_chapters = Symtab.empty: string Symtab.table, bibtex_entries = Symtab.empty: string list Symtab.table, timings = empty_timings, scala_functions = Symtab.empty: (bool * Position.T) Symtab.table}, {global_theories = Symtab.empty: string Symtab.table, loaded_theories = Symtab.empty: unit Symtab.table}); val global_session_base = Synchronized.var "Sessions.base" empty_session_base; fun init_session {session_positions, session_directories, session_chapters, bibtex_entries, command_timings, scala_functions, global_theories, loaded_theories} = Synchronized.change global_session_base (fn _ => ({session_positions = sort_by #1 (map (apsnd make_entry) session_positions), session_directories = - fold_rev (fn (dir, name) => Symtab.cons_list (name, Path.explode dir)) - session_directories Symtab.empty, + Symtab.build (session_directories |> fold_rev (fn (dir, name) => + Symtab.cons_list (name, Path.explode dir))), session_chapters = Symtab.make session_chapters, bibtex_entries = Symtab.make bibtex_entries, timings = make_timings command_timings, scala_functions = Symtab.make scala_functions}, {global_theories = Symtab.make global_theories, loaded_theories = Symtab.make_set loaded_theories})); fun init_session_yxml yxml = let val (session_positions, (session_directories, (session_chapters, (bibtex_entries, (command_timings, (scala_functions, (global_theories, loaded_theories))))))) = YXML.parse_body yxml |> let open XML.Decode in (pair (list (pair string properties)) (pair (list (pair string string)) (pair (list (pair string string)) (pair (list (pair string (list string))) (pair (list properties) (pair (list (pair string (pair bool properties))) (pair (list (pair string string)) (list string)))))))) end; in init_session {session_positions = session_positions, session_directories = session_directories, session_chapters = session_chapters, bibtex_entries = bibtex_entries, command_timings = command_timings, scala_functions = (map o apsnd o apsnd) Position.of_properties scala_functions, global_theories = global_theories, loaded_theories = loaded_theories} end; fun init_session_file path = init_session_yxml (File.read path) before File.rm path; fun finish_session_base () = Synchronized.change global_session_base (apfst (K (#1 empty_session_base))); fun get_session_base f = f (Synchronized.value global_session_base); fun get_session_base1 f = get_session_base (f o #1); fun get_session_base2 f = get_session_base (f o #2); fun global_theory a = Symtab.lookup (get_session_base2 #global_theories) a; fun loaded_theory a = Symtab.defined (get_session_base2 #loaded_theories) a; fun check_session ctxt arg = Completion.check_item "session" (fn (name, {pos, serial}) => Position.make_entity_markup false serial Markup.sessionN (name, pos)) (get_session_base1 #session_positions) ctxt arg; fun session_chapter name = the_default "Unsorted" (Symtab.lookup (get_session_base1 #session_chapters) name); fun last_timing tr = get_timings (get_session_base1 #timings) tr; (* Scala functions *) (*raw bootstrap environment*) fun scala_functions () = space_explode "," (getenv "ISABELLE_SCALA_FUNCTIONS"); (*regular resources*) fun scala_function a = (a, the_default (false, Position.none) (Symtab.lookup (get_session_base1 #scala_functions) a)); fun check_scala_function ctxt arg = let val funs = scala_functions () |> sort_strings |> map scala_function; val name = Completion.check_entity Markup.scala_functionN (map (apsnd #2) funs) ctxt arg; val multi = (case AList.lookup (op =) funs name of SOME (multi, _) => multi | NONE => false); in (name, multi) end; val _ = Theory.setup (Document_Output.antiquotation_verbatim_embedded \<^binding>\scala_function\ (Scan.lift Parse.embedded_position) (#1 oo check_scala_function) #> ML_Antiquotation.inline_embedded \<^binding>\scala_function\ (Args.context -- Scan.lift Parse.embedded_position >> (uncurry check_scala_function #> #1 #> ML_Syntax.print_string)) #> ML_Antiquotation.value_embedded \<^binding>\scala\ (Args.context -- Scan.lift Args.embedded_position >> (fn (ctxt, arg) => let val (name, multi) = check_scala_function ctxt arg; val func = if multi then "Scala.function" else "Scala.function1"; in ML_Syntax.atomic (func ^ " " ^ ML_Syntax.print_string name) end))); (* manage source files *) type files = {master_dir: Path.T, (*master directory of theory source*) imports: (string * Position.T) list, (*source specification of imports*) provided: (Path.T * SHA1.digest) list}; (*source path, digest*) fun make_files (master_dir, imports, provided): files = {master_dir = master_dir, imports = imports, provided = provided}; structure Files = Theory_Data ( type T = files; val empty = make_files (Path.current, [], []); val extend = I; fun merge ({master_dir, imports, provided = provided1}, {provided = provided2, ...}) = let val provided' = Library.merge (op =) (provided1, provided2) in make_files (master_dir, imports, provided') end ); fun map_files f = Files.map (fn {master_dir, imports, provided} => make_files (f (master_dir, imports, provided))); val master_directory = #master_dir o Files.get; val imports_of = #imports o Files.get; fun begin_theory master_dir {name, imports, keywords} parents = Theory.begin_theory name parents |> map_files (fn _ => (Path.explode (Path.implode_symbolic master_dir), imports, [])) |> Thy_Header.add_keywords keywords; (* theory files *) val thy_path = Path.ext "thy"; fun theory_qualifier theory = (case global_theory theory of SOME qualifier => qualifier | NONE => Long_Name.qualifier theory); fun theory_name qualifier theory = if Long_Name.is_qualified theory orelse is_some (global_theory theory) then theory else Long_Name.qualify qualifier theory; fun theory_bibtex_entries theory = Symtab.lookup_list (get_session_base1 #bibtex_entries) (theory_qualifier theory); fun find_theory_file thy_name = let val thy_file = thy_path (Path.basic (Long_Name.base_name thy_name)); val session = theory_qualifier thy_name; val dirs = Symtab.lookup_list (get_session_base1 #session_directories) session; in dirs |> get_first (fn dir => let val path = dir + thy_file in if File.is_file path then SOME path else NONE end) end; fun make_theory_node node_name theory = {node_name = node_name, master_dir = Path.dir node_name, theory_name = theory}; fun loaded_theory_node theory = {node_name = Path.basic theory, master_dir = Path.current, theory_name = theory}; fun import_name qualifier dir s = let val theory = theory_name qualifier (Thy_Header.import_name s); fun theory_node () = make_theory_node (File.full_path dir (thy_path (Path.expand (Path.explode s)))) theory; in if not (Thy_Header.is_base_name s) then theory_node () else if loaded_theory theory then loaded_theory_node theory else (case find_theory_file theory of SOME node_name => make_theory_node node_name theory | NONE => if Long_Name.is_qualified s then loaded_theory_node theory else theory_node ()) end; fun check_file dir file = File.check_file (File.full_path dir file); fun check_thy dir thy_name = let val thy_base_name = Long_Name.base_name thy_name; val master_file = (case find_theory_file thy_name of SOME path => check_file Path.current path | NONE => check_file dir (thy_path (Path.basic thy_base_name))); val text = File.read master_file; val {name = (name, pos), imports, keywords} = Thy_Header.read (Path.position master_file) text; val _ = thy_base_name <> name andalso error ("Bad theory name " ^ quote name ^ " for file " ^ Path.print (Path.base master_file) ^ Position.here pos); in {master = (master_file, SHA1.digest text), text = text, theory_pos = pos, imports = imports, keywords = keywords} end; (* load files *) fun parse_files make_paths = Scan.ahead Parse.not_eof -- Parse.path_input >> (fn (tok, source) => fn thy => (case Token.get_files tok of [] => let val master_dir = master_directory thy; val name = Input.string_of source; val pos = Input.pos_of source; val delimited = Input.is_delimited source; val src_paths = make_paths (Path.explode name); in map (Command.read_file master_dir pos delimited) src_paths end | files => map Exn.release files)); val parse_file = parse_files single >> (fn f => f #> the_single); fun provide (src_path, id) = map_files (fn (master_dir, imports, provided) => if AList.defined (op =) provided src_path then error ("Duplicate use of source file: " ^ Path.print src_path) else (master_dir, imports, (src_path, id) :: provided)); fun provide_file (file: Token.file) = provide (#src_path file, #digest file); fun provide_parse_files make_paths = parse_files make_paths >> (fn files => fn thy => let val fs = files thy; val thy' = fold (fn {src_path, digest, ...} => provide (src_path, digest)) fs thy; in (fs, thy') end); val provide_parse_file = provide_parse_files single >> (fn f => f #>> the_single); fun load_file thy src_path = let val full_path = check_file (master_directory thy) src_path; val text = File.read full_path; val id = SHA1.digest text; in ((full_path, id), text) end; fun loaded_files_current thy = #provided (Files.get thy) |> forall (fn (src_path, id) => (case try (load_file thy) src_path of NONE => false | SOME ((_, id'), _) => id = id')); (* formal check *) fun formal_check check_file ctxt opt_dir source = let val name = Input.string_of source; val pos = Input.pos_of source; val delimited = Input.is_delimited source; val _ = Context_Position.report ctxt pos (Markup.language_path delimited); fun err msg = error (msg ^ Position.here pos); val dir = (case opt_dir of SOME dir => dir | NONE => master_directory (Proof_Context.theory_of ctxt)); val path = dir + Path.explode name handle ERROR msg => err msg; val _ = Path.expand path handle ERROR msg => err msg; val _ = Context_Position.report ctxt pos (Markup.path (Path.implode_symbolic path)); val _ : Path.T = check_file path handle ERROR msg => err msg; in path end; val check_path = formal_check I; val check_file = formal_check File.check_file; val check_dir = formal_check File.check_dir; fun check_session_dir ctxt opt_dir s = let val dir = Path.expand (check_dir ctxt opt_dir s); val ok = File.is_file (dir + Path.explode("ROOT")) orelse File.is_file (dir + Path.explode("ROOTS")); in if ok then dir else error ("Bad session root directory (missing ROOT or ROOTS): " ^ Path.print dir ^ Position.here (Input.pos_of s)) end; (* antiquotations *) local fun document_antiq (check: Proof.context -> Path.T option -> Input.source -> Path.T) = Args.context -- Scan.lift Parse.path_input >> (fn (ctxt, source) => (check ctxt NONE source; Latex.string (Latex.output_ascii_breakable "/" (Input.string_of source)) |> Latex.enclose_text "\\isatt{" "}")); fun ML_antiq check = Args.context -- Scan.lift Parse.path_input >> (fn (ctxt, source) => check ctxt (SOME Path.current) source |> ML_Syntax.print_path); in val _ = Theory.setup (Document_Output.antiquotation_verbatim_embedded \<^binding>\session\ (Scan.lift Parse.embedded_position) check_session #> Document_Output.antiquotation_raw_embedded \<^binding>\path\ (document_antiq check_path) (K I) #> Document_Output.antiquotation_raw_embedded \<^binding>\file\ (document_antiq check_file) (K I) #> Document_Output.antiquotation_raw_embedded \<^binding>\dir\ (document_antiq check_dir) (K I) #> ML_Antiquotation.value_embedded \<^binding>\path\ (ML_antiq check_path) #> ML_Antiquotation.value_embedded \<^binding>\file\ (ML_antiq check_file) #> ML_Antiquotation.value_embedded \<^binding>\dir\ (ML_antiq check_dir) #> ML_Antiquotation.value_embedded \<^binding>\path_binding\ (Scan.lift (Parse.position Parse.path) >> (ML_Syntax.print_path_binding o Path.explode_binding)) #> ML_Antiquotation.value \<^binding>\master_dir\ (Args.theory >> (ML_Syntax.print_path o master_directory))); end; end; diff --git a/src/Pure/PIDE/xml.ML b/src/Pure/PIDE/xml.ML --- a/src/Pure/PIDE/xml.ML +++ b/src/Pure/PIDE/xml.ML @@ -1,422 +1,421 @@ (* Title: Pure/PIDE/xml.ML Author: David Aspinall Author: Stefan Berghofer Author: Makarius Untyped XML trees and representation of ML values. *) signature XML_DATA_OPS = sig type 'a A type 'a T type 'a V type 'a P val int_atom: int A val bool_atom: bool A val unit_atom: unit A val properties: Properties.T T val string: string T val int: int T val bool: bool T val unit: unit T val pair: 'a T -> 'b T -> ('a * 'b) T val triple: 'a T -> 'b T -> 'c T -> ('a * 'b * 'c) T val list: 'a T -> 'a list T val option: 'a T -> 'a option T val variant: 'a V list -> 'a T end; signature XML = sig type attributes = (string * string) list datatype tree = Elem of (string * attributes) * tree list | Text of string type body = tree list val blob: string list -> body val is_empty: tree -> bool val is_empty_body: body -> bool val xml_elemN: string val xml_nameN: string val xml_bodyN: string val wrap_elem: ((string * attributes) * tree list) * tree list -> tree val unwrap_elem: tree -> (((string * attributes) * tree list) * tree list) option val add_content: tree -> Buffer.T -> Buffer.T val content_of: body -> string val trim_blanks: body -> body val header: string val text: string -> string val element: string -> attributes -> string list -> string val output_markup: Markup.T -> Markup.output val string_of: tree -> string val pretty: int -> tree -> Pretty.T val parse_comments: string list -> unit * string list val parse_string : string -> string option val parse_element: string list -> tree * string list val parse_document: string list -> tree * string list val parse: string -> tree exception XML_ATOM of string exception XML_BODY of body structure Encode: sig include XML_DATA_OPS val tree: tree T end structure Decode: sig include XML_DATA_OPS val tree: tree T end end; structure XML: XML = struct (** XML trees **) open Output_Primitives.XML; val blob = map Text; fun is_empty (Text "") = true | is_empty _ = false; val is_empty_body = forall is_empty; (* wrapped elements *) val xml_elemN = "xml_elem"; val xml_nameN = "xml_name"; val xml_bodyN = "xml_body"; fun wrap_elem (((a, atts), body1), body2) = Elem ((xml_elemN, (xml_nameN, a) :: atts), Elem ((xml_bodyN, []), body1) :: body2); fun unwrap_elem (Elem ((name, (n, a) :: atts), Elem ((name', atts'), body1) :: body2)) = if name = xml_elemN andalso n = xml_nameN andalso name' = xml_bodyN andalso null atts' then SOME (((a, atts), body1), body2) else NONE | unwrap_elem _ = NONE; (* text content *) fun add_content tree = (case unwrap_elem tree of SOME (_, ts) => fold add_content ts | NONE => (case tree of Elem (_, ts) => fold add_content ts | Text s => Buffer.add s)); -fun content_of body = Buffer.empty |> fold add_content body |> Buffer.content; +val content_of = Buffer.build_content o fold add_content; (* trim blanks *) fun trim_blanks trees = trees |> maps (fn Elem (markup, body) => [Elem (markup, trim_blanks body)] | Text s => let val s' = s |> raw_explode |> trim Symbol.is_blank |> implode; in if s' = "" then [] else [Text s'] end); (** string representation **) val header = "\n"; (* escaped text *) fun decode "<" = "<" | decode ">" = ">" | decode "&" = "&" | decode "'" = "'" | decode """ = "\"" | decode c = c; fun encode "<" = "<" | encode ">" = ">" | encode "&" = "&" | encode "'" = "'" | encode "\"" = """ | encode c = c; val text = translate_string encode; (* elements *) fun elem name atts = space_implode " " (name :: map (fn (a, x) => a ^ "=\"" ^ text x ^ "\"") atts); fun element name atts body = let val b = implode body in if b = "" then enclose "<" "/>" (elem name atts) else enclose "<" ">" (elem name atts) ^ b ^ enclose "" name end; fun output_markup (markup as (name, atts)) = if Markup.is_empty markup then Markup.no_output else (enclose "<" ">" (elem name atts), enclose "" name); -(* output *) +(* output content *) -fun buffer_of depth tree = +fun content_depth depth = let fun traverse _ (Elem ((name, atts), [])) = Buffer.add "<" #> Buffer.add (elem name atts) #> Buffer.add "/>" | traverse d (Elem ((name, atts), ts)) = Buffer.add "<" #> Buffer.add (elem name atts) #> Buffer.add ">" #> traverse_body d ts #> Buffer.add " Buffer.add name #> Buffer.add ">" | traverse _ (Text s) = Buffer.add (text s) and traverse_body 0 _ = Buffer.add "..." | traverse_body d ts = fold (traverse (d - 1)) ts; - in Buffer.empty |> traverse depth tree end; + in Buffer.build_content o traverse depth end; -val string_of = Buffer.content o buffer_of ~1; +val string_of = content_depth ~1; -fun pretty depth tree = - Pretty.str (Buffer.content (buffer_of (Int.max (0, depth)) tree)); +fun pretty depth tree = Pretty.str (content_depth (Int.max (0, depth)) tree); val _ = ML_system_pp (fn depth => fn _ => Pretty.to_polyml o pretty (FixedInt.toInt depth)); (** XML parsing **) local fun err msg (xs, _) = fn () => "XML parsing error: " ^ msg () ^ "\nfound: " ^ quote (Symbol.beginning 100 xs); fun ignored _ = []; fun name_start_char c = Symbol.is_ascii_letter c orelse c = ":" orelse c = "_"; fun name_char c = name_start_char c orelse Symbol.is_ascii_digit c orelse c = "-" orelse c = "."; val parse_name = Scan.one name_start_char ::: Scan.many name_char; val blanks = Scan.many Symbol.is_blank; val special = $$ "&" ^^ (parse_name >> implode) ^^ $$ ";" >> decode; val regular = Scan.one Symbol.not_eof; fun regular_except x = Scan.one (fn c => Symbol.not_eof c andalso c <> x); val parse_chars = Scan.repeat1 (special || regular_except "<") >> implode; val parse_cdata = Scan.this_string "") regular) >> implode) --| Scan.this_string "]]>"; val parse_att = ((parse_name >> implode) --| (blanks -- $$ "=" -- blanks)) -- (($$ "\"" || $$ "'") :|-- (fn s => (Scan.repeat (special || regular_except s) >> implode) --| $$ s)); val parse_comment = Scan.this_string "") regular) -- Scan.this_string "-->" >> ignored; val parse_processing_instruction = Scan.this_string "") regular) -- Scan.this_string "?>" >> ignored; val parse_doctype = Scan.this_string "") regular) -- $$ ">" >> ignored; val parse_misc = Scan.one Symbol.is_blank >> ignored || parse_processing_instruction || parse_comment; val parse_optional_text = Scan.optional (parse_chars >> (single o Text)) []; in val parse_comments = blanks -- Scan.repeat (parse_comment -- blanks >> K ()) >> K (); val parse_string = Scan.read Symbol.stopper parse_chars o raw_explode; fun parse_content xs = (parse_optional_text @@@ Scan.repeats ((parse_element >> single || parse_cdata >> (single o Text) || parse_processing_instruction || parse_comment) @@@ parse_optional_text)) xs and parse_element xs = ($$ "<" |-- parse_name -- Scan.repeat (blanks |-- parse_att) --| blanks :-- (fn (name, _) => !! (err (fn () => "Expected > or />")) ($$ "/" -- $$ ">" >> ignored || $$ ">" |-- parse_content --| !! (err (fn () => "Expected ")) ($$ "<" -- $$ "/" -- Scan.this name -- blanks -- $$ ">"))) >> (fn ((name, atts), body) => Elem ((implode name, atts), body))) xs; val parse_document = (Scan.repeat parse_misc -- Scan.option parse_doctype -- Scan.repeat parse_misc) |-- parse_element; fun parse s = (case Scan.finite Symbol.stopper (Scan.error (!! (err (fn () => "Malformed element")) (blanks |-- parse_document --| blanks))) (raw_explode s) of (x, []) => x | (_, ys) => error ("XML parsing error: unprocessed input\n" ^ Symbol.beginning 100 ys)); end; (** XML as data representation language **) exception XML_ATOM of string; exception XML_BODY of tree list; structure Encode = struct type 'a A = 'a -> string; type 'a T = 'a -> body; type 'a V = 'a -> string list * body; type 'a P = 'a -> string list; (* atomic values *) fun int_atom i = Value.print_int i; fun bool_atom false = "0" | bool_atom true = "1"; fun unit_atom () = ""; (* structural nodes *) fun node ts = Elem ((":", []), ts); fun vector xs = map_index (fn (i, x) => (int_atom i, x)) xs; fun tagged (tag, (xs, ts)) = Elem ((int_atom tag, vector xs), ts); (* representation of standard types *) fun tree (t: tree) = [t]; fun properties props = [Elem ((":", props), [])]; fun string "" = [] | string s = [Text s]; val int = string o int_atom; val bool = string o bool_atom; val unit = string o unit_atom; fun pair f g (x, y) = [node (f x), node (g y)]; fun triple f g h (x, y, z) = [node (f x), node (g y), node (h z)]; fun list f xs = map (node o f) xs; fun option _ NONE = [] | option f (SOME x) = [node (f x)]; fun variant fs x = [tagged (the (get_index (fn f => SOME (f x) handle General.Match => NONE) fs))]; end; structure Decode = struct type 'a A = string -> 'a; type 'a T = body -> 'a; type 'a V = string list * body -> 'a; type 'a P = string list -> 'a; (* atomic values *) fun int_atom s = Value.parse_int s handle Fail _ => raise XML_ATOM s; fun bool_atom "0" = false | bool_atom "1" = true | bool_atom s = raise XML_ATOM s; fun unit_atom "" = () | unit_atom s = raise XML_ATOM s; (* structural nodes *) fun node (Elem ((":", []), ts)) = ts | node t = raise XML_BODY [t]; fun vector atts = map_index (fn (i, (a, x)) => if int_atom a = i then x else raise XML_ATOM a) atts; fun tagged (Elem ((name, atts), ts)) = (int_atom name, (vector atts, ts)) | tagged t = raise XML_BODY [t]; (* representation of standard types *) fun tree [t] = t | tree ts = raise XML_BODY ts; fun properties [Elem ((":", props), [])] = props | properties ts = raise XML_BODY ts; fun string [] = "" | string [Text s] = s | string ts = raise XML_BODY ts; val int = int_atom o string; val bool = bool_atom o string; val unit = unit_atom o string; fun pair f g [t1, t2] = (f (node t1), g (node t2)) | pair _ _ ts = raise XML_BODY ts; fun triple f g h [t1, t2, t3] = (f (node t1), g (node t2), h (node t3)) | triple _ _ _ ts = raise XML_BODY ts; fun list f ts = map (f o node) ts; fun option _ [] = NONE | option f [t] = SOME (f (node t)) | option _ ts = raise XML_BODY ts; fun variant fs [t] = let val (tag, (xs, ts)) = tagged t; val f = nth fs tag handle General.Subscript => raise XML_BODY [t]; in f (xs, ts) end | variant _ ts = raise XML_BODY ts; end; end; diff --git a/src/Pure/PIDE/yxml.ML b/src/Pure/PIDE/yxml.ML --- a/src/Pure/PIDE/yxml.ML +++ b/src/Pure/PIDE/yxml.ML @@ -1,175 +1,175 @@ (* Title: Pure/PIDE/yxml.ML Author: Makarius Efficient text representation of XML trees using extra characters X and Y -- no escaping, may nest marked text verbatim. Suitable for direct inlining into plain text. Markup ...body... is encoded as: X Y name Y att=val ... X ... body ... X Y X *) signature YXML = sig val X: Symbol.symbol val Y: Symbol.symbol val embed_controls: string -> string val detect: string -> bool val traverse: (string -> 'a -> 'a) -> XML.tree -> 'a -> 'a val tree_size: XML.tree -> int val body_size: XML.body -> int val string_of_body: XML.body -> string val string_of: XML.tree -> string val write_body: Path.T -> XML.body -> unit val output_markup: Markup.T -> string * string val output_markup_elem: Markup.T -> (string * string) * string val parse_body: string -> XML.body val parse: string -> XML.tree val content_of: string -> string val is_wellformed: string -> bool end; structure YXML: YXML = struct (** string representation **) (* idempotent recoding of certain low ASCII control characters *) fun pseudo_utf8 c = if Symbol.is_ascii_control c then chr 192 ^ chr (128 + ord c) else c; fun embed_controls str = if exists_string Symbol.is_ascii_control str then translate_string pseudo_utf8 str else str; (* markers *) val X_char = Char.chr 5; val Y_char = Char.chr 6; val X = String.str X_char; val Y = String.str Y_char; val XY = X ^ Y; val XYX = XY ^ X; fun detect s = Char.contains s X_char orelse Char.contains s Y_char; (* traverse *) fun traverse string = let fun attrib (a, x) = string Y #> string a #> string "=" #> string x; fun tree (XML.Elem (markup as (name, atts), ts)) = if Markup.is_empty markup then fold tree ts else string XY #> string name #> fold attrib atts #> string X #> fold tree ts #> string XYX | tree (XML.Text s) = string s; in tree end; fun tree_size tree = traverse (Integer.add o size) tree 0; fun body_size body = fold (Integer.add o tree_size) body 0; (* output *) -fun string_of_body body = Buffer.empty |> fold (traverse Buffer.add) body |> Buffer.content; +val string_of_body = Buffer.build_content o fold (traverse Buffer.add); val string_of = string_of_body o single; fun write_body path body = path |> File.open_output (fn file => fold (traverse (fn s => fn () => File.output file s)) body ()); (* markup elements *) val Z = chr 0; val Z_text = [XML.Text Z]; fun output_markup (markup as (name, atts)) = if Markup.is_empty markup then Markup.no_output else (XY ^ name ^ implode (map (fn (a, x) => Y ^ a ^ "=" ^ x) atts) ^ X, XYX); fun output_markup_elem markup = let val [bg1, bg2, en] = space_explode Z (string_of (XML.wrap_elem ((markup, Z_text), Z_text))) in ((bg1, bg2), en) end; (** efficient YXML parsing **) local (* splitting *) val split_string = Substring.full #> Substring.tokens (fn c => c = X_char) #> map (Substring.fields (fn c => c = Y_char) #> map Substring.string); (* structural errors *) fun err msg = raise Fail ("Malformed YXML: " ^ msg); fun err_attribute () = err "bad attribute"; fun err_element () = err "bad element"; fun err_unbalanced "" = err "unbalanced element" | err_unbalanced name = err ("unbalanced element " ^ quote name); (* stack operations *) fun add x ((elem, body) :: pending) = (elem, x :: body) :: pending; fun push "" _ _ = err_element () | push name atts pending = ((name, atts), []) :: pending; fun pop ((("", _), _) :: _) = err_unbalanced "" | pop ((markup, body) :: pending) = add (XML.Elem (markup, rev body)) pending; (* parsing *) fun parse_attrib s = (case first_field "=" s of NONE => err_attribute () | SOME ("", _) => err_attribute () | SOME att => att); fun parse_chunk ["", ""] = pop | parse_chunk ("" :: name :: atts) = push name (map parse_attrib atts) | parse_chunk txts = fold (add o XML.Text) txts; in fun parse_body source = (case fold parse_chunk (split_string source) [(("", []), [])] of [(("", _), result)] => rev result | ((name, _), _) :: _ => err_unbalanced name); fun parse source = (case parse_body source of [result] => result | [] => XML.Text "" | _ => err "multiple results"); end; val content_of = parse_body #> XML.content_of; fun is_wellformed s = string_of_body (parse_body s) = s handle Fail _ => false; end; diff --git a/src/Pure/Proof/extraction.ML b/src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML +++ b/src/Pure/Proof/extraction.ML @@ -1,862 +1,862 @@ (* Title: Pure/Proof/extraction.ML Author: Stefan Berghofer, TU Muenchen Extraction of programs from proofs. *) signature EXTRACTION = sig val set_preprocessor : (theory -> Proofterm.proof -> Proofterm.proof) -> theory -> theory val add_realizes_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory val add_realizes_eqns : string list -> theory -> theory val add_typeof_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory val add_typeof_eqns : string list -> theory -> theory val add_realizers_i : (string * (string list * term * Proofterm.proof)) list -> theory -> theory val add_realizers : (thm * (string list * string * string)) list -> theory -> theory val add_expand_thm : bool -> thm -> theory -> theory val add_types : (xstring * ((term -> term option) list * (term -> typ -> term -> typ -> term) option)) list -> theory -> theory val extract : (thm * string list) list -> theory -> theory val nullT : typ val nullt : term val mk_typ : typ -> term val etype_of : theory -> string list -> typ list -> term -> typ val realizes_of: theory -> string list -> term -> term -> term val abs_corr_shyps: theory -> thm -> string list -> term list -> Proofterm.proof -> Proofterm.proof end; structure Extraction : EXTRACTION = struct (**** tools ****) val typ = Simple_Syntax.read_typ; val add_syntax = Sign.root_path #> Sign.add_types_global [(Binding.make ("Type", \<^here>), 0, NoSyn), (Binding.make ("Null", \<^here>), 0, NoSyn)] #> Sign.add_consts [(Binding.make ("typeof", \<^here>), typ "'b \ Type", NoSyn), (Binding.make ("Type", \<^here>), typ "'a itself \ Type", NoSyn), (Binding.make ("Null", \<^here>), typ "Null", NoSyn), (Binding.make ("realizes", \<^here>), typ "'a \ 'b \ 'b", NoSyn)]; val nullT = Type ("Null", []); val nullt = Const ("Null", nullT); fun mk_typ T = Const ("Type", Term.itselfT T --> Type ("Type", [])) $ Logic.mk_type T; fun typeof_proc defaultS vs (Const ("typeof", _) $ u) = SOME (mk_typ (case strip_comb u of (Var ((a, i), _), _) => if member (op =) vs a then TFree ("'" ^ a ^ ":" ^ string_of_int i, defaultS) else nullT | (Free (a, _), _) => if member (op =) vs a then TFree ("'" ^ a, defaultS) else nullT | _ => nullT)) | typeof_proc _ _ _ = NONE; fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ _ $ t) = SOME t | rlz_proc (Const ("realizes", Type (_, [T, _])) $ r $ t) = (case strip_comb t of (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, T --> U), r :: ts)) | (Free (s, U), ts) => SOME (list_comb (Free (s, T --> U), r :: ts)) | _ => NONE) | rlz_proc _ = NONE; val unpack_ixn = apfst implode o apsnd (fst o read_int o tl) o chop_prefix (fn s => s <> ":") o raw_explode; type rules = {next: int, rs: ((term * term) list * (term * term)) list, net: (int * ((term * term) list * (term * term))) Net.net}; val empty_rules : rules = {next = 0, rs = [], net = Net.empty}; fun add_rule (r as (_, (lhs, _))) ({next, rs, net} : rules) = {next = next - 1, rs = r :: rs, net = Net.insert_term (K false) (Envir.eta_contract lhs, (next, r)) net}; fun merge_rules ({next, rs = rs1, net} : rules) ({rs = rs2, ...} : rules) = fold_rev add_rule (subtract (op =) rs1 rs2) {next = next, rs = rs1, net = net}; fun condrew thy rules procs = let fun rew tm = Pattern.rewrite_term thy [] (condrew' :: procs) tm and condrew' tm = let val cache = Unsynchronized.ref ([] : (term * term) list); fun lookup f x = (case AList.lookup (op =) (!cache) x of NONE => let val y = f x in (cache := (x, y) :: !cache; y) end | SOME y => y); in get_first (fn (_, (prems, (tm1, tm2))) => let fun ren t = the_default t (Term.rename_abs tm1 tm t); val inc = Logic.incr_indexes ([], [], maxidx_of_term tm + 1); val env as (Tenv, tenv) = Pattern.match thy (inc tm1, tm) (Vartab.empty, Vartab.empty); val prems' = map (apply2 (Envir.subst_term env o inc o ren)) prems; val env' = Envir.Envir {maxidx = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u) prems' ~1, tenv = tenv, tyenv = Tenv}; val env'' = fold (Pattern.unify (Context.Theory thy) o apply2 (lookup rew)) prems' env'; in SOME (Envir.norm_term env'' (inc (ren tm2))) end handle Pattern.MATCH => NONE | Pattern.Unif => NONE) (sort (int_ord o apply2 fst) (Net.match_term rules (Envir.eta_contract tm))) end; in rew end; val change_types = Proofterm.change_types o SOME; fun extr_name s vs = Long_Name.append "extr" (space_implode "_" (s :: vs)); fun corr_name s vs = extr_name s vs ^ "_correctness"; fun msg d s = writeln (Symbol.spaces d ^ s); -fun vars_of t = map Var (rev (Term.add_vars t [])); -fun frees_of t = map Free (rev (Term.add_frees t [])); +fun vars_of t = map Var (build_rev (Term.add_vars t)); +fun frees_of t = map Free (build_rev (Term.add_frees t)); fun vfs_of t = vars_of t @ frees_of t; val mkabs = fold_rev (fn v => fn t => Abs ("x", fastype_of v, abstract_over (v, t))); val mkabsp = fold_rev (fn t => fn prf => AbsP ("H", SOME t, prf)); fun strip_abs 0 t = t | strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t | strip_abs _ _ = error "strip_abs: not an abstraction"; val prf_subst_TVars = Proofterm.map_proof_types o typ_subst_TVars; fun relevant_vars types prop = List.foldr (fn (Var ((a, _), T), vs) => (case body_type T of Type (s, _) => if member (op =) types s then a :: vs else vs | _ => vs) | (_, vs) => vs) [] (vars_of prop); fun tname_of (Type (s, _)) = s | tname_of _ = ""; fun get_var_type t = let val vs = Term.add_vars t []; val fs = Term.add_frees t []; in fn Var (ixn, _) => (case AList.lookup (op =) vs ixn of NONE => error "get_var_type: no such variable in term" | SOME T => Var (ixn, T)) | Free (s, _) => (case AList.lookup (op =) fs s of NONE => error "get_var_type: no such variable in term" | SOME T => Free (s, T)) | _ => error "get_var_type: not a variable" end; fun read_term ctxt T s = let val ctxt' = ctxt |> Proof_Context.set_defsort [] |> Config.put Type_Infer.object_logic false |> Config.put Type_Infer_Context.const_sorts false; val parse = if T = propT then Syntax.parse_prop else Syntax.parse_term; in parse ctxt' s |> Type.constraint T |> Syntax.check_term ctxt' end; fun make_proof_body prf = let val (oracles, thms) = ([prf], ([], [])) |-> Proofterm.fold_proof_atoms false (fn Oracle (name, prop, _) => apfst (cons ((name, Position.none), SOME prop)) | PThm (header, thm_body) => apsnd (cons (Proofterm.make_thm header thm_body)) | _ => I); val body = PBody {oracles = Ord_List.make Proofterm.oracle_ord oracles, thms = Ord_List.make Proofterm.thm_ord thms, proof = prf}; in Proofterm.thm_body body end; (**** theory data ****) (* theory data *) structure ExtractionData = Theory_Data ( type T = {realizes_eqns : rules, typeof_eqns : rules, types : (string * ((term -> term option) list * (term -> typ -> term -> typ -> term) option)) list, realizers : (string list * (term * proof)) list Symtab.table, defs : thm list, expand : string list, prep : (theory -> proof -> proof) option} val empty = {realizes_eqns = empty_rules, typeof_eqns = empty_rules, types = [], realizers = Symtab.empty, defs = [], expand = [], prep = NONE}; val extend = I; fun merge ({realizes_eqns = realizes_eqns1, typeof_eqns = typeof_eqns1, types = types1, realizers = realizers1, defs = defs1, expand = expand1, prep = prep1}, {realizes_eqns = realizes_eqns2, typeof_eqns = typeof_eqns2, types = types2, realizers = realizers2, defs = defs2, expand = expand2, prep = prep2}) : T = {realizes_eqns = merge_rules realizes_eqns1 realizes_eqns2, typeof_eqns = merge_rules typeof_eqns1 typeof_eqns2, types = AList.merge (op =) (K true) (types1, types2), realizers = Symtab.merge_list (eq_set (op =) o apply2 #1) (realizers1, realizers2), defs = Library.merge Thm.eq_thm (defs1, defs2), expand = Library.merge (op =) (expand1, expand2), prep = if is_some prep1 then prep1 else prep2}; ); fun read_condeq thy = let val ctxt' = Proof_Context.init_global (add_syntax thy) in fn s => let val t = Logic.varify_global (read_term ctxt' propT s) in (map Logic.dest_equals (Logic.strip_imp_prems t), Logic.dest_equals (Logic.strip_imp_concl t)) handle TERM _ => error ("Not a (conditional) meta equality:\n" ^ s) end end; (** preprocessor **) fun set_preprocessor prep thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, ...} = ExtractionData.get thy in ExtractionData.put {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy end; (** equations characterizing realizability **) fun gen_add_realizes_eqns prep_eq eqns thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; in ExtractionData.put {realizes_eqns = fold_rev add_rule (map (prep_eq thy) eqns) realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = realizers, defs = defs, expand = expand, prep = prep} thy end val add_realizes_eqns_i = gen_add_realizes_eqns (K I); val add_realizes_eqns = gen_add_realizes_eqns read_condeq; (** equations characterizing type of extracted program **) fun gen_add_typeof_eqns prep_eq eqns thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; val eqns' = map (prep_eq thy) eqns in ExtractionData.put {realizes_eqns = realizes_eqns, realizers = realizers, typeof_eqns = fold_rev add_rule eqns' typeof_eqns, types = types, defs = defs, expand = expand, prep = prep} thy end val add_typeof_eqns_i = gen_add_typeof_eqns (K I); val add_typeof_eqns = gen_add_typeof_eqns read_condeq; fun thaw (T as TFree (a, S)) = if exists_string (fn s => s = ":") a then TVar (unpack_ixn a, S) else T | thaw (Type (a, Ts)) = Type (a, map thaw Ts) | thaw T = T; fun freeze (TVar ((a, i), S)) = TFree (a ^ ":" ^ string_of_int i, S) | freeze (Type (a, Ts)) = Type (a, map freeze Ts) | freeze T = T; fun freeze_thaw f x = map_types thaw (f (map_types freeze x)); fun etype_of thy vs Ts t = let val {typeof_eqns, ...} = ExtractionData.get thy; fun err () = error ("Unable to determine type of extracted program for\n" ^ Syntax.string_of_term_global thy t) in (case strip_abs_body (freeze_thaw (condrew thy (#net typeof_eqns) [typeof_proc [] vs]) (fold (Term.abs o pair "x") Ts (Const ("typeof", fastype_of1 (Ts, t) --> Type ("Type", [])) $ t))) of Const ("Type", _) $ u => (Logic.dest_type u handle TERM _ => err ()) | _ => err ()) end; (** realizers for axioms / theorems, together with correctness proofs **) fun gen_add_realizers prep_rlz rs thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy in ExtractionData.put {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = fold (Symtab.cons_list o prep_rlz thy) rs realizers, defs = defs, expand = expand, prep = prep} thy end fun prep_realizer thy = let val {realizes_eqns, typeof_eqns, defs, types, ...} = ExtractionData.get thy; val procs = maps (fst o snd) types; val rtypes = map fst types; val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); val thy' = add_syntax thy; val ctxt' = Proof_Context.init_global thy'; val rd = Proof_Syntax.read_proof thy' true false; in fn (thm, (vs, s1, s2)) => let val name = Thm.derivation_name thm; val _ = name <> "" orelse error "add_realizers: unnamed theorem"; val prop = Thm.unconstrainT thm |> Thm.prop_of |> Pattern.rewrite_term thy' (map (Logic.dest_equals o Thm.prop_of) defs) []; val vars = vars_of prop; val vars' = filter_out (fn v => member (op =) rtypes (tname_of (body_type (fastype_of v)))) vars; val shyps = maps (fn Var ((x, i), _) => if member (op =) vs x then Logic.mk_of_sort (TVar (("'" ^ x, i), []), Sign.defaultS thy') else []) vars; val T = etype_of thy' vs [] prop; val (T', thw) = Type.legacy_freeze_thaw_type (if T = nullT then nullT else map fastype_of vars' ---> T); val t = map_types thw (read_term ctxt' T' s1); val r' = freeze_thaw (condrew thy' eqns (procs @ [typeof_proc [] vs, rlz_proc])) (Const ("realizes", T --> propT --> propT) $ (if T = nullT then t else list_comb (t, vars')) $ prop); val r = Logic.list_implies (shyps, fold_rev Logic.all (map (get_var_type r') vars) r'); val prf = Proofterm.reconstruct_proof thy' r (rd s2); in (name, (vs, (t, prf))) end end; val add_realizers_i = gen_add_realizers (fn _ => fn (name, (vs, t, prf)) => (name, (vs, (t, prf)))); val add_realizers = gen_add_realizers prep_realizer; fun realizes_of thy vs t prop = let val thy' = add_syntax thy; val {realizes_eqns, typeof_eqns, defs, types, ...} = ExtractionData.get thy'; val procs = maps (rev o fst o snd) types; val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); val prop' = Pattern.rewrite_term thy' (map (Logic.dest_equals o Thm.prop_of) defs) [] prop; in freeze_thaw (condrew thy' eqns (procs @ [typeof_proc [] vs, rlz_proc])) (Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop') end; fun abs_corr_shyps thy thm vs xs prf = let val S = Sign.defaultS thy; val (ucontext, prop') = Logic.unconstrainT (Thm.shyps_of thm) (Thm.prop_of thm); val atyps = fold_types (fold_atyps (insert (op =))) (Thm.prop_of thm) []; val Ts = map_filter (fn ((v, i), _) => if member (op =) vs v then SOME (TVar (("'" ^ v, i), [])) else NONE) - (rev (Term.add_vars prop' [])); + (build_rev (Term.add_vars prop')); val cs = maps (fn T => map (pair T) S) Ts; val constraints' = map Logic.mk_of_class cs; fun typ_map T = Type.strip_sorts (map_atyps (fn U => if member (op =) atyps U then (#atyp_map ucontext) U else U) T); fun mk_hyp (T, c) = Hyp (Logic.mk_of_class (typ_map T, c)); val xs' = map (map_types typ_map) xs in prf |> Same.commit (Proofterm.map_proof_same (map_types typ_map) typ_map mk_hyp) |> fold_rev Proofterm.implies_intr_proof' (map snd (#constraints ucontext)) |> fold_rev Proofterm.forall_intr_proof' xs' |> fold_rev Proofterm.implies_intr_proof' constraints' end; (** expanding theorems / definitions **) fun add_expand_thm is_def thm thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; val name = Thm.derivation_name thm; val _ = name <> "" orelse error "add_expand_thm: unnamed theorem"; in thy |> ExtractionData.put (if is_def then {realizes_eqns = realizes_eqns, typeof_eqns = add_rule ([], Logic.dest_equals (map_types Type.strip_sorts (Thm.prop_of (Drule.abs_def thm)))) typeof_eqns, types = types, realizers = realizers, defs = insert Thm.eq_thm_prop (Thm.trim_context thm) defs, expand = expand, prep = prep} else {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = realizers, defs = defs, expand = insert (op =) name expand, prep = prep}) end; fun extraction_expand is_def = Thm.declaration_attribute (fn th => Context.mapping (add_expand_thm is_def th) I); (** types with computational content **) fun add_types tys thy = ExtractionData.map (fn {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} => {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = fold (AList.update (op =) o apfst (Sign.intern_type thy)) tys types, realizers = realizers, defs = defs, expand = expand, prep = prep}) thy; (** Pure setup **) val _ = Theory.setup (add_types [("prop", ([], NONE))] #> add_typeof_eqns ["(typeof (PROP P)) \ (Type (TYPE(Null))) \ \ \ (typeof (PROP Q)) \ (Type (TYPE('Q))) \ \ \ (typeof (PROP P \ PROP Q)) \ (Type (TYPE('Q)))", "(typeof (PROP Q)) \ (Type (TYPE(Null))) \ \ \ (typeof (PROP P \ PROP Q)) \ (Type (TYPE(Null)))", "(typeof (PROP P)) \ (Type (TYPE('P))) \ \ \ (typeof (PROP Q)) \ (Type (TYPE('Q))) \ \ \ (typeof (PROP P \ PROP Q)) \ (Type (TYPE('P \ 'Q)))", "(\x. typeof (PROP P (x))) \ (\x. Type (TYPE(Null))) \ \ \ (typeof (\x. PROP P (x))) \ (Type (TYPE(Null)))", "(\x. typeof (PROP P (x))) \ (\x. Type (TYPE('P))) \ \ \ (typeof (\x::'a. PROP P (x))) \ (Type (TYPE('a \ 'P)))", "(\x. typeof (f (x))) \ (\x. Type (TYPE('f))) \ \ \ (typeof (f)) \ (Type (TYPE('f)))"] #> add_realizes_eqns ["(typeof (PROP P)) \ (Type (TYPE(Null))) \ \ \ (realizes (r) (PROP P \ PROP Q)) \ \ \ (PROP realizes (Null) (PROP P) \ PROP realizes (r) (PROP Q))", "(typeof (PROP P)) \ (Type (TYPE('P))) \ \ \ (typeof (PROP Q)) \ (Type (TYPE(Null))) \ \ \ (realizes (r) (PROP P \ PROP Q)) \ \ \ (\x::'P. PROP realizes (x) (PROP P) \ PROP realizes (Null) (PROP Q))", "(realizes (r) (PROP P \ PROP Q)) \ \ \ (\x. PROP realizes (x) (PROP P) \ PROP realizes (r (x)) (PROP Q))", "(\x. typeof (PROP P (x))) \ (\x. Type (TYPE(Null))) \ \ \ (realizes (r) (\x. PROP P (x))) \ \ \ (\x. PROP realizes (Null) (PROP P (x)))", "(realizes (r) (\x. PROP P (x))) \ \ \ (\x. PROP realizes (r (x)) (PROP P (x)))"] #> Attrib.setup \<^binding>\extraction_expand\ (Scan.succeed (extraction_expand false)) "specify theorems to be expanded during extraction" #> Attrib.setup \<^binding>\extraction_expand_def\ (Scan.succeed (extraction_expand true)) "specify definitions to be expanded during extraction"); (**** extract program ****) val dummyt = Const ("dummy", dummyT); fun extract thm_vss thy = let val thy' = add_syntax thy; val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; val procs = maps (rev o fst o snd) types; val rtypes = map fst types; val typroc = typeof_proc []; fun expand_name ({name, ...}: Proofterm.thm_header) = if name = "" orelse member (op =) expand name then SOME "" else NONE; val prep = the_default (K I) prep thy' o Proof_Rewrite_Rules.elim_defs thy' false (map (Thm.transfer thy) defs) o Proofterm.expand_proof thy' expand_name; val rrews = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); fun find_inst prop Ts ts vs = let val rvs = relevant_vars rtypes prop; val vars = vars_of prop; val n = Int.min (length vars, length ts); fun add_args (Var ((a, i), _), t) (vs', tye) = if member (op =) rvs a then let val T = etype_of thy' vs Ts t in if T = nullT then (vs', tye) else (a :: vs', (("'" ^ a, i), T) :: tye) end else (vs', tye) in fold_rev add_args (take n vars ~~ take n ts) ([], []) end; fun mk_shyps tye = maps (fn (ixn, _) => Logic.mk_of_sort (TVar (ixn, []), Sign.defaultS thy)) tye; fun mk_sprfs cs tye = maps (fn (_, T) => Proof_Rewrite_Rules.expand_of_sort_proof thy' (map SOME cs) (T, Sign.defaultS thy)) tye; fun find (vs: string list) = Option.map snd o find_first (curry (eq_set (op =)) vs o fst); fun find' (s: string) = map_filter (fn (s', x) => if s = s' then SOME x else NONE); fun app_rlz_rews Ts vs t = strip_abs (length Ts) (freeze_thaw (condrew thy' rrews (procs @ [typroc vs, rlz_proc])) (fold (Term.abs o pair "x") Ts t)); fun realizes_null vs prop = app_rlz_rews [] vs (Const ("realizes", nullT --> propT --> propT) $ nullt $ prop); fun corr d vs ts Ts hs cs _ (PBound i) _ defs = (PBound i, defs) | corr d vs ts Ts hs cs t (Abst (s, SOME T, prf)) (Abst (_, _, prf')) defs = let val (corr_prf, defs') = corr d vs [] (T :: Ts) (dummyt :: hs) cs (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE) prf (Proofterm.incr_pboundvars 1 0 prf') defs in (Abst (s, SOME T, corr_prf), defs') end | corr d vs ts Ts hs cs t (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) defs = let val T = etype_of thy' vs Ts prop; val u = if T = nullT then (case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE) else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE); val (corr_prf, defs') = corr d vs [] (T :: Ts) (prop :: hs) (prop :: cs) u (Proofterm.incr_pboundvars 0 1 prf) (Proofterm.incr_pboundvars 0 1 prf') defs; val rlz = Const ("realizes", T --> propT --> propT) in ( if T = nullT then AbsP ("R", SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)), Proofterm.prf_subst_bounds [nullt] corr_prf) else Abst (s, SOME T, AbsP ("R", SOME (app_rlz_rews (T :: Ts) vs (rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf)), defs') end | corr d vs ts Ts hs cs t' (prf % SOME t) (prf' % _) defs = let val (Us, T) = strip_type (fastype_of1 (Ts, t)); val (corr_prf, defs') = corr d vs (t :: ts) Ts hs cs (if member (op =) rtypes (tname_of T) then t' else (case t' of SOME (u $ _) => SOME u | _ => NONE)) prf prf' defs; val u = if not (member (op =) rtypes (tname_of T)) then t else let val eT = etype_of thy' vs Ts t; val (r, Us') = if eT = nullT then (nullt, Us) else (Bound (length Us), eT :: Us); val u = list_comb (incr_boundvars (length Us') t, map Bound (length Us - 1 downto 0)); val u' = (case AList.lookup (op =) types (tname_of T) of SOME ((_, SOME f)) => f r eT u T | _ => Const ("realizes", eT --> T --> T) $ r $ u) in app_rlz_rews Ts vs (fold_rev (Term.abs o pair "x") Us' u') end in (corr_prf % SOME u, defs') end | corr d vs ts Ts hs cs t (prf1 %% prf2) (prf1' %% prf2') defs = let val prop = Proofterm.prop_of' hs prf2'; val T = etype_of thy' vs Ts prop; val (f, u, defs1) = if T = nullT then (t, NONE, defs) else (case t of SOME (f $ u) => (SOME f, SOME u, defs) | _ => let val (u, defs1) = extr d vs [] Ts hs prf2' defs in (NONE, SOME u, defs1) end) val ((corr_prf1, corr_prf2), defs2) = defs1 |> corr d vs [] Ts hs cs f prf1 prf1' ||>> corr d vs [] Ts hs cs u prf2 prf2'; in if T = nullT then (corr_prf1 %% corr_prf2, defs2) else (corr_prf1 % u %% corr_prf2, defs2) end | corr d vs ts Ts hs cs _ (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) _ defs = let val {pos, theory_name, name, prop, ...} = thm_header; val prf = Proofterm.thm_body_proof_open thm_body; val (vs', tye) = find_inst prop Ts ts vs; val shyps = mk_shyps tye; val sprfs = mk_sprfs cs tye; val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye; val T = etype_of thy' vs' [] prop; val defs' = if T = nullT then defs else snd (extr d vs ts Ts hs prf0 defs) in if T = nullT andalso realizes_null vs' prop aconv prop then (prf0, defs) else (case Symtab.lookup realizers name of NONE => (case find vs' (find' name defs') of NONE => let val _ = T = nullT orelse error "corr: internal error"; val _ = msg d ("Building correctness proof for " ^ quote name ^ (if null vs' then "" else " (relevant variables: " ^ commas_quote vs' ^ ")")); val prf' = prep (Proofterm.reconstruct_proof thy' prop prf); val (corr_prf0, defs'') = corr (d + 1) vs' [] [] [] (rev shyps) NONE prf' prf' defs'; val corr_prf = mkabsp shyps corr_prf0; val corr_prop = Proofterm.prop_of corr_prf; val corr_header = Proofterm.thm_header (serial ()) pos theory_name (corr_name name vs') corr_prop (SOME (map TVar (Term.add_tvars corr_prop [] |> rev))); val corr_prf' = Proofterm.proof_combP (Proofterm.proof_combt (PThm (corr_header, make_proof_body corr_prf), vfs_of corr_prop), map PBound (length shyps - 1 downto 0)) |> fold_rev Proofterm.forall_intr_proof' (map (get_var_type corr_prop) (vfs_of prop)) |> mkabsp shyps in (Proofterm.proof_combP (prf_subst_TVars tye' corr_prf', sprfs), (name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'') end | SOME (_, (_, prf')) => (Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs), defs')) | SOME rs => (case find vs' rs of SOME (_, prf') => (Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs), defs') | NONE => error ("corr: no realizer for instance of theorem " ^ quote name ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))))) end | corr d vs ts Ts hs cs _ (prf0 as PAxm (s, prop, SOME Ts')) _ defs = let val (vs', tye) = find_inst prop Ts ts vs; val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye in if etype_of thy' vs' [] prop = nullT andalso realizes_null vs' prop aconv prop then (prf0, defs) else case find vs' (Symtab.lookup_list realizers s) of SOME (_, prf) => (Proofterm.proof_combP (prf_subst_TVars tye' prf, mk_sprfs cs tye), defs) | NONE => error ("corr: no realizer for instance of axiom " ^ quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))) end | corr d vs ts Ts hs _ _ _ _ defs = error "corr: bad proof" and extr d vs ts Ts hs (PBound i) defs = (Bound i, defs) | extr d vs ts Ts hs (Abst (s, SOME T, prf)) defs = let val (t, defs') = extr d vs [] (T :: Ts) (dummyt :: hs) (Proofterm.incr_pboundvars 1 0 prf) defs in (Abs (s, T, t), defs') end | extr d vs ts Ts hs (AbsP (s, SOME t, prf)) defs = let val T = etype_of thy' vs Ts t; val (t, defs') = extr d vs [] (T :: Ts) (t :: hs) (Proofterm.incr_pboundvars 0 1 prf) defs in (if T = nullT then subst_bound (nullt, t) else Abs (s, T, t), defs') end | extr d vs ts Ts hs (prf % SOME t) defs = let val (u, defs') = extr d vs (t :: ts) Ts hs prf defs in (if member (op =) rtypes (tname_of (body_type (fastype_of1 (Ts, t)))) then u else u $ t, defs') end | extr d vs ts Ts hs (prf1 %% prf2) defs = let val (f, defs') = extr d vs [] Ts hs prf1 defs; val prop = Proofterm.prop_of' hs prf2; val T = etype_of thy' vs Ts prop in if T = nullT then (f, defs') else let val (t, defs'') = extr d vs [] Ts hs prf2 defs' in (f $ t, defs'') end end | extr d vs ts Ts hs (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) defs = let val {pos, theory_name, name = s, prop, ...} = thm_header; val prf = Proofterm.thm_body_proof_open thm_body; val (vs', tye) = find_inst prop Ts ts vs; val shyps = mk_shyps tye; val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye in case Symtab.lookup realizers s of NONE => (case find vs' (find' s defs) of NONE => let val _ = msg d ("Extracting " ^ quote s ^ (if null vs' then "" else " (relevant variables: " ^ commas_quote vs' ^ ")")); val prf' = prep (Proofterm.reconstruct_proof thy' prop prf); val (t, defs') = extr (d + 1) vs' [] [] [] prf' defs; val (corr_prf, defs'') = corr (d + 1) vs' [] [] [] (rev shyps) (SOME t) prf' prf' defs'; val nt = Envir.beta_norm t; val args = filter_out (fn v => member (op =) rtypes (tname_of (body_type (fastype_of v)))) (vfs_of prop); val args' = filter (fn v => Logic.occs (v, nt)) args; val t' = mkabs args' nt; val T = fastype_of t'; val cname = extr_name s vs'; val c = Const (cname, T); val u = mkabs args (list_comb (c, args')); val eqn = Logic.mk_equals (c, t'); val rlz = Const ("realizes", fastype_of nt --> propT --> propT); val lhs = app_rlz_rews [] vs' (rlz $ nt $ prop); val rhs = app_rlz_rews [] vs' (rlz $ list_comb (c, args') $ prop); val f = app_rlz_rews [] vs' (Abs ("x", T, rlz $ list_comb (Bound 0, args') $ prop)); val corr_prf' = mkabsp shyps (change_types [] Proofterm.equal_elim_axm %> lhs %> rhs %% (change_types [propT] Proofterm.symmetric_axm %> rhs %> lhs %% (change_types [T, propT] Proofterm.combination_axm %> f %> f %> c %> t' %% (change_types [T --> propT] Proofterm.reflexive_axm %> f) %% PAxm (Thm.def_name cname, eqn, SOME (map TVar (Term.add_tvars eqn [] |> rev))))) %% corr_prf); val corr_prop = Proofterm.prop_of corr_prf'; val corr_header = Proofterm.thm_header (serial ()) pos theory_name (corr_name s vs') corr_prop (SOME (map TVar (Term.add_tvars corr_prop [] |> rev))); val corr_prf'' = Proofterm.proof_combP (Proofterm.proof_combt (PThm (corr_header, make_proof_body corr_prf), vfs_of corr_prop), map PBound (length shyps - 1 downto 0)) |> fold_rev Proofterm.forall_intr_proof' (map (get_var_type corr_prop) (vfs_of prop)) |> mkabsp shyps in (subst_TVars tye' u, (s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'') end | SOME ((_, u), _) => (subst_TVars tye' u, defs)) | SOME rs => (case find vs' rs of SOME (t, _) => (subst_TVars tye' t, defs) | NONE => error ("extr: no realizer for instance of theorem " ^ quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts)))))) end | extr d vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) defs = let val (vs', tye) = find_inst prop Ts ts vs; val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye in case find vs' (Symtab.lookup_list realizers s) of SOME (t, _) => (subst_TVars tye' t, defs) | NONE => error ("extr: no realizer for instance of axiom " ^ quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))) end | extr d vs ts Ts hs _ defs = error "extr: bad proof"; fun prep_thm vs raw_thm = let val thm = Thm.transfer thy raw_thm; val prop = Thm.prop_of thm; val prf = Thm.proof_of thm; val name = Thm.derivation_name thm; val _ = name <> "" orelse error "extraction: unnamed theorem"; val _ = etype_of thy' vs [] prop <> nullT orelse error ("theorem " ^ quote name ^ " has no computational content") in Proofterm.reconstruct_proof thy' prop prf end; val defs = fold (fn (thm, vs) => snd o (extr 0 vs [] [] [] o prep_thm vs) thm) thm_vss []; fun add_def (s, (vs, ((t, u)))) thy = let val ft = Type.legacy_freeze t; val fu = Type.legacy_freeze u; val head = head_of (strip_abs_body fu); val b = Binding.qualified_name (extr_name s vs); in thy |> Sign.add_consts [(b, fastype_of ft, NoSyn)] |> Global_Theory.add_defs false [((Binding.qualified_name (Thm.def_name (extr_name s vs)), Logic.mk_equals (head, ft)), [])] |-> (fn [def_thm] => Spec_Rules.add_global b Spec_Rules.equational [Thm.term_of (Thm.lhs_of def_thm)] [def_thm] #> Code.declare_default_eqns_global [(def_thm, true)]) end; fun add_corr (s, (vs, prf)) thy = let val corr_prop = Proofterm.prop_of prf; in thy |> Global_Theory.store_thm (Binding.qualified_name (corr_name s vs), Thm.varifyT_global (funpow (length (vars_of corr_prop)) (Thm.forall_elim_var 0) (Thm.forall_intr_frees (Proof_Checker.thm_of_proof thy (fst (Proofterm.freeze_thaw_prf prf)))))) |> snd end; fun add_def_and_corr (s, (vs, ((t, u), (prf, _)))) thy = if is_none (Sign.const_type thy (extr_name s vs)) then thy |> not (t = nullt) ? add_def (s, (vs, ((t, u)))) |> add_corr (s, (vs, prf)) else thy; in thy |> Sign.root_path |> fold_rev add_def_and_corr defs |> Sign.restore_naming thy end; val etype_of = etype_of o add_syntax; end; diff --git a/src/Pure/Proof/proof_checker.ML b/src/Pure/Proof/proof_checker.ML --- a/src/Pure/Proof/proof_checker.ML +++ b/src/Pure/Proof/proof_checker.ML @@ -1,152 +1,152 @@ (* Title: Pure/Proof/proof_checker.ML Author: Stefan Berghofer, TU Muenchen Simple proof checker based only on the core inference rules of Isabelle/Pure. *) signature PROOF_CHECKER = sig val thm_of_proof : theory -> Proofterm.proof -> thm end; structure Proof_Checker : PROOF_CHECKER = struct (***** construct a theorem out of a proof term *****) fun lookup_thm thy = - let val tab = fold_rev Symtab.update (Global_Theory.all_thms_of thy true) Symtab.empty in + let val tab = Symtab.build (Global_Theory.all_thms_of thy true |> fold_rev Symtab.update) in fn s => (case Symtab.lookup tab s of NONE => error ("Unknown theorem " ^ quote s) | SOME thm => thm) end; val beta_eta_convert = Conv.fconv_rule Drule.beta_eta_conversion; (* equality modulo renaming of type variables *) fun is_equal t t' = let val atoms = fold_types (fold_atyps (insert (op =))) t []; val atoms' = fold_types (fold_atyps (insert (op =))) t' [] in length atoms = length atoms' andalso map_types (map_atyps (the o AList.lookup (op =) (atoms ~~ atoms'))) t aconv t' end; fun pretty_prf thy vs Hs prf = let val prf' = prf |> Proofterm.prf_subst_bounds (map Free vs) |> Proofterm.prf_subst_pbounds (map (Hyp o Thm.prop_of) Hs) in (Proof_Syntax.pretty_proof (Syntax.init_pretty_global thy) prf', Syntax.pretty_term_global thy (Proofterm.prop_of prf')) end; fun pretty_term thy vs _ t = let val t' = subst_bounds (map Free vs, t) in (Syntax.pretty_term_global thy t', Syntax.pretty_typ_global thy (fastype_of t')) end; fun appl_error thy prt vs Hs s f a = let val (pp_f, pp_fT) = pretty_prf thy vs Hs f; val (pp_a, pp_aT) = prt thy vs Hs a in error (cat_lines [s, "", Pretty.string_of (Pretty.block [Pretty.str "Operator:", Pretty.brk 2, pp_f, Pretty.str " ::", Pretty.brk 1, pp_fT]), Pretty.string_of (Pretty.block [Pretty.str "Operand:", Pretty.brk 3, pp_a, Pretty.str " ::", Pretty.brk 1, pp_aT]), ""]) end; fun thm_of_proof thy = let val lookup = lookup_thm thy in fn prf => let val prf_names = Proofterm.fold_proof_terms Term.declare_term_frees prf Name.context; fun thm_of_atom thm Ts = let - val tvars = Term.add_tvars (Thm.full_prop_of thm) [] |> rev; - val (fmap, thm') = Thm.varifyT_global' [] thm; + val tvars = build_rev (Term.add_tvars (Thm.full_prop_of thm)); + val (fmap, thm') = Thm.varifyT_global' Term_Subst.TFrees.empty thm; val ctye = (tvars @ map (fn ((_, S), ixn) => (ixn, S)) fmap ~~ map (Thm.global_ctyp_of thy) Ts); in - Thm.instantiate (ctye, []) (forall_intr_vars (Thm.forall_intr_frees thm')) + Thm.instantiate (ctye, []) (Thm.forall_intr_vars (Thm.forall_intr_frees thm')) end; fun thm_of _ _ (PThm ({name, prop = prop', types = SOME Ts, ...}, _)) = let val thm = Thm.unconstrainT (Drule.implies_intr_hyps (lookup name)); val prop = Thm.prop_of thm; val _ = if is_equal prop prop' then () else error ("Duplicate use of theorem name " ^ quote name ^ "\n" ^ Syntax.string_of_term_global thy prop ^ "\n\n" ^ Syntax.string_of_term_global thy prop'); in thm_of_atom thm Ts end | thm_of _ _ (PAxm (name, _, SOME Ts)) = thm_of_atom (Thm.axiom thy name) Ts | thm_of _ Hs (PBound i) = nth Hs i | thm_of (vs, names) Hs (Abst (s, SOME T, prf)) = let val (x, names') = Name.variant s names; val thm = thm_of ((x, T) :: vs, names') Hs prf in Thm.forall_intr (Thm.global_cterm_of thy (Free (x, T))) thm end | thm_of (vs, names) Hs (prf % SOME t) = let val thm = thm_of (vs, names) Hs prf; val ct = Thm.global_cterm_of thy (Term.subst_bounds (map Free vs, t)); in Thm.forall_elim ct thm handle THM (s, _, _) => appl_error thy pretty_term vs Hs s prf t end | thm_of (vs, names) Hs (AbsP (_, SOME t, prf)) = let val ct = Thm.global_cterm_of thy (Term.subst_bounds (map Free vs, t)); val thm = thm_of (vs, names) (Thm.assume ct :: Hs) prf; in Thm.implies_intr ct thm end | thm_of vars Hs (prf %% prf') = let val thm = beta_eta_convert (thm_of vars Hs prf); val thm' = beta_eta_convert (thm_of vars Hs prf'); in Thm.implies_elim thm thm' handle THM (s, _, _) => appl_error thy pretty_prf (fst vars) Hs s prf prf' end | thm_of _ _ (Hyp t) = Thm.assume (Thm.global_cterm_of thy t) | thm_of _ _ (PClass (T, c)) = if Sign.of_sort thy (T, [c]) then Thm.of_class (Thm.global_ctyp_of thy T, c) else error ("thm_of_proof: bad PClass proof " ^ Syntax.string_of_term_global thy (Logic.mk_of_class (T, c))) | thm_of _ _ _ = error "thm_of_proof: partial proof term"; in beta_eta_convert (thm_of ([], prf_names) [] prf) end end; end; diff --git a/src/Pure/Proof/proof_rewrite_rules.ML b/src/Pure/Proof/proof_rewrite_rules.ML --- a/src/Pure/Proof/proof_rewrite_rules.ML +++ b/src/Pure/Proof/proof_rewrite_rules.ML @@ -1,394 +1,394 @@ (* Title: Pure/Proof/proof_rewrite_rules.ML Author: Stefan Berghofer, TU Muenchen Simplification functions for proof terms involving meta level rules. *) signature PROOF_REWRITE_RULES = sig val rew : bool -> typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option val rprocs : bool -> (typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option) list val rewrite_terms : (term -> term) -> Proofterm.proof -> Proofterm.proof val elim_defs : theory -> bool -> thm list -> Proofterm.proof -> Proofterm.proof val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof val hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof val un_hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof val expand_of_sort_proof : theory -> term option list -> typ * sort -> Proofterm.proof list val expand_of_class_proof : theory -> term option list -> typ * class -> proof val expand_of_class : theory -> typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option val standard_preproc: (proof * proof) list -> theory -> proof -> proof end; structure Proof_Rewrite_Rules : PROOF_REWRITE_RULES = struct fun rew b _ _ = let fun ?? x = if b then SOME x else NONE; fun ax (prf as PAxm (s, prop, _)) Ts = if b then PAxm (s, prop, SOME Ts) else prf; fun ty T = if b then (case T of Type (_, [Type (_, [U, _]), _]) => SOME U | _ => NONE) else NONE; val equal_intr_axm = ax Proofterm.equal_intr_axm []; val equal_elim_axm = ax Proofterm.equal_elim_axm []; val symmetric_axm = ax Proofterm.symmetric_axm [propT]; fun rew' (PThm ({name = "Pure.protectD", ...}, _) % _ %% (PThm ({name = "Pure.protectI", ...}, _) % _ %% prf)) = SOME prf | rew' (PThm ({name = "Pure.conjunctionD1", ...}, _) % _ % _ %% (PThm ({name = "Pure.conjunctionI", ...}, _) % _ % _ %% prf %% _)) = SOME prf | rew' (PThm ({name = "Pure.conjunctionD2", ...}, _) % _ % _ %% (PThm ({name = "Pure.conjunctionI", ...}, _) % _ % _ %% _ %% prf)) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_intr", _, _) % A % B %% prf1 %% prf2)) = SOME (equal_intr_axm % B % A %% prf2 %% prf1) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% ((tg as PThm ({name = "Pure.protectI", ...}, _)) % _ %% prf2)) = SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1)) %% ((tg as PThm ({name = "Pure.protectI", ...}, _)) % _ %% prf2)) = SOME (tg %> B %% (equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf1) %% prf2)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) = let val _ $ A $ C = Envir.beta_norm X; val _ $ B $ D = Envir.beta_norm Y in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? B, Proofterm.equal_elim_axm %> C %> D %% Proofterm.incr_pboundvars 2 0 prf2 %% (PBound 1 %% (equal_elim_axm %> B %> A %% (Proofterm.symmetric_axm % ?? A % ?? B %% Proofterm.incr_pboundvars 2 0 prf1) %% PBound 0))))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2))) = let val _ $ A $ C = Envir.beta_norm Y; val _ $ B $ D = Envir.beta_norm X in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? A, equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% Proofterm.incr_pboundvars 2 0 prf2) %% (PBound 1 %% (equal_elim_axm %> A %> B %% Proofterm.incr_pboundvars 2 0 prf1 %% PBound 0))))) end | rew' (PAxm ("Pure.combination", _, _) % SOME (imp as (imp' as Const ("Pure.imp", T)) $ X) % _ % Y % Z %% (PAxm ("Pure.reflexive", _, _) % _) %% prf) = SOME (ax Proofterm.combination_axm [propT, propT] %> imp % ?? imp % Y % Z %% (ax Proofterm.combination_axm [propT, propT --> propT] %> imp' % ?? imp' % ?? X % ?? X %% (ax Proofterm.reflexive_axm [T] % ?? imp') %% (ax Proofterm.reflexive_axm [propT] % ?? X)) %% prf) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% (PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf))) = let val Const (_, T) $ P = Envir.beta_norm X; val _ $ Q = Envir.beta_norm Y; in SOME (AbsP ("H", ?? X, Abst ("x", ty T, equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %% (Proofterm.incr_pboundvars 1 1 prf %> Bound 0) %% (PBound 0 %> Bound 0)))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% (PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf)))) = let val Const (_, T) $ P = Envir.beta_norm X; val _ $ Q = Envir.beta_norm Y; val t = incr_boundvars 1 P $ Bound 0; val u = incr_boundvars 1 Q $ Bound 0 in SOME (AbsP ("H", ?? X, Abst ("x", ty T, equal_elim_axm %> t %> u %% (symmetric_axm % ?? u % ?? t %% (Proofterm.incr_pboundvars 1 1 prf %> Bound 0)) %% (PBound 0 %> Bound 0)))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %% (PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2) %% prf3) = SOME (equal_elim_axm %> B %> C %% prf2 %% (equal_elim_axm %> A %> B %% prf1 %% prf3)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2)) %% prf3) = SOME (equal_elim_axm %> B %> C %% (symmetric_axm % ?? C % ?? B %% prf1) %% (equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf2) %% prf3)) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _)) %% prf) = SOME prf | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% prf)) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) = SOME (equal_elim_axm %> C %> D %% prf2 %% (equal_elim_axm %> A %> C %% prf3 %% (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) = SOME (equal_elim_axm %> A %> B %% prf1 %% (equal_elim_axm %> C %> A %% (symmetric_axm % ?? A % ?? C %% prf3) %% (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) = SOME (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% (equal_elim_axm %> B %> D %% prf3 %% (equal_elim_axm %> A %> B %% prf1 %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) = SOME (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% (equal_elim_axm %> D %> B %% (symmetric_axm % ?? B % ?? D %% prf3) %% (equal_elim_axm %> C %> D %% prf2 %% prf4))) | rew' ((prf as PAxm ("Pure.combination", _, _) % SOME ((eq as Const ("Pure.eq", T)) $ t) % _ % _ % _) %% (PAxm ("Pure.reflexive", _, _) % _)) = let val (U, V) = (case T of Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT)) in SOME (prf %% (ax Proofterm.combination_axm [U, V] %> eq % ?? eq % ?? t % ?? t %% (ax Proofterm.reflexive_axm [T] % ?? eq) %% (ax Proofterm.reflexive_axm [U] % ?? t))) end | rew' _ = NONE; in rew' #> Option.map (rpair Proofterm.no_skel) end; fun rprocs b = [rew b]; val _ = Theory.setup (fold Proofterm.add_prf_rproc (rprocs false)); (**** apply rewriting function to all terms in proof ****) fun rewrite_terms r = let fun rew_term Ts t = let val frees = map Free (Name.invent (Term.declare_term_frees t Name.context) "xa" (length Ts) ~~ Ts); val t' = r (subst_bounds (frees, t)); fun strip [] t = t | strip (_ :: xs) (Abs (_, _, t)) = strip xs t; in strip Ts (fold lambda frees t') end; fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2 | rew Ts (prf % SOME t) = rew Ts prf % SOME (rew_term Ts t) | rew Ts (Abst (s, SOME T, prf)) = Abst (s, SOME T, rew (T :: Ts) prf) | rew Ts (AbsP (s, SOME t, prf)) = AbsP (s, SOME (rew_term Ts t), rew Ts prf) | rew _ prf = prf in rew [] end; (**** eliminate definitions in proof ****) fun vars_of t = rev (fold_aterms (fn v as Var _ => insert (op =) v | _ => I) t []); fun insert_refl defs Ts (prf1 %% prf2) = let val (prf1', b) = insert_refl defs Ts prf1 in if b then (prf1', true) else (prf1' %% fst (insert_refl defs Ts prf2), false) end | insert_refl defs Ts (Abst (s, SOME T, prf)) = (Abst (s, SOME T, fst (insert_refl defs (T :: Ts) prf)), false) | insert_refl defs Ts (AbsP (s, t, prf)) = (AbsP (s, t, fst (insert_refl defs Ts prf)), false) | insert_refl defs Ts prf = (case Proofterm.strip_combt prf of (PThm ({name = s, prop, types = SOME Ts, ...}, _), ts) => if member (op =) defs s then let val vs = vars_of prop; - val tvars = Term.add_tvars prop [] |> rev; + val tvars = build_rev (Term.add_tvars prop); val (_, rhs) = Logic.dest_equals (Logic.strip_imp_concl prop); val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts) (fold_rev (fn x => fn b => Abs ("", dummyT, abstract_over (x, b))) vs rhs), map the ts); in (Proofterm.change_types (SOME [fastype_of1 (Ts, rhs')]) Proofterm.reflexive_axm %> rhs', true) end else (prf, false) | (_, []) => (prf, false) | (prf', ts) => (Proofterm.proof_combt' (fst (insert_refl defs Ts prf'), ts), false)); fun elim_defs thy r defs prf = let val defs' = map (Logic.dest_equals o map_types Type.strip_sorts o Thm.prop_of o Drule.abs_def) defs; val defnames = map Thm.derivation_name defs; val prf' = if r then let val cnames = map (fst o dest_Const o fst) defs'; val expand = Proofterm.fold_proof_atoms true (fn PThm ({serial, name, prop, ...}, _) => if member (op =) defnames name orelse not (exists_Const (member (op =) cnames o #1) prop) then I else Inttab.update (serial, "") | _ => I) [prf] Inttab.empty; in Proofterm.expand_proof thy (Inttab.lookup expand o #serial) prf end else prf; in rewrite_terms (Pattern.rewrite_term thy defs' []) (fst (insert_refl defnames [] prf')) end; (**** eliminate all variables that don't occur in the proposition ****) fun elim_vars mk_default prf = let val prop = Proofterm.prop_of prf; - val tv = Term.add_vars prop []; - val tf = Term.add_frees prop []; + val tv = Term_Subst.Vars.build (Term_Subst.add_vars prop); + val tf = Term_Subst.Frees.build (Term_Subst.add_frees prop); - fun hidden_variable (Var v) = not (member (op =) tv v) - | hidden_variable (Free f) = not (member (op =) tf f) + fun hidden_variable (Var v) = not (Term_Subst.Vars.defined tv v) + | hidden_variable (Free f) = not (Term_Subst.Frees.defined tf f) | hidden_variable _ = false; fun mk_default' T = fold_rev (Term.abs o pair "x") (binder_types T) (mk_default (body_type T)); fun elim_varst (t $ u) = elim_varst t $ elim_varst u | elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t) - | elim_varst (t as Free (x, T)) = if member (op =) tf (x, T) then t else mk_default' T - | elim_varst (t as Var (xi, T)) = if member (op =) tv (xi, T) then t else mk_default' T + | elim_varst (t as Free (x, T)) = if Term_Subst.Frees.defined tf (x, T) then t else mk_default' T + | elim_varst (t as Var (xi, T)) = if Term_Subst.Vars.defined tv (xi, T) then t else mk_default' T | elim_varst t = t; in Proofterm.map_proof_terms (fn t => if Term.exists_subterm hidden_variable t then Envir.beta_norm (elim_varst t) else t) I prf end; (**** convert between hhf and non-hhf form ****) fun hhf_proof P Q prf = let val params = Logic.strip_params Q; val Hs = Logic.strip_assums_hyp P; val Hs' = Logic.strip_assums_hyp Q; val k = length Hs; val l = length params; fun mk_prf i j Hs Hs' (Const ("Pure.all", _) $ Abs (_, _, P)) prf = mk_prf i (j - 1) Hs Hs' P (prf %> Bound j) | mk_prf i j (H :: Hs) (H' :: Hs') (Const ("Pure.imp", _) $ _ $ P) prf = mk_prf (i - 1) j Hs Hs' P (prf %% un_hhf_proof H' H (PBound i)) | mk_prf _ _ _ _ _ prf = prf in prf |> Proofterm.incr_pboundvars k l |> mk_prf (k - 1) (l - 1) Hs Hs' P |> fold_rev (fn P => fn prf => AbsP ("H", SOME P, prf)) Hs' |> fold_rev (fn (s, T) => fn prf => Abst (s, SOME T, prf)) params end and un_hhf_proof P Q prf = let val params = Logic.strip_params Q; val Hs = Logic.strip_assums_hyp P; val Hs' = Logic.strip_assums_hyp Q; val k = length Hs; val l = length params; fun mk_prf (Const ("Pure.all", _) $ Abs (s, T, P)) prf = Abst (s, SOME T, mk_prf P prf) | mk_prf (Const ("Pure.imp", _) $ P $ Q) prf = AbsP ("H", SOME P, mk_prf Q prf) | mk_prf _ prf = prf in prf |> Proofterm.incr_pboundvars k l |> fold (fn i => fn prf => prf %> Bound i) (l - 1 downto 0) |> fold (fn ((H, H'), i) => fn prf => prf %% hhf_proof H' H (PBound i)) (Hs ~~ Hs' ~~ (k - 1 downto 0)) |> mk_prf Q end; (**** expand PClass proofs ****) fun expand_of_sort_proof thy hyps (T, S) = let val of_class_hyps = map (fn SOME t => try Logic.dest_of_class t | NONE => NONE) hyps; fun of_class_index p = (case find_index (fn SOME h => h = p | NONE => false) of_class_hyps of ~1 => raise Fail "expand_of_class_proof: missing class hypothesis" | i => PBound i); val sorts = AList.coalesce (op =) (rev (map_filter I of_class_hyps)); fun get_sort T = the_default [] (AList.lookup (op =) sorts T); val subst = map_atyps (fn T as TVar (ai, _) => TVar (ai, get_sort T) | T as TFree (a, _) => TFree (a, get_sort T)); fun reconstruct prop = Proofterm.reconstruct_proof thy prop #> Proofterm.expand_proof thy Proofterm.expand_name_empty #> Same.commit (Proofterm.map_proof_same Same.same Same.same of_class_index); in map2 reconstruct (Logic.mk_of_sort (T, S)) (Proofterm.of_sort_proof (Sign.classes_of thy) (Thm.classrel_proof thy) (Thm.arity_proof thy) (PClass o apfst Type.strip_sorts) (subst T, S)) end; fun expand_of_class_proof thy hyps (T, c) = hd (expand_of_sort_proof thy hyps (T, [c])); fun expand_of_class thy (_: typ list) hyps (PClass (T, c)) = SOME (expand_of_class_proof thy hyps (T, c), Proofterm.no_skel) | expand_of_class _ _ _ _ = NONE; (* standard preprocessor *) fun standard_preproc rews thy = Proofterm.rewrite_proof thy (rews, rprocs true @ [expand_of_class thy]); end; diff --git a/src/Pure/Syntax/ast.ML b/src/Pure/Syntax/ast.ML --- a/src/Pure/Syntax/ast.ML +++ b/src/Pure/Syntax/ast.ML @@ -1,247 +1,247 @@ (* Title: Pure/Syntax/ast.ML Author: Markus Wenzel, TU Muenchen Abstract syntax trees, translation rules, matching and normalization of asts. *) signature AST = sig datatype ast = Constant of string | Variable of string | Appl of ast list val mk_appl: ast -> ast list -> ast exception AST of string * ast list val pretty_ast: ast -> Pretty.T val pretty_rule: ast * ast -> Pretty.T val strip_positions: ast -> ast val head_of_rule: ast * ast -> string val rule_error: ast * ast -> string option val fold_ast: string -> ast list -> ast val fold_ast_p: string -> ast list * ast -> ast val unfold_ast: string -> ast -> ast list val unfold_ast_p: string -> ast -> ast list * ast val trace: bool Config.T val stats: bool Config.T val normalize: Proof.context -> (string -> (ast * ast) list) -> ast -> ast end; structure Ast: AST = struct (** abstract syntax trees **) (*asts come in two flavours: - ordinary asts representing terms and typs: Variables are (often) treated like Constants; - patterns used as lhs and rhs in rules: Variables are placeholders for proper asts*) datatype ast = Constant of string | (*"not", "_abs", "fun"*) Variable of string | (*x, ?x, 'a, ?'a*) Appl of ast list; (*(f x y z), ("fun" 'a 'b), ("_abs" x t)*) (*the list of subasts of an Appl node has to contain at least 2 elements, i.e. there are no empty asts or nullary applications; use mk_appl for convenience*) fun mk_appl f [] = f | mk_appl f args = Appl (f :: args); (*exception for system errors involving asts*) exception AST of string * ast list; (** print asts in a LISP-like style **) fun pretty_ast (Constant a) = Pretty.quote (Pretty.str a) | pretty_ast (Variable x) = (case Term_Position.decode x of SOME pos => Term_Position.pretty pos | NONE => Pretty.str x) | pretty_ast (Appl asts) = Pretty.enclose "(" ")" (Pretty.breaks (map pretty_ast asts)); fun pretty_rule (lhs, rhs) = Pretty.block [pretty_ast lhs, Pretty.str " \", Pretty.brk 1, pretty_ast rhs]; val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_ast); (* strip_positions *) fun strip_positions (Appl ((t as Constant c) :: u :: (v as Variable x) :: asts)) = if member (op =) Term_Position.markers c andalso is_some (Term_Position.decode x) then mk_appl (strip_positions u) (map strip_positions asts) else Appl (map strip_positions (t :: u :: v :: asts)) | strip_positions (Appl asts) = Appl (map strip_positions asts) | strip_positions ast = ast; (* head_of_ast and head_of_rule *) fun head_of_ast (Constant a) = a | head_of_ast (Appl (Constant a :: _)) = a | head_of_ast _ = ""; fun head_of_rule (lhs, _) = head_of_ast lhs; (** check translation rules **) fun rule_error (lhs, rhs) = let fun add_vars (Constant _) = I | add_vars (Variable x) = cons x | add_vars (Appl asts) = fold add_vars asts; val lvars = add_vars lhs []; val rvars = add_vars rhs []; in if has_duplicates (op =) lvars then SOME "duplicate vars in lhs" else if not (subset (op =) (rvars, lvars)) then SOME "rhs contains extra variables" else NONE end; (** ast translation utilities **) (* fold asts *) fun fold_ast _ [] = raise Match | fold_ast _ [y] = y | fold_ast c (x :: xs) = Appl [Constant c, x, fold_ast c xs]; fun fold_ast_p c = uncurry (fold_rev (fn x => fn xs => Appl [Constant c, x, xs])); (* unfold asts *) fun unfold_ast c (y as Appl [Constant c', x, xs]) = if c = c' then x :: unfold_ast c xs else [y] | unfold_ast _ y = [y]; fun unfold_ast_p c (y as Appl [Constant c', x, xs]) = if c = c' then apfst (cons x) (unfold_ast_p c xs) else ([], y) | unfold_ast_p _ y = ([], y); (** normalization of asts **) (* match *) fun match ast pat = let exception NO_MATCH; fun mtch (Constant a) (Constant b) env = if a = b then env else raise NO_MATCH | mtch (Variable a) (Constant b) env = if a = b then env else raise NO_MATCH | mtch ast (Variable x) env = Symtab.update (x, ast) env | mtch (Appl asts) (Appl pats) env = mtch_lst asts pats env | mtch _ _ _ = raise NO_MATCH and mtch_lst (ast :: asts) (pat :: pats) env = mtch_lst asts pats (mtch ast pat env) | mtch_lst [] [] env = env | mtch_lst _ _ _ = raise NO_MATCH; val (head, args) = (case (ast, pat) of (Appl asts, Appl pats) => let val a = length asts and p = length pats in if a > p then (Appl (take p asts), drop p asts) else (ast, []) end | _ => (ast, [])); in - SOME (mtch head pat Symtab.empty, args) handle NO_MATCH => NONE + SOME (Symtab.build (mtch head pat), args) handle NO_MATCH => NONE end; (* normalize *) val trace = Config.declare_bool ("syntax_ast_trace", \<^here>) (K false); val stats = Config.declare_bool ("syntax_ast_stats", \<^here>) (K false); fun message head body = Pretty.string_of (Pretty.block [Pretty.str head, Pretty.brk 1, body]); (*the normalizer works yoyo-like: top-down, bottom-up, top-down, ...*) fun normalize ctxt get_rules pre_ast = let val trace = Config.get ctxt trace; val stats = Config.get ctxt stats; val passes = Unsynchronized.ref 0; val failed_matches = Unsynchronized.ref 0; val changes = Unsynchronized.ref 0; fun subst _ (ast as Constant _) = ast | subst env (Variable x) = the (Symtab.lookup env x) | subst env (Appl asts) = Appl (map (subst env) asts); fun try_rules ((lhs, rhs) :: pats) ast = (case match ast lhs of SOME (env, args) => (Unsynchronized.inc changes; SOME (mk_appl (subst env rhs) args)) | NONE => (Unsynchronized.inc failed_matches; try_rules pats ast)) | try_rules [] _ = NONE; val try_headless_rules = try_rules (get_rules ""); fun try ast a = (case try_rules (get_rules a) ast of NONE => try_headless_rules ast | some => some); fun rewrite (ast as Constant a) = try ast a | rewrite (ast as Variable a) = try ast a | rewrite (ast as Appl (Constant a :: _)) = try ast a | rewrite (ast as Appl (Variable a :: _)) = try ast a | rewrite ast = try_headless_rules ast; fun rewrote old_ast new_ast = if trace then tracing (message "rewrote:" (pretty_rule (old_ast, new_ast))) else (); fun norm_root ast = (case rewrite ast of SOME new_ast => (rewrote ast new_ast; norm_root new_ast) | NONE => ast); fun norm ast = (case norm_root ast of Appl sub_asts => let val old_changes = ! changes; val new_ast = Appl (map norm sub_asts); in if old_changes = ! changes then new_ast else norm_root new_ast end | atomic_ast => atomic_ast); fun normal ast = let val old_changes = ! changes; val new_ast = norm ast; in Unsynchronized.inc passes; if old_changes = ! changes then new_ast else normal new_ast end; val _ = if trace then tracing (message "pre:" (pretty_ast pre_ast)) else (); val post_ast = normal pre_ast; val _ = if trace orelse stats then tracing (message "post:" (pretty_ast post_ast) ^ "\nnormalize: " ^ string_of_int (! passes) ^ " passes, " ^ string_of_int (! changes) ^ " changes, " ^ string_of_int (! failed_matches) ^ " matches failed") else (); in post_ast end; end; diff --git a/src/Pure/System/options.ML b/src/Pure/System/options.ML --- a/src/Pure/System/options.ML +++ b/src/Pure/System/options.ML @@ -1,223 +1,223 @@ (* Title: Pure/System/options.ML Author: Makarius System options with external string representation. *) signature OPTIONS = sig val boolT: string val intT: string val realT: string val stringT: string val unknownT: string type T val empty: T val dest: T -> (string * Position.T) list val typ: T -> string -> string val bool: T -> string -> bool val int: T -> string -> int val real: T -> string -> real val seconds: T -> string -> Time.time val string: T -> string -> string val put_bool: string -> bool -> T -> T val put_int: string -> int -> T -> T val put_real: string -> real -> T -> T val put_string: string -> string -> T -> T val declare: {pos: Position.T, name: string, typ: string, value: string} -> T -> T val update: string -> string -> T -> T val encode: T XML.Encode.T val decode: T XML.Decode.T val default: unit -> T val default_typ: string -> string val default_bool: string -> bool val default_int: string -> int val default_real: string -> real val default_seconds: string -> Time.time val default_string: string -> string val default_put_bool: string -> bool -> unit val default_put_int: string -> int -> unit val default_put_real: string -> real -> unit val default_put_string: string -> string -> unit val get_default: string -> string val put_default: string -> string -> unit val set_default: T -> unit val reset_default: unit -> unit val load_default: unit -> unit end; structure Options: OPTIONS = struct (* representation *) val boolT = "bool"; val intT = "int"; val realT = "real"; val stringT = "string"; val unknownT = "unknown"; datatype T = Options of {pos: Position.T, typ: string, value: string} Symtab.table; val empty = Options Symtab.empty; fun dest (Options tab) = Symtab.fold (fn (name, {pos, ...}) => cons (name, pos)) tab [] |> sort_by #1; (* check *) fun check_name (Options tab) name = let val opt = Symtab.lookup tab name in if is_some opt andalso #typ (the opt) <> unknownT then the opt else error ("Unknown system option " ^ quote name) end; fun check_type options name typ = let val opt = check_name options name in if #typ opt = typ then opt else error ("Ill-typed system option " ^ quote name ^ " : " ^ #typ opt ^ " vs. " ^ typ) end; (* typ *) fun typ options name = #typ (check_name options name); (* basic operations *) fun put T print name x (options as Options tab) = let val opt = check_type options name T in Options (Symtab.update (name, {pos = #pos opt, typ = #typ opt, value = print x}) tab) end; fun get T parse options name = let val opt = check_type options name T in (case parse (#value opt) of SOME x => x | NONE => error ("Malformed value for system option " ^ quote name ^ " : " ^ T ^ " =\n" ^ quote (#value opt))) end; (* internal lookup and update *) val bool = get boolT (try Value.parse_bool); val int = get intT (try Value.parse_int); val real = get realT (try Value.parse_real); val seconds = Time.fromReal oo real; val string = get stringT SOME; val put_bool = put boolT Value.print_bool; val put_int = put intT Value.print_int; val put_real = put realT Value.print_real; val put_string = put stringT I; (* external updates *) fun check_value options name = let val opt = check_name options name in if #typ opt = boolT then ignore (bool options name) else if #typ opt = intT then ignore (int options name) else if #typ opt = realT then ignore (real options name) else if #typ opt = stringT then ignore (string options name) else () end; fun declare {pos, name, typ, value} (Options tab) = let val options' = (case Symtab.lookup tab name of SOME other => error ("Duplicate declaration of system option " ^ quote name ^ Position.here pos ^ Position.here (#pos other)) | NONE => Options (Symtab.update (name, {pos = pos, typ = typ, value = value}) tab)); val _ = typ = boolT orelse typ = intT orelse typ = realT orelse typ = stringT orelse error ("Unknown type for system option " ^ quote name ^ " : " ^ quote typ ^ Position.here pos); val _ = check_value options' name; in options' end; fun update name value (options as Options tab) = let val opt = check_name options name; val options' = Options (Symtab.update (name, {pos = #pos opt, typ = #typ opt, value = value}) tab); val _ = check_value options' name; in options' end; (* XML data *) fun encode (Options tab) = let val opts = - (tab, []) |-> Symtab.fold (fn (name, {pos, typ, value}) => - cons (Position.properties_of pos, (name, (typ, value)))); + build (tab |> Symtab.fold (fn (name, {pos, typ, value}) => + cons (Position.properties_of pos, (name, (typ, value))))); open XML.Encode; in list (pair properties (pair string (pair string string))) opts end; fun decode body = let open XML.Decode; val decode_options = list (pair properties (pair string (pair string string))) #> map (fn (props, (name, (typ, value))) => {pos = Position.of_properties props, name = name, typ = typ, value = value}); in fold declare (decode_options body) empty end; (** global default **) val global_default = Synchronized.var "Options.default" (NONE: T option); fun err_no_default () = error "Missing default for system options within Isabelle process"; fun change_default f x y = Synchronized.change global_default (fn SOME options => SOME (f x y options) | NONE => err_no_default ()); fun default () = (case Synchronized.value global_default of SOME options => options | NONE => err_no_default ()); fun default_typ name = typ (default ()) name; fun default_bool name = bool (default ()) name; fun default_int name = int (default ()) name; fun default_real name = real (default ()) name; fun default_seconds name = seconds (default ()) name; fun default_string name = string (default ()) name; val default_put_bool = change_default put_bool; val default_put_int = change_default put_int; val default_put_real = change_default put_real; val default_put_string = change_default put_string; fun get_default name = let val options = default () in get (typ options name) SOME options name end; val put_default = change_default update; fun set_default options = Synchronized.change global_default (K (SOME options)); fun reset_default () = Synchronized.change global_default (K NONE); fun load_default () = (case getenv "ISABELLE_PROCESS_OPTIONS" of "" => () | name => let val path = Path.explode name in (case try File.read path of SOME s => set_default (decode (YXML.parse_body s)) | NONE => ()) end); val _ = load_default (); val _ = ML_Print_Depth.set_print_depth (default_int "ML_print_depth"); end; diff --git a/src/Pure/Thy/export_theory.ML b/src/Pure/Thy/export_theory.ML --- a/src/Pure/Thy/export_theory.ML +++ b/src/Pure/Thy/export_theory.ML @@ -1,438 +1,438 @@ (* Title: Pure/Thy/export_theory.ML Author: Makarius Export foundational theory content and locale/class structure. *) signature EXPORT_THEORY = sig val export_enabled: Thy_Info.presentation_context -> bool val export_body: theory -> string -> XML.body -> unit end; structure Export_Theory: EXPORT_THEORY = struct (* approximative syntax *) val get_syntax = Syntax.get_approx o Proof_Context.syn_of; fun get_syntax_type ctxt = get_syntax ctxt o Lexicon.mark_type; fun get_syntax_const ctxt = get_syntax ctxt o Lexicon.mark_const; fun get_syntax_fixed ctxt = get_syntax ctxt o Lexicon.mark_fixed; fun get_syntax_param ctxt loc x = let val thy = Proof_Context.theory_of ctxt in if Class.is_class thy loc then (case AList.lookup (op =) (Class.these_params thy [loc]) x of NONE => NONE | SOME (_, (c, _)) => get_syntax_const ctxt c) else get_syntax_fixed ctxt x end; val encode_syntax = XML.Encode.variant [fn NONE => ([], []), fn SOME (Syntax.Prefix delim) => ([delim], []), fn SOME (Syntax.Infix {assoc, delim, pri}) => let val ass = (case assoc of Printer.No_Assoc => 0 | Printer.Left_Assoc => 1 | Printer.Right_Assoc => 2); open XML.Encode Term_XML.Encode; in ([], triple int string int (ass, delim, pri)) end]; (* free variables: not declared in the context *) val is_free = not oo Name.is_declared; fun add_frees used = fold_aterms (fn Free (x, T) => is_free used x ? insert (op =) (x, T) | _ => I); fun add_tfrees used = (fold_types o fold_atyps) (fn TFree (a, S) => is_free used a ? insert (op =) (a, S) | _ => I); (* locales *) fun locale_content thy loc = let val ctxt = Locale.init loc thy; val args = Locale.params_of thy loc |> map (fn ((x, T), _) => ((x, T), get_syntax_param ctxt loc x)); val axioms = let val (asm, defs) = Locale.specification_of thy loc; val cprops = map (Thm.cterm_of ctxt) (the_list asm @ defs); val (intro1, intro2) = Locale.intros_of thy loc; val intros_tac = Method.try_intros_tac ctxt (the_list intro1 @ the_list intro2) []; val res = Goal.init (Conjunction.mk_conjunction_balanced cprops) |> (ALLGOALS Goal.conjunction_tac THEN intros_tac) |> try Seq.hd; in (case res of SOME goal => Thm.prems_of goal | NONE => raise Fail ("Cannot unfold locale " ^ quote loc)) end; - val typargs = rev (fold Term.add_tfrees (map (Free o #1) args @ axioms) []); + val typargs = build_rev (fold Term.add_tfrees (map (Free o #1) args @ axioms)); in {typargs = typargs, args = args, axioms = axioms} end; fun get_locales thy = Locale.get_locales thy |> map_filter (fn loc => if Experiment.is_experiment thy loc then NONE else SOME (loc, ())); fun get_dependencies prev_thys thy = Locale.dest_dependencies prev_thys thy |> map_filter (fn dep => if Experiment.is_experiment thy (#source dep) orelse Experiment.is_experiment thy (#target dep) then NONE else let val (type_params, params) = Locale.parameters_of thy (#source dep); val typargs = fold (Term.add_tfreesT o #2 o #1) params type_params; val substT = typargs |> map_filter (fn v => let val T = TFree v; val T' = Morphism.typ (#morphism dep) T; in if T = T' then NONE else SOME (v, T') end); val subst = params |> map_filter (fn (v, _) => let val t = Free v; val t' = Morphism.term (#morphism dep) t; in if t aconv t' then NONE else SOME (v, t') end); in SOME (dep, (substT, subst)) end); (* presentation *) fun export_enabled (context: Thy_Info.presentation_context) = Options.bool (#options context) "export_theory"; fun export_body thy name body = if XML.is_empty_body body then () else Export.export thy (Path.binding0 (Path.make ["theory", name])) body; val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => let val rep_tsig = Type.rep_tsig (Sign.tsig_of thy); val consts = Sign.consts_of thy; val thy_ctxt = Proof_Context.init_global thy; val pos_properties = Thy_Info.adjust_pos_properties context; val enabled = export_enabled context; (* strict parents *) val parents = Theory.parents_of thy; val _ = Export.export thy \<^path_binding>\theory/parents\ (XML.Encode.string (cat_lines (map Context.theory_long_name parents))); (* spec rules *) fun spec_rule_content {pos, name, rough_classification, terms, rules} = let val spec = terms @ map Thm.plain_prop_of rules |> Term_Subst.zero_var_indexes_list |> map Logic.unvarify_global; in {props = pos_properties pos, name = name, rough_classification = rough_classification, - typargs = rev (fold Term.add_tfrees spec []), - args = rev (fold Term.add_frees spec []), + typargs = build_rev (fold Term.add_tfrees spec), + args = build_rev (fold Term.add_frees spec), terms = map (fn t => (t, Term.type_of t)) (take (length terms) spec), rules = drop (length terms) spec} end; (* entities *) fun make_entity_markup name xname pos serial = let val props = pos_properties pos @ Markup.serial_properties serial; in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end; fun entity_markup space name = let val xname = Name_Space.extern_shortest thy_ctxt space name; val {serial, pos, ...} = Name_Space.the_entry space name; in make_entity_markup name xname pos serial end; fun export_entities export_name get_space decls export = let val parent_spaces = map get_space parents; val space = get_space thy; in - (decls, []) |-> fold (fn (name, decl) => + build (decls |> fold (fn (name, decl) => if exists (fn space => Name_Space.declared space name) parent_spaces then I else (case export name decl of NONE => I | SOME make_body => let val i = #serial (Name_Space.the_entry space name); val body = if enabled then make_body () else []; - in cons (i, XML.Elem (entity_markup space name, body)) end)) + in cons (i, XML.Elem (entity_markup space name, body)) end))) |> sort (int_ord o apply2 #1) |> map #2 |> export_body thy export_name end; (* types *) val encode_type = let open XML.Encode Term_XML.Encode in triple encode_syntax (list string) (option typ) end; val _ = export_entities "types" Sign.type_space (Name_Space.dest_table (#types rep_tsig)) (fn c => (fn Type.LogicalType n => SOME (fn () => encode_type (get_syntax_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE)) | Type.Abbreviation (args, U, false) => SOME (fn () => encode_type (get_syntax_type thy_ctxt c, args, SOME U)) | _ => NONE)); (* consts *) val encode_term = Term_XML.Encode.term consts; val encode_const = let open XML.Encode Term_XML.Encode in pair encode_syntax (pair (list string) (pair typ (pair (option encode_term) bool))) end; val _ = export_entities "consts" Sign.const_space (#constants (Consts.dest consts)) (fn c => fn (T, abbrev) => SOME (fn () => let val syntax = get_syntax_const thy_ctxt c; val U = Logic.unvarifyT_global T; val U0 = Type.strip_sorts U; val trim_abbrev = Proofterm.standard_vars_term Name.context #> map_types Type.strip_sorts; val abbrev' = Option.map trim_abbrev abbrev; val args = map (#1 o dest_TFree) (Consts.typargs consts (c, U0)); val propositional = Object_Logic.is_propositional thy_ctxt (Term.body_type U0); in encode_const (syntax, (args, (U0, (abbrev', propositional)))) end)); (* axioms *) fun standard_prop used extra_shyps raw_prop raw_proof = let val (prop, proof) = Proofterm.standard_vars used (raw_prop, raw_proof); val args = rev (add_frees used prop []); val typargs = rev (add_tfrees used prop []); val used_typargs = fold (Name.declare o #1) typargs used; val sorts = Name.invent used_typargs Name.aT (length extra_shyps) ~~ extra_shyps; in ((sorts @ typargs, args, prop), proof) end; fun standard_prop_of thm = standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm); val encode_prop = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair string typ)) encode_term end; fun encode_axiom used prop = encode_prop (#1 (standard_prop used [] prop NONE)); val _ = export_entities "axioms" Theory.axiom_space (Theory.all_axioms_of thy) (fn _ => fn prop => SOME (fn () => encode_axiom Name.context prop)); (* theorems and proof terms *) val clean_thm = Thm.check_hyps (Context.Theory thy) #> Thm.strip_shyps; val prep_thm = clean_thm #> Thm.unconstrainT #> Thm.strip_shyps; val lookup_thm_id = Global_Theory.lookup_thm_id thy; fun expand_name thm_id (header: Proofterm.thm_header) = if #serial header = #serial thm_id then "" else (case lookup_thm_id (Proofterm.thm_header_id header) of NONE => "" | SOME thm_name => Thm_Name.print thm_name); fun entity_markup_thm (serial, (name, i)) = let val space = Facts.space_of (Global_Theory.facts_of thy); val xname = Name_Space.extern_shortest thy_ctxt space name; val {pos, ...} = Name_Space.the_entry space name; in make_entity_markup (Thm_Name.print (name, i)) (Thm_Name.print (xname, i)) pos serial end; fun encode_thm thm_id raw_thm = let val deps = map (Thm_Name.print o #2) (Thm_Deps.thm_deps thy [raw_thm]); val thm = prep_thm raw_thm; val proof0 = if Proofterm.export_standard_enabled () then Proof_Syntax.standard_proof_of {full = true, expand_name = SOME o expand_name thm_id} thm else if Proofterm.export_enabled () then Thm.reconstruct_proof_of thm else MinProof; val (prop, SOME proof) = standard_prop_of thm (SOME proof0); val _ = Thm.expose_proofs thy [thm]; in (prop, deps, proof) |> let open XML.Encode Term_XML.Encode; val encode_proof = Proofterm.encode_standard_proof consts; in triple encode_prop (list string) encode_proof end end; fun export_thm (thm_id, thm_name) = let val markup = entity_markup_thm (#serial thm_id, thm_name); val body = if enabled then Global_Theory.get_thm_name thy (thm_name, Position.none) |> encode_thm thm_id else []; in XML.Elem (markup, body) end; val _ = export_body thy "thms" (map export_thm (Global_Theory.dest_thm_names thy)); (* type classes *) val encode_class = let open XML.Encode Term_XML.Encode in pair (list (pair string typ)) (list (encode_axiom Name.context)) end; val _ = export_entities "classes" Sign.class_space (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig))))) (fn name => fn () => SOME (fn () => (case try (Axclass.get_info thy) name of NONE => ([], []) | SOME {params, axioms, ...} => (params, map (Thm.plain_prop_of o clean_thm) axioms)) |> encode_class)); (* sort algebra *) val _ = if enabled then let val prop = encode_axiom Name.context o Logic.varify_global; val encode_classrel = let open XML.Encode in list (pair prop (pair string string)) end; val encode_arities = let open XML.Encode Term_XML.Encode in list (pair prop (triple string (list sort) string)) end; val export_classrel = maps (fn (c, cs) => map (pair c) cs) #> map (`Logic.mk_classrel) #> encode_classrel; val export_arities = map (`Logic.mk_arity) #> encode_arities; val {classrel, arities} = Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents) (#2 (#classes rep_tsig)); in if null classrel then () else export_body thy "classrel" (export_classrel classrel); if null arities then () else export_body thy "arities" (export_arities arities) end else (); (* locales *) fun encode_locale used = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair (pair string typ) encode_syntax)) (list (encode_axiom used)) end; val _ = export_entities "locales" Locale.locale_space (get_locales thy) (fn loc => fn () => SOME (fn () => let val {typargs, args, axioms} = locale_content thy loc; val used = fold Name.declare (map #1 typargs @ map (#1 o #1) args) Name.context; in encode_locale used (typargs, args, axioms) end handle ERROR msg => cat_error msg ("The error(s) above occurred in locale " ^ quote (Locale.markup_name thy_ctxt loc)))); (* locale dependencies *) fun encode_locale_dependency (dep: Locale.locale_dependency, subst) = (#source dep, (#target dep, (#prefix dep, subst))) |> let open XML.Encode Term_XML.Encode; val encode_subst = pair (list (pair (pair string sort) typ)) (list (pair (pair string typ) (term consts))); in pair string (pair string (pair (list (pair string bool)) encode_subst)) end; val _ = if enabled then get_dependencies parents thy |> map_index (fn (i, dep) => let val xname = string_of_int (i + 1); val name = Long_Name.implode [Context.theory_name thy, xname]; val markup = make_entity_markup name xname (#pos (#1 dep)) (#serial (#1 dep)); val body = encode_locale_dependency dep; in XML.Elem (markup, body) end) |> export_body thy "locale_dependencies" else (); (* constdefs *) val _ = if enabled then let val constdefs = Defs.dest_constdefs (map Theory.defs_of (Theory.parents_of thy)) (Theory.defs_of thy) |> sort_by #1; val encode = let open XML.Encode in list (pair string string) end; in if null constdefs then () else export_body thy "constdefs" (encode constdefs) end else (); (* spec rules *) val encode_specs = let open XML.Encode Term_XML.Encode in list (fn {props, name, rough_classification, typargs, args, terms, rules} => pair properties (pair string (pair Spec_Rules.encode_rough_classification (pair (list (pair string sort)) (pair (list (pair string typ)) (pair (list (pair encode_term typ)) (list encode_term)))))) (props, (name, (rough_classification, (typargs, (args, (terms, rules))))))) end; val _ = if enabled then (case Spec_Rules.dest_theory thy of [] => () | spec_rules => export_body thy "spec_rules" (encode_specs (map spec_rule_content spec_rules))) else (); in () end); end; diff --git a/src/Pure/Tools/find_consts.ML b/src/Pure/Tools/find_consts.ML --- a/src/Pure/Tools/find_consts.ML +++ b/src/Pure/Tools/find_consts.ML @@ -1,172 +1,172 @@ (* Title: Pure/Tools/find_consts.ML Author: Timothy Bourke and Gerwin Klein, NICTA Hoogle-like (https://www-users.cs.york.ac.uk/~ndm/hoogle) searching by type over constants, but matching is not fuzzy. *) signature FIND_CONSTS = sig datatype criterion = Strict of string | Loose of string | Name of string val pretty_consts: Proof.context -> (bool * criterion) list -> Pretty.T val query_parser: (bool * criterion) list parser val read_query: Position.T -> string -> (bool * criterion) list val find_consts : Proof.context -> (bool * criterion) list -> unit end; structure Find_Consts : FIND_CONSTS = struct (* search criteria *) datatype criterion = Strict of string | Loose of string | Name of string; (* matching types/consts *) fun matches_subtype thy typat = Term.exists_subtype (fn ty => Sign.typ_instance thy (ty, typat)); fun check_const pred (c, (ty, _)) = if pred (c, ty) then SOME (Term.size_of_typ ty) else NONE; fun opt_not f (c as (_, (ty, _))) = if is_some (f c) then NONE else SOME (Term.size_of_typ ty); fun filter_const _ _ NONE = NONE | filter_const c f (SOME rank) = (case f c of NONE => NONE | SOME i => SOME (Int.min (rank, i))); (* pretty results *) fun pretty_criterion (b, c) = let fun prfx s = if b then s else "-" ^ s; val show_pat = quote o #1 o Input.source_content o Syntax.read_input; in (case c of Strict pat => Pretty.str (prfx "strict: " ^ show_pat pat) | Loose pat => Pretty.str (prfx (show_pat pat)) | Name name => Pretty.str (prfx "name: " ^ quote name)) end; fun pretty_const ctxt (c, ty) = let val ty' = Logic.unvarifyT_global ty; val markup = Name_Space.markup (Proof_Context.const_space ctxt) c; in Pretty.block [Pretty.mark markup (Pretty.str c), Pretty.str " ::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt ty')] end; (* find_consts *) fun pretty_consts ctxt raw_criteria = let val thy = Proof_Context.theory_of ctxt; val low_ranking = 10000; fun make_pattern crit = let val raw_T = Syntax.parse_typ ctxt crit; val t = Syntax.check_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt) (Term.dummy_pattern raw_T); in Term.type_of t end; fun make_match (Strict arg) = let val qty = make_pattern arg; in fn (_, (ty, _)) => let - val tye = Sign.typ_match thy (qty, ty) Vartab.empty; + val tye = Vartab.build (Sign.typ_match thy (qty, ty)); val sub_size = Vartab.fold (fn (_, (_, t)) => fn n => Term.size_of_typ t + n) tye 0; in SOME sub_size end handle Type.TYPE_MATCH => NONE end | make_match (Loose arg) = check_const (matches_subtype thy (make_pattern arg) o snd) | make_match (Name arg) = check_const (match_string arg o fst); fun make_criterion (b, crit) = (if b then I else opt_not) (make_match crit); val criteria = map make_criterion raw_criteria; fun user_visible (c, _) = if Consts.is_concealed (Proof_Context.consts_of ctxt) c then NONE else SOME low_ranking; fun eval_entry c = fold (filter_const c) (user_visible :: criteria) (SOME low_ranking); val {constants, ...} = Consts.dest (Sign.consts_of thy); val matches = fold (fn c => (case eval_entry c of NONE => I | SOME rank => cons (rank, c))) constants [] |> sort (prod_ord (rev_order o int_ord) (string_ord o apply2 fst)) |> map (apsnd fst o snd); val position_markup = Position.markup (Position.thread_data ()) Markup.position; in Pretty.block (Pretty.fbreaks (Pretty.mark position_markup (Pretty.keyword1 "find_consts") :: map pretty_criterion raw_criteria)) :: Pretty.str "" :: (if null matches then [Pretty.str "found nothing"] else Pretty.str ("found " ^ string_of_int (length matches) ^ " constant(s):") :: grouped 10 Par_List.map (Pretty.item o single o pretty_const ctxt) matches) end |> Pretty.chunks; fun find_consts ctxt args = Pretty.writeln (pretty_consts ctxt args); (* command syntax *) local val criterion = Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.name) >> Name || Parse.reserved "strict" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.typ) >> Strict || Parse.typ >> Loose; val query_keywords = Keyword.add_keywords [((":", \<^here>), Keyword.no_spec)] Keyword.empty_keywords; in val query_parser = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion); fun read_query pos str = Token.explode query_keywords pos str |> filter Token.is_proper |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query_parser --| Scan.ahead Parse.eof))) |> #1; end; (* PIDE query operation *) val _ = Query_Operation.register {name = "find_consts", pri = Task_Queue.urgent_pri} (fn {state, args, writeln_result, ...} => (case try Toplevel.context_of state of SOME ctxt => let val [query_arg] = args; val query = read_query Position.none query_arg; in writeln_result (Pretty.string_of (pretty_consts ctxt query)) end | NONE => error "Unknown context")); end; diff --git a/src/Pure/Tools/rule_insts.ML b/src/Pure/Tools/rule_insts.ML --- a/src/Pure/Tools/rule_insts.ML +++ b/src/Pure/Tools/rule_insts.ML @@ -1,346 +1,354 @@ (* Title: Pure/Tools/rule_insts.ML Author: Makarius Rule instantiations -- operations within implicit rule / subgoal context. *) signature RULE_INSTS = sig val where_rule: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> thm val of_rule: Proof.context -> string option list * string option list -> (binding * string option * mixfix) list -> thm -> thm val read_instantiate: Proof.context -> ((indexname * Position.T) * string) list -> string list -> thm -> thm val read_term: string -> Proof.context -> term * Proof.context val goal_context: term -> Proof.context -> (string * typ) list * Proof.context val res_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val eres_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val cut_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val forw_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val dres_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val thin_tac: Proof.context -> string -> (binding * string option * mixfix) list -> int -> tactic val subgoal_tac: Proof.context -> string -> (binding * string option * mixfix) list -> int -> tactic val make_elim_preserve: Proof.context -> thm -> thm val method: (Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic) -> (Proof.context -> thm list -> int -> tactic) -> (Proof.context -> Proof.method) context_parser end; structure Rule_Insts: RULE_INSTS = struct (** read instantiations **) local fun error_var msg (xi, pos) = error (msg ^ quote (Term.string_of_vname xi) ^ Position.here pos); -fun the_sort tvars (xi, pos) : sort = - (case AList.lookup (op =) tvars xi of +fun the_sort tvars (ai, pos) : sort = + (case Term_Subst.TVars.get_first (fn ((bi, S), ()) => if ai = bi then SOME S else NONE) tvars of SOME S => S - | NONE => error_var "No such type variable in theorem: " (xi, pos)); + | NONE => error_var "No such type variable in theorem: " (ai, pos)); fun the_type vars (xi, pos) : typ = - (case AList.lookup (op =) vars xi of + (case Vartab.lookup vars xi of SOME T => T | NONE => error_var "No such variable in theorem: " (xi, pos)); fun read_type ctxt tvars ((xi, pos), s) = let val S = the_sort tvars (xi, pos); val T = Syntax.read_typ ctxt s; in if Sign.of_sort (Proof_Context.theory_of ctxt) (T, S) then ((xi, S), T) else error_var "Bad sort for instantiation of type variable: " (xi, pos) end; -fun make_instT f v = +fun make_instT f (tvars: Term_Subst.TVars.set) = let - val T = TVar v; - val T' = f T; - in if T = T' then NONE else SOME (v, T') end; + fun add v = + let + val T = TVar v; + val T' = f T; + in if T = T' then I else cons (v, T') end; + in Term_Subst.TVars.fold (add o #1) tvars [] end; -fun make_inst f v = +fun make_inst f vars = let - val t = Var v; - val t' = f t; - in if t aconv t' then NONE else SOME (v, t') end; + fun add v = + let + val t = Var v; + val t' = f t; + in if t aconv t' then I else cons (v, t') end; + in fold add vars [] end; fun read_terms ss Ts ctxt = let fun parse T = if T = propT then Syntax.parse_prop ctxt else Syntax.parse_term ctxt; val (ts, ctxt') = fold_map Variable.fix_dummy_patterns (map2 parse Ts ss) ctxt; val ts' = map2 (Type.constraint o Type_Infer.paramify_vars) Ts ts |> Syntax.check_terms ctxt' |> Variable.polymorphic ctxt'; val Ts' = map Term.fastype_of ts'; - val tyenv = fold (Sign.typ_match (Proof_Context.theory_of ctxt)) (Ts ~~ Ts') Vartab.empty; + val tyenv = Vartab.build (fold (Sign.typ_match (Proof_Context.theory_of ctxt)) (Ts ~~ Ts')); val tyenv' = Vartab.fold (fn (xi, (S, T)) => cons ((xi, S), T)) tyenv []; in ((ts', tyenv'), ctxt') end; in fun read_term s ctxt = let val (t, ctxt') = Variable.fix_dummy_patterns (Syntax.parse_term ctxt s) ctxt; val t' = Syntax.check_term ctxt' t; in (t', ctxt') end; fun read_insts thm raw_insts raw_fixes ctxt = let val (type_insts, term_insts) = List.partition (fn (((x, _), _), _) => String.isPrefix "'" x) raw_insts; - val tvars = Thm.fold_terms Term.add_tvars thm []; - val vars = Thm.fold_terms Term.add_vars thm []; + val tvars = Term_Subst.TVars.build (Thm.fold_terms {hyps = false} Term_Subst.add_tvars thm); + val vars = Term_Subst.Vars.build (Thm.fold_terms {hyps = false} Term_Subst.add_vars thm); (*eigen-context*) val (_, ctxt1) = ctxt - |> fold (Variable.declare_internal o Logic.mk_type o TVar) tvars - |> fold (Variable.declare_internal o Var) vars + |> Term_Subst.TVars.fold (Variable.declare_internal o Logic.mk_type o TVar o #1) tvars + |> Term_Subst.Vars.fold (Variable.declare_internal o Var o #1) vars |> Proof_Context.add_fixes_cmd raw_fixes; (*explicit type instantiations*) val instT1 = Term_Subst.instantiateT (Term_Subst.TVars.table (map (read_type ctxt1 tvars) type_insts)); - val vars1 = map (apsnd instT1) vars; + val vars1 = + Vartab.build (vars |> Term_Subst.Vars.fold (fn ((v, T), ()) => + Vartab.insert (K true) (v, instT1 T))); (*term instantiations*) val (xs, ss) = split_list term_insts; val Ts = map (the_type vars1) xs; val ((ts, inferred), ctxt2) = read_terms ss Ts ctxt1; (*implicit type instantiations*) val instT2 = Term_Subst.instantiateT (Term_Subst.TVars.table inferred); - val vars2 = map (apsnd instT2) vars1; + val vars2 = Vartab.fold (fn (v, T) => cons (v, instT2 T)) vars1 []; val inst2 = Term_Subst.instantiate (Term_Subst.TVars.empty, - fold2 (fn (xi, _) => fn t => Term_Subst.Vars.add ((xi, Term.fastype_of t), t)) - xs ts Term_Subst.Vars.empty) + Term_Subst.Vars.build + (fold2 (fn (xi, _) => fn t => Term_Subst.Vars.add ((xi, Term.fastype_of t), t)) xs ts)) #> Envir.beta_norm; - val inst_tvars = map_filter (make_instT (instT2 o instT1)) tvars; - val inst_vars = map_filter (make_inst inst2) vars2; + val inst_tvars = make_instT (instT2 o instT1) tvars; + val inst_vars = make_inst inst2 vars2; in ((inst_tvars, inst_vars), ctxt2) end; end; (** forward rules **) fun where_rule ctxt raw_insts raw_fixes thm = let val ((inst_tvars, inst_vars), ctxt') = read_insts thm raw_insts raw_fixes ctxt; in thm |> Drule.instantiate_normalize (map (apsnd (Thm.ctyp_of ctxt')) inst_tvars, map (apsnd (Thm.cterm_of ctxt')) inst_vars) |> singleton (Variable.export ctxt' ctxt) |> Rule_Cases.save thm end; fun of_rule ctxt (args, concl_args) fixes thm = let fun zip_vars _ [] = [] | zip_vars (_ :: xs) (NONE :: rest) = zip_vars xs rest | zip_vars ((x, _) :: xs) (SOME t :: rest) = ((x, Position.none), t) :: zip_vars xs rest | zip_vars [] _ = error "More instantiations than variables in theorem"; val insts = - zip_vars (rev (Term.add_vars (Thm.full_prop_of thm) [])) args @ - zip_vars (rev (Term.add_vars (Thm.concl_of thm) [])) concl_args; + zip_vars (build_rev (Term.add_vars (Thm.full_prop_of thm))) args @ + zip_vars (build_rev (Term.add_vars (Thm.concl_of thm))) concl_args; in where_rule ctxt insts fixes thm end; fun read_instantiate ctxt insts xs = where_rule ctxt insts (map (fn x => (Binding.name x, NONE, NoSyn)) xs); (** attributes **) (* where: named instantiation *) val named_insts = Parse.and_list1 (Parse.position Args.var -- (Args.$$$ "=" |-- Parse.!!! Args.embedded_inner_syntax)) -- Parse.for_fixes; val _ = Theory.setup (Attrib.setup \<^binding>\where\ (Scan.lift named_insts >> (fn args => Thm.rule_attribute [] (fn context => uncurry (where_rule (Context.proof_of context)) args))) "named instantiation of theorem"); (* of: positional instantiation (terms only) *) local val inst = Args.maybe Args.embedded_inner_syntax; val concl = Args.$$$ "concl" -- Args.colon; val insts = Scan.repeat (Scan.unless concl inst) -- Scan.optional (concl |-- Scan.repeat inst) []; in val _ = Theory.setup (Attrib.setup \<^binding>\of\ (Scan.lift (insts -- Parse.for_fixes) >> (fn args => Thm.rule_attribute [] (fn context => uncurry (of_rule (Context.proof_of context)) args))) "positional instantiation of theorem"); end; (** tactics **) (* goal context *) fun goal_context goal ctxt = let val ((_, params), ctxt') = ctxt |> Variable.declare_constraints goal |> Variable.improper_fixes |> Variable.focus_params NONE goal ||> Variable.restore_proper_fixes ctxt; in (params, ctxt') end; (* resolution after lifting and instantiation; may refer to parameters of the subgoal *) fun bires_inst_tac bires_flag ctxt raw_insts raw_fixes thm i st = CSUBGOAL (fn (cgoal, _) => let (*goal context*) val (params, goal_ctxt) = goal_context (Thm.term_of cgoal) ctxt; val paramTs = map #2 params; (*instantiation context*) val ((inst_tvars, inst_vars), inst_ctxt) = read_insts thm raw_insts raw_fixes goal_ctxt; val fixed = map #1 (fold (Variable.add_newly_fixed inst_ctxt goal_ctxt o #2) inst_vars []); (* lift and instantiate rule *) val inc = Thm.maxidx_of st + 1; val lift_type = Logic.incr_tvar inc; fun lift_var ((a, j), T) = ((a, j + inc), paramTs ---> lift_type T); fun lift_term t = fold_rev Term.absfree params (Logic.incr_indexes (fixed, paramTs, inc) t); val inst_tvars' = inst_tvars |> map (fn (((a, i), S), T) => (((a, i + inc), S), Thm.ctyp_of inst_ctxt (lift_type T))); val inst_vars' = inst_vars |> map (fn (v, t) => (lift_var v, Thm.cterm_of inst_ctxt (lift_term t))); val thm' = Thm.lift_rule cgoal thm |> Drule.instantiate_normalize (inst_tvars', inst_vars') |> singleton (Variable.export inst_ctxt ctxt); in compose_tac ctxt (bires_flag, thm', Thm.nprems_of thm) i end) i st; val res_inst_tac = bires_inst_tac false; val eres_inst_tac = bires_inst_tac true; (* forward resolution *) fun make_elim_preserve ctxt rl = let val maxidx = Thm.maxidx_of rl; fun var x = ((x, 0), propT); fun cvar xi = Thm.cterm_of ctxt (Var (xi, propT)); val revcut_rl' = Drule.instantiate_normalize ([], [(var "V", cvar ("V", maxidx + 1)), (var "W", cvar ("W", maxidx + 1))]) Drule.revcut_rl; in (case Seq.list_of (Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false} (false, rl, Thm.nprems_of rl) 1 revcut_rl') of [th] => th | _ => raise THM ("make_elim_preserve", 1, [rl])) end; (*instantiate and cut -- for atomic fact*) fun cut_inst_tac ctxt insts fixes rule = res_inst_tac ctxt insts fixes (make_elim_preserve ctxt rule); (*forward tactic applies a rule to an assumption without deleting it*) fun forw_inst_tac ctxt insts fixes rule = cut_inst_tac ctxt insts fixes rule THEN' assume_tac ctxt; (*dresolve tactic applies a rule to replace an assumption*) fun dres_inst_tac ctxt insts fixes rule = eres_inst_tac ctxt insts fixes (make_elim_preserve ctxt rule); (* derived tactics *) (*deletion of an assumption*) fun thin_tac ctxt s fixes = eres_inst_tac ctxt [((("V", 0), Position.none), s)] fixes Drule.thin_rl; (*Introduce the given proposition as lemma and subgoal*) fun subgoal_tac ctxt A fixes = DETERM o res_inst_tac ctxt [((("psi", 0), Position.none), A)] fixes cut_rl; (* method wrapper *) fun method inst_tac tac = Args.goal_spec -- Scan.optional (Scan.lift (named_insts --| Args.$$$ "in")) ([], []) -- Attrib.thms >> (fn ((quant, (insts, fixes)), thms) => fn ctxt => METHOD (fn facts => if null insts andalso null fixes then quant (Method.insert_tac ctxt facts THEN' tac ctxt thms) else (case thms of [thm] => quant (Method.insert_tac ctxt facts THEN' inst_tac ctxt insts fixes thm) | _ => error "Cannot have instantiations with multiple rules"))); (* setup *) (*warning: rule_tac etc. refer to dynamic subgoal context!*) val _ = Theory.setup (Method.setup \<^binding>\rule_tac\ (method res_inst_tac resolve_tac) "apply rule (dynamic instantiation)" #> Method.setup \<^binding>\erule_tac\ (method eres_inst_tac eresolve_tac) "apply rule in elimination manner (dynamic instantiation)" #> Method.setup \<^binding>\drule_tac\ (method dres_inst_tac dresolve_tac) "apply rule in destruct manner (dynamic instantiation)" #> Method.setup \<^binding>\frule_tac\ (method forw_inst_tac forward_tac) "apply rule in forward manner (dynamic instantiation)" #> Method.setup \<^binding>\cut_tac\ (method cut_inst_tac (K cut_rules_tac)) "cut rule (dynamic instantiation)" #> Method.setup \<^binding>\subgoal_tac\ (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.embedded_inner_syntax -- Parse.for_fixes) >> (fn (quant, (props, fixes)) => fn ctxt => SIMPLE_METHOD'' quant (EVERY' (map (fn prop => subgoal_tac ctxt prop fixes) props)))) "insert subgoal (dynamic instantiation)" #> Method.setup \<^binding>\thin_tac\ (Args.goal_spec -- Scan.lift (Args.embedded_inner_syntax -- Parse.for_fixes) >> (fn (quant, (prop, fixes)) => fn ctxt => SIMPLE_METHOD'' quant (thin_tac ctxt prop fixes))) "remove premise (dynamic instantiation)"); end; diff --git a/src/Pure/consts.ML b/src/Pure/consts.ML --- a/src/Pure/consts.ML +++ b/src/Pure/consts.ML @@ -1,348 +1,348 @@ (* Title: Pure/consts.ML Author: Makarius Polymorphic constants: declarations, abbreviations, additional type constraints. *) signature CONSTS = sig type T val eq_consts: T * T -> bool val change_base: bool -> T -> T val change_ignore: T -> T val retrieve_abbrevs: T -> string list -> term -> (term * term) list val dest: T -> {const_space: Name_Space.T, constants: (string * (typ * term option)) list, constraints: (string * typ) list} val the_const: T -> string -> string * typ (*exception TYPE*) val the_abbreviation: T -> string -> typ * term (*exception TYPE*) val type_scheme: T -> string -> typ (*exception TYPE*) val is_monomorphic: T -> string -> bool (*exception TYPE*) val the_constraint: T -> string -> typ (*exception TYPE*) val space_of: T -> Name_Space.T val alias: Name_Space.naming -> binding -> string -> T -> T val is_concealed: T -> string -> bool val intern: T -> xstring -> string val intern_syntax: T -> xstring -> string val check_const: Context.generic -> T -> xstring * Position.T list -> term * Position.report list val certify: Context.generic -> Type.tsig -> bool -> T -> term -> term (*exception TYPE*) val typargs: T -> string * typ -> typ list val instance: T -> string * typ list -> typ val dummy_types: T -> term -> term val declare: Context.generic -> binding * typ -> T -> T val constrain: string * typ option -> T -> T val abbreviate: Context.generic -> Type.tsig -> string -> binding * term -> T -> (term * term) * T val revert_abbrev: string -> string -> T -> T val hide: bool -> string -> T -> T val empty: T val merge: T * T -> T end; structure Consts: CONSTS = struct (** consts type **) (* datatype T *) type decl = {T: typ, typargs: int list list}; type abbrev = {rhs: term, normal_rhs: term, force_expand: bool}; datatype T = Consts of {decls: (decl * abbrev option) Name_Space.table, constraints: typ Symtab.table, rev_abbrevs: (term * term) Item_Net.T Symtab.table}; fun eq_consts (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1}, Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) = pointer_eq (decls1, decls2) andalso pointer_eq (constraints1, constraints2) andalso pointer_eq (rev_abbrevs1, rev_abbrevs2); fun make_consts (decls, constraints, rev_abbrevs) = Consts {decls = decls, constraints = constraints, rev_abbrevs = rev_abbrevs}; fun map_consts f (Consts {decls, constraints, rev_abbrevs}) = make_consts (f (decls, constraints, rev_abbrevs)); fun change_base begin = map_consts (fn (decls, constraints, rev_abbrevs) => (Name_Space.change_base begin decls, constraints, rev_abbrevs)); val change_ignore = map_consts (fn (decls, constraints, rev_abbrevs) => (Name_Space.change_ignore decls, constraints, rev_abbrevs)); (* reverted abbrevs *) val empty_abbrevs = Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') (single o #1); fun update_abbrevs mode abbrs = Symtab.map_default (mode, empty_abbrevs) (Item_Net.update abbrs); fun retrieve_abbrevs (Consts {rev_abbrevs, ...}) modes = let val nets = map_filter (Symtab.lookup rev_abbrevs) modes in fn t => let val retrieve = if Term.could_beta_eta_contract t then Item_Net.retrieve else Item_Net.retrieve_matching in maps (fn net => retrieve net t) nets end end; (* dest consts *) fun dest (Consts {decls, constraints, ...}) = {const_space = Name_Space.space_of_table decls, constants = Name_Space.fold_table (fn (c, ({T, ...}, abbr)) => cons (c, (T, Option.map #rhs abbr))) decls [], constraints = Symtab.dest constraints}; (* lookup consts *) fun the_entry (Consts {decls, ...}) c = (case Name_Space.lookup_key decls c of SOME entry => entry | NONE => raise TYPE ("Unknown constant: " ^ quote c, [], [])); fun the_const consts c = (case the_entry consts c of (c', ({T, ...}, NONE)) => (c', T) | _ => raise TYPE ("Not a logical constant: " ^ quote c, [], [])); fun the_abbreviation consts c = (case the_entry consts c of (_, ({T, ...}, SOME {rhs, ...})) => (T, rhs) | _ => raise TYPE ("Not an abbreviated constant: " ^ quote c, [], [])); fun the_decl consts = #1 o #2 o the_entry consts; val type_scheme = #T oo the_decl; val type_arguments = #typargs oo the_decl; val is_monomorphic = null oo type_arguments; fun the_constraint (consts as Consts {constraints, ...}) c = (case Symtab.lookup constraints c of SOME T => T | NONE => type_scheme consts c); (* name space and syntax *) fun space_of (Consts {decls, ...}) = Name_Space.space_of_table decls; fun alias naming binding name = map_consts (fn (decls, constraints, rev_abbrevs) => ((Name_Space.alias_table naming binding name decls), constraints, rev_abbrevs)); val is_concealed = Name_Space.is_concealed o space_of; val intern = Name_Space.intern o space_of; fun intern_syntax consts s = (case try Lexicon.unmark_const s of SOME c => c | NONE => intern consts s); (* check_const *) fun check_const context consts (xname, ps) = let val Consts {decls, ...} = consts; val ((c, reports), _) = Name_Space.check_reports context decls (xname, ps); val T = type_scheme consts c handle TYPE (msg, _, _) => error (msg ^ Position.here_list ps); in (Const (c, T), reports) end; (* certify *) fun certify context tsig do_expand consts = let fun err msg (c, T) = raise TYPE (msg ^ " " ^ quote c ^ " :: " ^ Syntax.string_of_typ (Syntax.init_pretty context) T, [], []); val certT = Type.cert_typ tsig; fun cert tm = let val (head, args) = Term.strip_comb tm; val args' = map cert args; fun comb head' = Term.list_comb (head', args'); in (case head of Abs (x, T, t) => comb (Abs (x, certT T, cert t)) | Const (c, T) => let val T' = certT T; val (_, ({T = U, ...}, abbr)) = the_entry consts c; fun expand u = Term.betapplys (Envir.expand_atom T' (U, u) handle TYPE _ => err "Illegal type for abbreviation" (c, T), args'); in if not (Type.raw_instance (T', U)) then err "Illegal type for constant" (c, T) else (case abbr of SOME {rhs, normal_rhs, force_expand} => if do_expand then expand normal_rhs else if force_expand then expand rhs else comb head | _ => comb head) end | _ => comb head) end; in cert end; (* typargs -- view actual const type as instance of declaration *) local fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos | args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos) | args_of (TFree _) _ = I and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is | args_of_list [] _ _ = I; fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is | subscript T [] = T | subscript _ _ = raise Subscript; in fun typargs_of T = map #2 (rev (args_of T [] [])); fun typargs consts (c, T) = map (subscript T) (type_arguments consts c); end; fun instance consts (c, Ts) = let val declT = type_scheme consts c; val args = typargs consts (c, declT); val inst = - fold2 (fn a => fn T => Term_Subst.TVars.add (Term.dest_TVar a, T)) - args Ts Term_Subst.TVars.empty + Term_Subst.TVars.build + (fold2 (fn a => fn T => Term_Subst.TVars.add (Term.dest_TVar a, T)) args Ts) handle ListPair.UnequalLengths => raise TYPE ("Consts.instance", Ts, [Const (c, dummyT)]); in declT |> Term_Subst.instantiateT inst end; fun dummy_types consts = let fun dummy (Const (c, T)) = Const (c, instance consts (c, replicate (length (typargs consts (c, T))) dummyT)) | dummy (Free (x, _)) = Free (x, dummyT) | dummy (Var (xi, _)) = Var (xi, dummyT) | dummy (b as Bound _) = b | dummy (t $ u) = dummy t $ dummy u | dummy (Abs (a, _, b)) = Abs (a, dummyT, dummy b); in dummy end; (** build consts **) (* name space *) fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs) => (Name_Space.hide_table fully c decls, constraints, rev_abbrevs)); (* declarations *) fun declare context (b, declT) = map_consts (fn (decls, constraints, rev_abbrevs) => let val decl = {T = declT, typargs = typargs_of declT}; val _ = Binding.check b; val (_, decls') = decls |> Name_Space.define context true (b, (decl, NONE)); in (decls', constraints, rev_abbrevs) end); (* constraints *) fun constrain (c, C) consts = consts |> map_consts (fn (decls, constraints, rev_abbrevs) => (#2 (the_entry consts c) handle TYPE (msg, _, _) => error msg; (decls, constraints |> (case C of SOME T => Symtab.update (c, T) | NONE => Symtab.delete_safe c), rev_abbrevs))); (* abbreviations *) local fun strip_abss (t as Abs (x, T, b)) = if Term.is_dependent b then strip_abss b |>> cons (x, T) (* FIXME decr!? *) else ([], t) | strip_abss t = ([], t); fun rev_abbrev lhs rhs = let val (xs, body) = strip_abss (Envir.beta_eta_contract rhs); val vars = fold (fn (x, T) => cons (Var ((x, 0), T))) (Term.rename_wrt_term body xs) []; in (Term.subst_bounds (rev vars, body), Term.list_comb (lhs, vars)) end; in fun abbreviate context tsig mode (b, raw_rhs) consts = let val cert_term = certify context tsig false consts; val expand_term = certify context tsig true consts; val force_expand = mode = Print_Mode.internal; val _ = Term.exists_subterm Term.is_Var raw_rhs andalso error ("Illegal schematic variables on rhs of abbreviation " ^ Binding.print b); val rhs = raw_rhs |> Term.map_types (Type.cert_typ tsig) |> cert_term |> Term.close_schematic_term; val normal_rhs = expand_term rhs; val T = Term.fastype_of rhs; val lhs = Const (Name_Space.full_name (Name_Space.naming_of context) b, T); in consts |> map_consts (fn (decls, constraints, rev_abbrevs) => let val decl = {T = T, typargs = typargs_of T}; val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand}; val _ = Binding.check b; val (_, decls') = decls |> Name_Space.define context true (b, (decl, SOME abbr)); val rev_abbrevs' = rev_abbrevs |> update_abbrevs mode (rev_abbrev lhs rhs); in (decls', constraints, rev_abbrevs') end) |> pair (lhs, rhs) end; fun revert_abbrev mode c consts = consts |> map_consts (fn (decls, constraints, rev_abbrevs) => let val (T, rhs) = the_abbreviation consts c; val rev_abbrevs' = rev_abbrevs |> update_abbrevs mode (rev_abbrev (Const (c, T)) rhs); in (decls, constraints, rev_abbrevs') end); end; (* empty and merge *) val empty = make_consts (Name_Space.empty_table Markup.constantN, Symtab.empty, Symtab.empty); fun merge (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1}, Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) = let val decls' = Name_Space.merge_tables (decls1, decls2); val constraints' = Symtab.merge (K true) (constraints1, constraints2); val rev_abbrevs' = Symtab.join (K Item_Net.merge) (rev_abbrevs1, rev_abbrevs2); in make_consts (decls', constraints', rev_abbrevs') end; end; diff --git a/src/Pure/context.ML b/src/Pure/context.ML --- a/src/Pure/context.ML +++ b/src/Pure/context.ML @@ -1,768 +1,768 @@ (* Title: Pure/context.ML Author: Markus Wenzel, TU Muenchen Generic theory contexts with unique identity, arbitrarily typed data, monotonic development graph and history support. Generic proof contexts with arbitrarily typed data. Firm naming conventions: thy, thy', thy1, thy2: theory ctxt, ctxt', ctxt1, ctxt2: Proof.context context: Context.generic *) signature BASIC_CONTEXT = sig type theory exception THEORY of string * theory list structure Proof: sig type context end structure Proof_Context: sig val theory_of: Proof.context -> theory val init_global: theory -> Proof.context val get_global: theory -> string -> Proof.context end end; signature CONTEXT = sig include BASIC_CONTEXT (*theory context*) type theory_id val theory_id: theory -> theory_id val timing: bool Unsynchronized.ref val parents_of: theory -> theory list val ancestors_of: theory -> theory list val theory_id_ord: theory_id ord val theory_id_long_name: theory_id -> string val theory_id_name: theory_id -> string val theory_long_name: theory -> string val theory_name: theory -> string val theory_name': {long: bool} -> theory -> string val PureN: string val pretty_thy: theory -> Pretty.T val pretty_abbrev_thy: theory -> Pretty.T val get_theory: {long: bool} -> theory -> string -> theory val eq_thy_id: theory_id * theory_id -> bool val eq_thy: theory * theory -> bool val proper_subthy_id: theory_id * theory_id -> bool val proper_subthy: theory * theory -> bool val subthy_id: theory_id * theory_id -> bool val subthy: theory * theory -> bool val trace_theories: bool Unsynchronized.ref val theories_trace: unit -> {active_positions: Position.T list, active: int, total: int} val join_thys: theory * theory -> theory val begin_thy: string -> theory list -> theory val finish_thy: theory -> theory val theory_data_size: theory -> (Position.T * int) list (*proof context*) val raw_transfer: theory -> Proof.context -> Proof.context (*certificate*) datatype certificate = Certificate of theory | Certificate_Id of theory_id val certificate_theory: certificate -> theory val certificate_theory_id: certificate -> theory_id val eq_certificate: certificate * certificate -> bool val join_certificate: certificate * certificate -> certificate (*generic context*) datatype generic = Theory of theory | Proof of Proof.context val cases: (theory -> 'a) -> (Proof.context -> 'a) -> generic -> 'a val mapping: (theory -> theory) -> (Proof.context -> Proof.context) -> generic -> generic val mapping_result: (theory -> 'a * theory) -> (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic val the_theory: generic -> theory val the_proof: generic -> Proof.context val map_theory: (theory -> theory) -> generic -> generic val map_proof: (Proof.context -> Proof.context) -> generic -> generic val map_theory_result: (theory -> 'a * theory) -> generic -> 'a * generic val map_proof_result: (Proof.context -> 'a * Proof.context) -> generic -> 'a * generic val theory_map: (generic -> generic) -> theory -> theory val proof_map: (generic -> generic) -> Proof.context -> Proof.context val theory_of: generic -> theory (*total*) val proof_of: generic -> Proof.context (*total*) (*thread data*) val get_generic_context: unit -> generic option val put_generic_context: generic option -> unit val setmp_generic_context: generic option -> ('a -> 'b) -> 'a -> 'b val the_generic_context: unit -> generic val the_global_context: unit -> theory val the_local_context: unit -> Proof.context val >> : (generic -> generic) -> unit val >>> : (generic -> 'a * generic) -> 'a end; signature PRIVATE_CONTEXT = sig include CONTEXT structure Theory_Data: sig val declare: Position.T -> Any.T -> (Any.T -> Any.T) -> (theory * theory -> Any.T * Any.T -> Any.T) -> serial val get: serial -> (Any.T -> 'a) -> theory -> 'a val put: serial -> ('a -> Any.T) -> 'a -> theory -> theory end structure Proof_Data: sig val declare: (theory -> Any.T) -> serial val get: serial -> (Any.T -> 'a) -> Proof.context -> 'a val put: serial -> ('a -> Any.T) -> 'a -> Proof.context -> Proof.context end end; structure Context: PRIVATE_CONTEXT = struct (*** theory context ***) (*private copy avoids potential conflict of table exceptions*) structure Datatab = Table(type key = int val ord = int_ord); (** datatype theory **) type stage = int * int Synchronized.var; fun init_stage () : stage = (0, Synchronized.var "Context.stage" 1); fun next_stage ((_, state): stage) = (Synchronized.change_result state (fn n => (n, n + 1)), state); abstype theory_id = Theory_Id of (*identity*) {id: serial, (*identifier*) ids: Inttab.set} * (*cumulative identifiers -- symbolic body content*) (*history*) {name: string, (*official theory name*) stage: stage option} (*index and counter for anonymous updates*) with fun rep_theory_id (Theory_Id args) = args; val make_theory_id = Theory_Id; end; datatype theory = Theory of (*identity*) {theory_id: theory_id, token: Position.T Unsynchronized.ref} * (*ancestry*) {parents: theory list, (*immediate predecessors*) ancestors: theory list} * (*all predecessors -- canonical reverse order*) (*data*) Any.T Datatab.table; (*body content*) exception THEORY of string * theory list; fun rep_theory (Theory args) = args; val theory_identity = #1 o rep_theory; val theory_id = #theory_id o theory_identity; val identity_of_id = #1 o rep_theory_id; val identity_of = identity_of_id o theory_id; val history_of_id = #2 o rep_theory_id; val history_of = history_of_id o theory_id; val ancestry_of = #2 o rep_theory; val data_of = #3 o rep_theory; fun make_identity id ids = {id = id, ids = ids}; fun make_history name = {name = name, stage = SOME (init_stage ())}; fun make_ancestry parents ancestors = {parents = parents, ancestors = ancestors}; val theory_id_ord = int_ord o apply2 (#id o identity_of_id); val theory_id_long_name = #name o history_of_id; val theory_id_name = Long_Name.base_name o theory_id_long_name; val theory_long_name = #name o history_of; val theory_name = Long_Name.base_name o theory_long_name; fun theory_name' {long} = if long then theory_long_name else theory_name; val parents_of = #parents o ancestry_of; val ancestors_of = #ancestors o ancestry_of; (* names *) val PureN = "Pure"; fun display_name thy_id = (case history_of_id thy_id of {name, stage = NONE} => name | {name, stage = SOME (i, _)} => name ^ ":" ^ string_of_int i); fun display_names thy = let val name = display_name (theory_id thy); val ancestor_names = map theory_long_name (ancestors_of thy); in rev (name :: ancestor_names) end; val pretty_thy = Pretty.str_list "{" "}" o display_names; val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty_thy); fun pretty_abbrev_thy thy = let val names = display_names thy; val n = length names; val abbrev = if n > 5 then "..." :: List.drop (names, n - 5) else names; in Pretty.str_list "{" "}" abbrev end; fun get_theory long thy name = if theory_name' long thy <> name then (case find_first (fn thy' => theory_name' long thy' = name) (ancestors_of thy) of SOME thy' => thy' | NONE => error ("Unknown ancestor theory " ^ quote name)) else if is_none (#stage (history_of thy)) then thy else error ("Unfinished theory " ^ quote name); (* build ids *) fun insert_id id ids = Inttab.update (id, ()) ids; val merge_ids = apply2 (theory_id #> rep_theory_id #> #1) #> (fn ({id = id1, ids = ids1, ...}, {id = id2, ids = ids2, ...}) => Inttab.merge (K true) (ids1, ids2) |> insert_id id1 |> insert_id id2); (* equality and inclusion *) val eq_thy_id = op = o apply2 (#id o identity_of_id); val eq_thy = op = o apply2 (#id o identity_of); val proper_subthy_id = apply2 (rep_theory_id #> #1) #> (fn ({id, ...}, {ids, ...}) => Inttab.defined ids id); val proper_subthy = proper_subthy_id o apply2 theory_id; fun subthy_id p = eq_thy_id p orelse proper_subthy_id p; val subthy = subthy_id o apply2 theory_id; (* consistent ancestors *) fun eq_thy_consistent (thy1, thy2) = eq_thy (thy1, thy2) orelse (theory_name thy1 = theory_name thy2 andalso raise THEORY ("Duplicate theory name", [thy1, thy2])); fun extend_ancestors thy thys = if member eq_thy_consistent thys thy then raise THEORY ("Duplicate theory node", thy :: thys) else thy :: thys; val merge_ancestors = merge eq_thy_consistent; (** theory data **) (* data kinds and access methods *) val timing = Unsynchronized.ref false; local type kind = {pos: Position.T, empty: Any.T, extend: Any.T -> Any.T, merge: theory * theory -> Any.T * Any.T -> Any.T}; val kinds = Synchronized.var "Theory_Data" (Datatab.empty: kind Datatab.table); fun invoke name f k x = (case Datatab.lookup (Synchronized.value kinds) k of SOME kind => if ! timing andalso name <> "" then Timing.cond_timeit true ("Theory_Data." ^ name ^ Position.here (#pos kind)) (fn () => f kind x) else f kind x | NONE => raise Fail "Invalid theory data identifier"); in fun invoke_pos k = invoke "" (K o #pos) k (); fun invoke_empty k = invoke "" (K o #empty) k (); val invoke_extend = invoke "extend" #extend; fun invoke_merge thys = invoke "merge" (fn kind => #merge kind thys); fun declare_theory_data pos empty extend merge = let val k = serial (); val kind = {pos = pos, empty = empty, extend = extend, merge = merge}; val _ = Synchronized.change kinds (Datatab.update (k, kind)); in k end; val extend_data = Datatab.map invoke_extend; fun merge_data thys = Datatab.join (invoke_merge thys) o apply2 extend_data; end; (** build theories **) (* create theory *) val trace_theories = Unsynchronized.ref false; local val theories = Synchronized.var "theory_tokens" ([]: Position.T Unsynchronized.ref option Unsynchronized.ref list); val dummy_token = Unsynchronized.ref Position.none; fun make_token () = if ! trace_theories then let val token = Unsynchronized.ref (Position.thread_data ()); val _ = Synchronized.change theories (cons (Weak.weak (SOME token))); in token end else dummy_token; in fun theories_trace () = let val trace = Synchronized.value theories; val _ = ML_Heap.full_gc (); val active_positions = fold (fn Unsynchronized.ref (SOME pos) => cons (! pos) | _ => I) trace []; in {active_positions = active_positions, active = length active_positions, total = length trace} end; fun create_thy ids history ancestry data = let val theory_id = make_theory_id (make_identity (serial ()) ids, history); val token = make_token (); in Theory ({theory_id = theory_id, token = token}, ancestry, data) end; end; (* primitives *) val pre_pure_thy = create_thy Inttab.empty (make_history PureN) (make_ancestry [] []) Datatab.empty; local fun change_thy finish f thy = let val ({id, ids}, {name, stage}) = rep_theory_id (theory_id thy); val Theory (_, ancestry, data) = thy; val (ancestry', data') = if is_none stage then (make_ancestry [thy] (extend_ancestors thy (ancestors_of thy)), extend_data data) else (ancestry, data); val history' = {name = name, stage = if finish then NONE else Option.map next_stage stage}; val ids' = insert_id id ids; val data'' = f data'; in create_thy ids' history' ancestry' data'' end; in val update_thy = change_thy false; val extend_thy = update_thy I; val finish_thy = change_thy true I #> tap extend_thy; end; (* join: anonymous theory nodes *) local fun bad_join (thy1, thy2) = raise THEORY ("Cannot join theories", [thy1, thy2]); fun join_history thys = apply2 history_of thys |> (fn ({name, stage}, {name = name', stage = stage'}) => if name = name' andalso is_some stage andalso is_some stage' then {name = name, stage = Option.map next_stage stage} else bad_join thys); fun join_ancestry thys = apply2 ancestry_of thys |> (fn (ancestry as {parents, ancestors}, {parents = parents', ancestors = ancestors'}) => if eq_list eq_thy (parents, parents') andalso eq_list eq_thy (ancestors, ancestors') then ancestry else bad_join thys); in fun join_thys thys = let val ids = merge_ids thys; val history = join_history thys; val ancestry = join_ancestry thys; val data = merge_data thys (apply2 data_of thys); in create_thy ids history ancestry data end; end; (* merge: named theory nodes *) local fun merge_thys thys = let val ids = merge_ids thys; val history = make_history ""; val ancestry = make_ancestry [] []; val data = merge_data thys (apply2 data_of thys); in create_thy ids history ancestry data end; fun maximal_thys thys = thys |> filter_out (fn thy => exists (fn thy' => proper_subthy (thy, thy')) thys); in fun begin_thy name imports = if name = "" then error ("Bad theory name: " ^ quote name) else let val parents = maximal_thys (distinct eq_thy imports); val ancestors = Library.foldl merge_ancestors ([], map ancestors_of parents) |> fold extend_ancestors parents; val thy0 = (case parents of [] => error "Missing theory imports" | [thy] => extend_thy thy | thy :: thys => Library.foldl merge_thys (thy, thys)); val ({ids, ...}, _) = rep_theory_id (theory_id thy0); val history = make_history name; val ancestry = make_ancestry parents ancestors; in create_thy ids history ancestry (data_of thy0) |> tap finish_thy end; end; (* theory data *) structure Theory_Data = struct val declare = declare_theory_data; fun get k dest thy = (case Datatab.lookup (data_of thy) k of SOME x => x | NONE => invoke_empty k) |> dest; fun put k mk x = update_thy (Datatab.update (k, mk x)); fun obj_size k thy = Datatab.lookup (data_of thy) k |> Option.map ML_Heap.obj_size; end; fun theory_data_size thy = - (data_of thy, []) |-> Datatab.fold_rev (fn (k, _) => + build (data_of thy |> Datatab.fold_rev (fn (k, _) => (case Theory_Data.obj_size k thy of NONE => I - | SOME n => (cons (invoke_pos k, n)))); + | SOME n => (cons (invoke_pos k, n))))); (*** proof context ***) (* datatype Proof.context *) structure Proof = struct datatype context = Context of Any.T Datatab.table * theory; end; (* proof data kinds *) local val kinds = Synchronized.var "Proof_Data" (Datatab.empty: (theory -> Any.T) Datatab.table); fun init_data thy = Synchronized.value kinds |> Datatab.map (fn _ => fn init => init thy); fun init_new_data thy = Synchronized.value kinds |> Datatab.fold (fn (k, init) => fn data => if Datatab.defined data k then data else Datatab.update (k, init thy) data); fun init_fallback k thy = (case Datatab.lookup (Synchronized.value kinds) k of SOME init => init thy | NONE => raise Fail "Invalid proof data identifier"); in fun raw_transfer thy' (Proof.Context (data, thy)) = let val _ = subthy (thy, thy') orelse error "Cannot transfer proof context: not a super theory"; val data' = init_new_data thy' data; in Proof.Context (data', thy') end; structure Proof_Context = struct fun theory_of (Proof.Context (_, thy)) = thy; fun init_global thy = Proof.Context (init_data thy, thy); fun get_global thy name = init_global (get_theory {long = false} thy name); end; structure Proof_Data = struct fun declare init = let val k = serial (); val _ = Synchronized.change kinds (Datatab.update (k, init)); in k end; fun get k dest (Proof.Context (data, thy)) = (case Datatab.lookup data k of SOME x => x | NONE => init_fallback k thy) |> dest; fun put k mk x (Proof.Context (data, thy)) = Proof.Context (Datatab.update (k, mk x) data, thy); end; end; (*** theory certificate ***) datatype certificate = Certificate of theory | Certificate_Id of theory_id; fun certificate_theory (Certificate thy) = thy | certificate_theory (Certificate_Id thy_id) = error ("No content for theory certificate " ^ display_name thy_id); fun certificate_theory_id (Certificate thy) = theory_id thy | certificate_theory_id (Certificate_Id thy_id) = thy_id; fun eq_certificate (Certificate thy1, Certificate thy2) = eq_thy (thy1, thy2) | eq_certificate (Certificate_Id thy_id1, Certificate_Id thy_id2) = eq_thy_id (thy_id1, thy_id2) | eq_certificate _ = false; fun join_certificate (cert1, cert2) = let val (thy_id1, thy_id2) = apply2 certificate_theory_id (cert1, cert2) in if eq_thy_id (thy_id1, thy_id2) then (case cert1 of Certificate _ => cert1 | _ => cert2) else if proper_subthy_id (thy_id2, thy_id1) then cert1 else if proper_subthy_id (thy_id1, thy_id2) then cert2 else error ("Cannot join unrelated theory certificates " ^ display_name thy_id1 ^ " and " ^ display_name thy_id2) end; (*** generic context ***) datatype generic = Theory of theory | Proof of Proof.context; fun cases f _ (Theory thy) = f thy | cases _ g (Proof prf) = g prf; fun mapping f g = cases (Theory o f) (Proof o g); fun mapping_result f g = cases (apsnd Theory o f) (apsnd Proof o g); val the_theory = cases I (fn _ => error "Ill-typed context: theory expected"); val the_proof = cases (fn _ => error "Ill-typed context: proof expected") I; fun map_theory f = Theory o f o the_theory; fun map_proof f = Proof o f o the_proof; fun map_theory_result f = apsnd Theory o f o the_theory; fun map_proof_result f = apsnd Proof o f o the_proof; fun theory_map f = the_theory o f o Theory; fun proof_map f = the_proof o f o Proof; val theory_of = cases I Proof_Context.theory_of; val proof_of = cases Proof_Context.init_global I; (** thread data **) local val generic_context_var = Thread_Data.var () : generic Thread_Data.var in fun get_generic_context () = Thread_Data.get generic_context_var; val put_generic_context = Thread_Data.put generic_context_var; fun setmp_generic_context opt_context = Thread_Data.setmp generic_context_var opt_context; fun the_generic_context () = (case get_generic_context () of SOME context => context | _ => error "Unknown context"); val the_global_context = theory_of o the_generic_context; val the_local_context = proof_of o the_generic_context; end; fun >>> f = let val (res, context') = f (the_generic_context ()); val _ = put_generic_context (SOME context'); in res end; nonfix >>; fun >> f = >>> (fn context => ((), f context)); val _ = put_generic_context (SOME (Theory pre_pure_thy)); end; structure Basic_Context: BASIC_CONTEXT = Context; open Basic_Context; (*** type-safe interfaces for data declarations ***) (** theory data **) signature THEORY_DATA'_ARGS = sig type T val empty: T val extend: T -> T val merge: theory * theory -> T * T -> T end; signature THEORY_DATA_ARGS = sig type T val empty: T val extend: T -> T val merge: T * T -> T end; signature THEORY_DATA = sig type T val get: theory -> T val put: T -> theory -> theory val map: (T -> T) -> theory -> theory end; functor Theory_Data'(Data: THEORY_DATA'_ARGS): THEORY_DATA = struct type T = Data.T; exception Data of T; val kind = let val pos = Position.thread_data () in Context.Theory_Data.declare pos (Data Data.empty) (fn Data x => let val y = Data.extend x; val _ = if pointer_eq (x, y) then () else raise Fail ("Theory_Data extend needs to be identity" ^ Position.here pos) in Data y end) (fn thys => fn (Data x1, Data x2) => Data (Data.merge thys (x1, x2))) end; val get = Context.Theory_Data.get kind (fn Data x => x); val put = Context.Theory_Data.put kind Data; fun map f thy = put (f (get thy)) thy; end; functor Theory_Data(Data: THEORY_DATA_ARGS): THEORY_DATA = Theory_Data' ( type T = Data.T; val empty = Data.empty; val extend = Data.extend; fun merge _ = Data.merge; ); (** proof data **) signature PROOF_DATA_ARGS = sig type T val init: theory -> T end; signature PROOF_DATA = sig type T val get: Proof.context -> T val put: T -> Proof.context -> Proof.context val map: (T -> T) -> Proof.context -> Proof.context end; functor Proof_Data(Data: PROOF_DATA_ARGS): PROOF_DATA = struct type T = Data.T; exception Data of T; val kind = Context.Proof_Data.declare (Data o Data.init); val get = Context.Proof_Data.get kind (fn Data x => x); val put = Context.Proof_Data.put kind Data; fun map f prf = put (f (get prf)) prf; end; (** generic data **) signature GENERIC_DATA_ARGS = sig type T val empty: T val extend: T -> T val merge: T * T -> T end; signature GENERIC_DATA = sig type T val get: Context.generic -> T val put: T -> Context.generic -> Context.generic val map: (T -> T) -> Context.generic -> Context.generic end; functor Generic_Data(Data: GENERIC_DATA_ARGS): GENERIC_DATA = struct structure Thy_Data = Theory_Data(Data); structure Prf_Data = Proof_Data(type T = Data.T val init = Thy_Data.get); type T = Data.T; fun get (Context.Theory thy) = Thy_Data.get thy | get (Context.Proof prf) = Prf_Data.get prf; fun put x (Context.Theory thy) = Context.Theory (Thy_Data.put x thy) | put x (Context.Proof prf) = Context.Proof (Prf_Data.put x prf); fun map f ctxt = put (f (get ctxt)) ctxt; end; (*hide private interface*) structure Context: CONTEXT = Context; diff --git a/src/Pure/defs.ML b/src/Pure/defs.ML --- a/src/Pure/defs.ML +++ b/src/Pure/defs.ML @@ -1,283 +1,283 @@ (* Title: Pure/defs.ML Author: Makarius Global well-formedness checks for overloaded definitions (mixed constants and types). Recall that constant definitions may be explained syntactically within Pure, but type definitions require particular set-theoretic semantics. *) signature DEFS = sig datatype item_kind = Const | Type type item = item_kind * string type entry = item * typ list val item_kind_ord: item_kind ord val plain_args: typ list -> bool type context = Proof.context * (Name_Space.T * Name_Space.T) val global_context: theory -> context val space: context -> item_kind -> Name_Space.T val pretty_item: context -> item -> Pretty.T val pretty_args: Proof.context -> typ list -> Pretty.T list val pretty_entry: context -> entry -> Pretty.T type T type spec = {def: string option, description: string, pos: Position.T, lhs: typ list, rhs: entry list} val all_specifications_of: T -> (item * spec list) list val specifications_of: T -> item -> spec list val dest: T -> {restricts: (entry * string) list, reducts: (entry * entry list) list} val dest_constdefs: T list -> T -> (string * string) list val empty: T val merge: context -> T * T -> T val define: context -> bool -> string option -> string -> entry -> entry list -> T -> T val get_deps: T -> item -> (typ list * entry list) list end; structure Defs: DEFS = struct (* specification items *) datatype item_kind = Const | Type; type item = item_kind * string; type entry = item * typ list; fun item_kind_ord (Const, Type) = LESS | item_kind_ord (Type, Const) = GREATER | item_kind_ord _ = EQUAL; structure Itemtab = Table(type key = item val ord = prod_ord item_kind_ord fast_string_ord); (* pretty printing *) type context = Proof.context * (Name_Space.T * Name_Space.T); fun global_context thy = (Syntax.init_pretty_global thy, (Sign.const_space thy, Sign.type_space thy)); fun space ((_, spaces): context) kind = if kind = Const then #1 spaces else #2 spaces; fun pretty_item (context as (ctxt, _)) (kind, name) = let val prt_name = Name_Space.pretty ctxt (space context kind) name in if kind = Const then prt_name else Pretty.block [Pretty.keyword1 "type", Pretty.brk 1, prt_name] end; fun pretty_args ctxt args = if null args then [] else [Pretty.list "(" ")" (map (Syntax.pretty_typ ctxt o Logic.unvarifyT_global) args)]; fun pretty_entry context (c, args) = Pretty.block (pretty_item context c :: pretty_args (#1 context) args); (* type arguments *) fun plain_args args = forall Term.is_TVar args andalso not (has_duplicates (op =) args); fun disjoint_args (Ts, Us) = not (Type.could_unifys (Ts, Us)) orelse - ((Type.raw_unifys (Ts, map (Logic.incr_tvar (maxidx_of_typs Ts + 1)) Us) Vartab.empty; false) + ((Vartab.build (Type.raw_unifys (Ts, map (Logic.incr_tvar (maxidx_of_typs Ts + 1)) Us)); false) handle Type.TUNIFY => true); fun match_args (Ts, Us) = if Type.could_matches (Ts, Us) then Option.map Envir.subst_type - (SOME (Type.raw_matches (Ts, Us) Vartab.empty) handle Type.TYPE_MATCH => NONE) + (SOME (Vartab.build (Type.raw_matches (Ts, Us))) handle Type.TYPE_MATCH => NONE) else NONE; (* datatype defs *) type spec = {def: string option, description: string, pos: Position.T, lhs: typ list, rhs: entry list}; type def = {specs: spec Inttab.table, (*source specifications*) restricts: (typ list * string) list, (*global restrictions imposed by incomplete patterns*) reducts: (typ list * entry list) list}; (*specifications as reduction system*) fun make_def (specs, restricts, reducts) = {specs = specs, restricts = restricts, reducts = reducts}: def; fun map_def c f = Itemtab.default (c, make_def (Inttab.empty, [], [])) #> Itemtab.map_entry c (fn {specs, restricts, reducts}: def => make_def (f (specs, restricts, reducts))); datatype T = Defs of def Itemtab.table; fun lookup_list which defs c = (case Itemtab.lookup defs c of SOME (def: def) => which def | NONE => []); fun all_specifications_of (Defs defs) = (map o apsnd) (map snd o Inttab.dest o #specs) (Itemtab.dest defs); fun specifications_of (Defs defs) = lookup_list (map snd o Inttab.dest o #specs) defs; val restricts_of = lookup_list #restricts; val reducts_of = lookup_list #reducts; fun dest (Defs defs) = let val restricts = Itemtab.fold (fn (c, {restricts, ...}) => fold (fn (args, description) => cons ((c, args), description)) restricts) defs []; val reducts = Itemtab.fold (fn (c, {reducts, ...}) => fold (fn (args, deps) => cons ((c, args), deps)) reducts) defs []; in {restricts = restricts, reducts = reducts} end; fun dest_constdefs prevs (Defs defs) = let fun prev_spec c i = prevs |> exists (fn Defs prev_defs => (case Itemtab.lookup prev_defs c of NONE => false | SOME {specs, ...} => Inttab.defined specs i)); in - (defs, []) |-> Itemtab.fold (fn (c, {specs, ...}) => + build (defs |> Itemtab.fold (fn (c, {specs, ...}) => specs |> Inttab.fold (fn (i, spec) => if #1 c = Const andalso is_some (#def spec) andalso not (prev_spec c i) - then cons (#2 c, the (#def spec)) else I)) + then cons (#2 c, the (#def spec)) else I))) end; val empty = Defs Itemtab.empty; (* specifications *) fun disjoint_specs context c (i, {description = a, pos = pos_a, lhs = Ts, ...}: spec) = Inttab.forall (fn (j, {description = b, pos = pos_b, lhs = Us, ...}: spec) => i = j orelse disjoint_args (Ts, Us) orelse error ("Clash of specifications for " ^ Pretty.unformatted_string_of (pretty_item context c) ^ ":\n" ^ " " ^ quote a ^ Position.here pos_a ^ "\n" ^ " " ^ quote b ^ Position.here pos_b)); fun join_specs context c ({specs = specs1, restricts, reducts}, {specs = specs2, ...}: def) = let val specs' = Inttab.fold (fn spec2 => (disjoint_specs context c spec2 specs1; Inttab.update spec2)) specs2 specs1; in make_def (specs', restricts, reducts) end; fun update_specs context c spec = map_def c (fn (specs, restricts, reducts) => (disjoint_specs context c spec specs; (Inttab.update spec specs, restricts, reducts))); (* normalized dependencies: reduction with well-formedness check *) local val prt = Pretty.string_of oo pretty_entry; fun err context (c, Ts) (d, Us) s1 s2 = error (s1 ^ " dependency of " ^ prt context (c, Ts) ^ " -> " ^ prt context (d, Us) ^ s2); fun acyclic context (c, Ts) (d, Us) = c <> d orelse is_none (match_args (Ts, Us)) orelse err context (c, Ts) (d, Us) "Circular" ""; fun reduction context defs const deps = let fun reduct Us (Ts, rhs) = (case match_args (Ts, Us) of NONE => NONE | SOME subst => SOME (map (apsnd (map subst)) rhs)); fun reducts (d, Us) = get_first (reduct Us) (reducts_of defs d); val reds = map (`reducts) deps; val deps' = if forall (is_none o #1) reds then NONE else SOME (fold_rev (fn (NONE, dp) => insert (op =) dp | (SOME dps, _) => fold (insert (op =)) dps) reds []); val _ = forall (acyclic context const) (the_default deps deps'); in deps' end; fun restriction context defs (c, Ts) (d, Us) = plain_args Us orelse (case find_first (fn (Rs, _) => not (disjoint_args (Rs, Us))) (restricts_of defs d) of SOME (Rs, description) => err context (c, Ts) (d, Us) "Malformed" ("\n(restriction " ^ prt context (d, Rs) ^ " from " ^ quote description ^ ")") | NONE => true); in fun normalize context = let fun check_def defs (c, {reducts, ...}: def) = reducts |> forall (fn (Ts, deps) => forall (restriction context defs (c, Ts)) deps); fun check_defs defs = Itemtab.forall (check_def defs) defs; fun norm_update (c, {reducts, ...}: def) (changed, defs) = let val reducts' = reducts |> map (fn (Ts, deps) => (Ts, perhaps (reduction context defs (c, Ts)) deps)); in if reducts = reducts' then (changed, defs) else (true, defs |> map_def c (fn (specs, restricts, _) => (specs, restricts, reducts'))) end; fun norm_loop defs = (case Itemtab.fold norm_update defs (false, defs) of (true, defs') => norm_loop defs' | (false, _) => defs); in norm_loop #> tap check_defs end; fun dependencies context (c, args) restr deps = map_def c (fn (specs, restricts, reducts) => let val restricts' = Library.merge (op =) (restricts, restr); val reducts' = insert (op =) (args, deps) reducts; in (specs, restricts', reducts') end) #> normalize context; end; (* merge *) fun merge context (Defs defs1, Defs defs2) = let fun add_deps (c, args) restr deps defs = if AList.defined (op =) (reducts_of defs c) args then defs else dependencies context (c, args) restr deps defs; fun add_def (c, {restricts, reducts, ...}: def) = fold (fn (args, deps) => add_deps (c, args) restricts deps) reducts; in Defs (Itemtab.join (join_specs context) (defs1, defs2) |> normalize context |> Itemtab.fold add_def defs2) end; (* define *) fun define context unchecked def description (c, args) deps (Defs defs) = let val pos = Position.thread_data (); val restr = if plain_args args orelse (case args of [Term.Type (_, rec_args)] => plain_args rec_args | _ => false) then [] else [(args, description)]; val spec = (serial (), {def = def, description = description, pos = pos, lhs = args, rhs = deps}); val defs' = defs |> update_specs context c spec; in Defs (defs' |> (if unchecked then I else dependencies context (c, args) restr deps)) end; fun get_deps (Defs defs) c = reducts_of defs c; end; diff --git a/src/Pure/drule.ML b/src/Pure/drule.ML --- a/src/Pure/drule.ML +++ b/src/Pure/drule.ML @@ -1,836 +1,834 @@ (* Title: Pure/drule.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Derived rules and other operations on theorems. *) infix 0 RL RLN MRS OF COMP INCR_COMP COMP_INCR; signature BASIC_DRULE = sig val mk_implies: cterm * cterm -> cterm val list_implies: cterm list * cterm -> cterm val strip_imp_prems: cterm -> cterm list val strip_imp_concl: cterm -> cterm val cprems_of: thm -> cterm list val forall_intr_list: cterm list -> thm -> thm - val forall_intr_vars: thm -> thm val forall_elim_list: cterm list -> thm -> thm val lift_all: Proof.context -> cterm -> thm -> thm val implies_elim_list: thm -> thm list -> thm val implies_intr_list: cterm list -> thm -> thm val instantiate_normalize: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> thm -> thm val instantiate'_normalize: ctyp option list -> cterm option list -> thm -> thm val infer_instantiate_types: Proof.context -> ((indexname * typ) * cterm) list -> thm -> thm val infer_instantiate: Proof.context -> (indexname * cterm) list -> thm -> thm val infer_instantiate': Proof.context -> cterm option list -> thm -> thm val zero_var_indexes_list: thm list -> thm list val zero_var_indexes: thm -> thm val implies_intr_hyps: thm -> thm val rotate_prems: int -> thm -> thm val rearrange_prems: int list -> thm -> thm val RLN: thm list * (int * thm list) -> thm list val RL: thm list * thm list -> thm list val MRS: thm list * thm -> thm val OF: thm * thm list -> thm val COMP: thm * thm -> thm val INCR_COMP: thm * thm -> thm val COMP_INCR: thm * thm -> thm val size_of_thm: thm -> int val reflexive_thm: thm val symmetric_thm: thm val transitive_thm: thm val extensional: thm -> thm val asm_rl: thm val cut_rl: thm val revcut_rl: thm val thin_rl: thm end; signature DRULE = sig include BASIC_DRULE val outer_params: term -> (string * typ) list val generalize: Symtab.set * Symtab.set -> thm -> thm val list_comb: cterm * cterm list -> cterm val strip_comb: cterm -> cterm * cterm list val beta_conv: cterm -> cterm -> cterm val flexflex_unique: Proof.context option -> thm -> thm val export_without_context: thm -> thm val export_without_context_open: thm -> thm val store_thm: binding -> thm -> thm val store_standard_thm: binding -> thm -> thm val store_thm_open: binding -> thm -> thm val store_standard_thm_open: binding -> thm -> thm val multi_resolve: Proof.context option -> thm list -> thm -> thm Seq.seq val multi_resolves: Proof.context option -> thm list -> thm list -> thm Seq.seq val compose: thm * int * thm -> thm val equals_cong: thm val imp_cong: thm val swap_prems_eq: thm val imp_cong_rule: thm -> thm -> thm val arg_cong_rule: cterm -> thm -> thm val binop_cong_rule: cterm -> thm -> thm -> thm val fun_cong_rule: thm -> cterm -> thm val beta_eta_conversion: cterm -> thm val eta_contraction_rule: thm -> thm val norm_hhf_eq: thm val norm_hhf_eqs: thm list val is_norm_hhf: term -> bool val norm_hhf: theory -> term -> term val norm_hhf_cterm: Proof.context -> cterm -> cterm val protect: cterm -> cterm val protectI: thm val protectD: thm val protect_cong: thm val implies_intr_protected: cterm list -> thm -> thm val termI: thm val mk_term: cterm -> thm val dest_term: thm -> cterm val cterm_rule: (thm -> thm) -> cterm -> cterm - val cterm_add_frees: cterm -> cterm list -> cterm list - val cterm_add_vars: cterm -> cterm list -> cterm list val dummy_thm: thm val free_dummy_thm: thm val is_sort_constraint: term -> bool val sort_constraintI: thm val sort_constraint_eq: thm val with_subgoal: int -> (thm -> thm) -> thm -> thm val comp_no_flatten: thm * int -> int -> thm -> thm val rename_bvars: (string * string) list -> thm -> thm val rename_bvars': string option list -> thm -> thm val incr_indexes: thm -> thm -> thm val incr_indexes2: thm -> thm -> thm -> thm val triv_forall_equality: thm val distinct_prems_rl: thm val equal_intr_rule: thm val equal_elim_rule1: thm val equal_elim_rule2: thm val remdups_rl: thm val abs_def: thm -> thm end; structure Drule: DRULE = struct (** some cterm->cterm operations: faster than calling cterm_of! **) (* A1\...An\B goes to [A1,...,An], where B is not an implication *) fun strip_imp_prems ct = let val (cA, cB) = Thm.dest_implies ct in cA :: strip_imp_prems cB end handle TERM _ => []; (* A1\...An\B goes to B, where B is not an implication *) fun strip_imp_concl ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => strip_imp_concl (Thm.dest_arg ct) | _ => ct); (*The premises of a theorem, as a cterm list*) val cprems_of = strip_imp_prems o Thm.cprop_of; fun certify t = Thm.global_cterm_of (Context.the_global_context ()) t; val implies = certify Logic.implies; fun mk_implies (A, B) = Thm.apply (Thm.apply implies A) B; (*cterm version of list_implies: [A1,...,An], B goes to \A1;...;An\\B *) fun list_implies([], B) = B | list_implies(A::As, B) = mk_implies (A, list_implies(As,B)); (*cterm version of list_comb: maps (f, [t1,...,tn]) to f(t1,...,tn) *) fun list_comb (f, []) = f | list_comb (f, t::ts) = list_comb (Thm.apply f t, ts); (*cterm version of strip_comb: maps f(t1,...,tn) to (f, [t1,...,tn]) *) fun strip_comb ct = let fun stripc (p as (ct, cts)) = let val (ct1, ct2) = Thm.dest_comb ct in stripc (ct1, ct2 :: cts) end handle CTERM _ => p in stripc (ct, []) end; (*Beta-conversion for cterms, where x is an abstraction. Simply returns the rhs of the meta-equality returned by the beta_conversion rule.*) fun beta_conv x y = Thm.dest_arg (Thm.cprop_of (Thm.beta_conversion false (Thm.apply x y))); (** Standardization of rules **) (*Generalization over a list of variables*) val forall_intr_list = fold_rev Thm.forall_intr; -(*Generalization over Vars -- canonical order*) -fun forall_intr_vars th = fold Thm.forall_intr (Thm.add_vars th []) th; - fun outer_params t = let val vs = Term.strip_all_vars t in Name.variant_list [] (map (Name.clean o #1) vs) ~~ map #2 vs end; (*lift vars wrt. outermost goal parameters -- reverses the effect of gen_all modulo higher-order unification*) fun lift_all ctxt raw_goal raw_th = let val thy = Proof_Context.theory_of ctxt; val goal = Thm.transfer_cterm thy raw_goal; val th = Thm.transfer thy raw_th; val maxidx = Thm.maxidx_of th; val ps = outer_params (Thm.term_of goal) |> map (fn (x, T) => Var ((x, maxidx + 1), Logic.incr_tvar (maxidx + 1) T)); val Ts = map Term.fastype_of ps; val inst = - Thm.fold_terms Term.add_vars th [] - |> map (fn (xi, T) => ((xi, T), Thm.cterm_of ctxt (Term.list_comb (Var (xi, Ts ---> T), ps)))); + Term_Subst.Vars.build (th |> (Thm.fold_terms {hyps = false} o Term.fold_aterms) + (fn t => fn inst => + (case t of + Var (xi, T) => + if Term_Subst.Vars.defined inst (xi, T) then inst + else + let val ct = Thm.cterm_of ctxt (Term.list_comb (Var (xi, Ts ---> T), ps)) + in Term_Subst.Vars.update ((xi, T), ct) inst end + | _ => inst))); in th - |> Thm.instantiate ([], inst) + |> Thm.instantiate ([], Term_Subst.Vars.dest inst) |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) ps end; (*direct generalization*) fun generalize names th = Thm.generalize names (Thm.maxidx_of th + 1) th; (*specialization over a list of cterms*) val forall_elim_list = fold Thm.forall_elim; (*maps A1,...,An |- B to \A1;...;An\ \ B*) val implies_intr_list = fold_rev Thm.implies_intr; (*maps \A1;...;An\ \ B and [A1,...,An] to B*) fun implies_elim_list impth ths = fold Thm.elim_implies ths impth; (*Reset Var indexes to zero, renaming to preserve distinctness*) fun zero_var_indexes_list [] = [] | zero_var_indexes_list ths = let val (instT, inst) = Term_Subst.zero_var_indexes_inst Name.context (map Thm.full_prop_of ths); - val tvars = fold Thm.add_tvars ths []; - fun the_tvar v = the (find_first (fn cT => v = dest_TVar (Thm.typ_of cT)) tvars); + val tvars = fold Thm.add_tvars ths Term_Subst.TVars.empty; + val the_tvar = the o Term_Subst.TVars.lookup tvars; val instT' = - (instT, []) |-> Term_Subst.TVars.fold (fn (v, TVar (b, _)) => - cons (v, Thm.rename_tvar b (the_tvar v))); + build (instT |> Term_Subst.TVars.fold (fn (v, TVar (b, _)) => + cons (v, Thm.rename_tvar b (the_tvar v)))); - val vars = fold (Thm.add_vars o Thm.instantiate (instT', [])) ths []; - fun the_var v = the (find_first (fn ct => v = dest_Var (Thm.term_of ct)) vars); + val vars = fold (Thm.add_vars o Thm.instantiate (instT', [])) ths Term_Subst.Vars.empty; + val the_var = the o Term_Subst.Vars.lookup vars; val inst' = - (inst, []) |-> Term_Subst.Vars.fold (fn (v, Var (b, _)) => - cons (v, Thm.var (b, Thm.ctyp_of_cterm (the_var v)))); + build (inst |> Term_Subst.Vars.fold (fn (v, Var (b, _)) => + cons (v, Thm.var (b, Thm.ctyp_of_cterm (the_var v))))); in map (Thm.adjust_maxidx_thm ~1 o Thm.instantiate (instT', inst')) ths end; val zero_var_indexes = singleton zero_var_indexes_list; (** Standard form of object-rule: no hypotheses, flexflex constraints, Frees, or outer quantifiers; all generality expressed by Vars of index 0.**) (*Discharge all hypotheses.*) fun implies_intr_hyps th = fold Thm.implies_intr (Thm.chyps_of th) th; (*Squash a theorem's flexflex constraints provided it can be done uniquely. This step can lose information.*) fun flexflex_unique opt_ctxt th = if null (Thm.tpairs_of th) then th else (case distinct Thm.eq_thm (Seq.list_of (Thm.flexflex_rule opt_ctxt th)) of [th] => th | [] => raise THM ("flexflex_unique: impossible constraints", 0, [th]) | _ => raise THM ("flexflex_unique: multiple unifiers", 0, [th])); (* old-style export without context *) val export_without_context_open = implies_intr_hyps #> Thm.forall_intr_frees #> `Thm.maxidx_of #-> (fn maxidx => Thm.forall_elim_vars (maxidx + 1) #> Thm.strip_shyps #> zero_var_indexes #> Thm.varifyT_global); val export_without_context = flexflex_unique NONE #> export_without_context_open #> Thm.close_derivation \<^here>; (*Rotates a rule's premises to the left by k*) fun rotate_prems 0 = I | rotate_prems k = Thm.permute_prems 0 k; fun with_subgoal i f = rotate_prems (i - 1) #> f #> rotate_prems (1 - i); (*Permute prems, where the i-th position in the argument list (counting from 0) gives the position within the original thm to be transferred to position i. Any remaining trailing positions are left unchanged.*) val rearrange_prems = let fun rearr new [] thm = thm | rearr new (p :: ps) thm = rearr (new + 1) (map (fn q => if new <= q andalso q < p then q + 1 else q) ps) (Thm.permute_prems (new + 1) (new - p) (Thm.permute_prems new (p - new) thm)) in rearr 0 end; (*Resolution: multiple arguments, multiple results*) local fun res opt_ctxt th i rule = (Thm.biresolution opt_ctxt false [(false, th)] i rule handle THM _ => Seq.empty) |> Seq.map Thm.solve_constraints; fun multi_res _ _ [] rule = Seq.single rule | multi_res opt_ctxt i (th :: ths) rule = Seq.maps (res opt_ctxt th i) (multi_res opt_ctxt (i + 1) ths rule); in fun multi_resolve opt_ctxt = multi_res opt_ctxt 1; fun multi_resolves opt_ctxt facts rules = Seq.maps (multi_resolve opt_ctxt facts) (Seq.of_list rules); end; (*For joining lists of rules*) fun thas RLN (i, thbs) = let val resolve = Thm.biresolution NONE false (map (pair false) thas) i fun resb thb = Seq.list_of (resolve thb) handle THM _ => [] in maps resb thbs |> map Thm.solve_constraints end; fun thas RL thbs = thas RLN (1, thbs); (*Isar-style multi-resolution*) fun bottom_rl OF rls = (case Seq.chop 2 (multi_resolve NONE rls bottom_rl) of ([th], _) => Thm.solve_constraints th | ([], _) => raise THM ("OF: no unifiers", 0, bottom_rl :: rls) | _ => raise THM ("OF: multiple unifiers", 0, bottom_rl :: rls)); (*Resolve a list of rules against bottom_rl from right to left; makes proof trees*) fun rls MRS bottom_rl = bottom_rl OF rls; (*compose Q and \...,Qi,Q(i+1),...\ \ R to \...,Q(i+1),...\ \ R with no lifting or renaming! Q may contain \ or meta-quantifiers ALWAYS deletes premise i *) fun compose (tha, i, thb) = Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, tha, 0) i thb |> Seq.list_of |> distinct Thm.eq_thm |> (fn [th] => Thm.solve_constraints th | _ => raise THM ("compose: unique result expected", i, [tha, thb])); (** theorem equality **) (*Useful "distance" function for BEST_FIRST*) val size_of_thm = size_of_term o Thm.full_prop_of; (*** Meta-Rewriting Rules ***) val read_prop = certify o Simple_Syntax.read_prop; fun store_thm name th = Context.>>> (Context.map_theory_result (Global_Theory.store_thm (name, th))); fun store_thm_open name th = Context.>>> (Context.map_theory_result (Global_Theory.store_thm_open (name, th))); fun store_standard_thm name th = store_thm name (export_without_context th); fun store_standard_thm_open name th = store_thm_open name (export_without_context_open th); val reflexive_thm = let val cx = certify (Var(("x",0),TVar(("'a",0),[]))) in store_standard_thm_open (Binding.make ("reflexive", \<^here>)) (Thm.reflexive cx) end; val symmetric_thm = let val xy = read_prop "x::'a \ y::'a"; val thm = Thm.implies_intr xy (Thm.symmetric (Thm.assume xy)); in store_standard_thm_open (Binding.make ("symmetric", \<^here>)) thm end; val transitive_thm = let val xy = read_prop "x::'a \ y::'a"; val yz = read_prop "y::'a \ z::'a"; val xythm = Thm.assume xy; val yzthm = Thm.assume yz; val thm = Thm.implies_intr yz (Thm.transitive xythm yzthm); in store_standard_thm_open (Binding.make ("transitive", \<^here>)) thm end; fun extensional eq = let val eq' = Thm.abstract_rule "x" (Thm.dest_arg (fst (Thm.dest_equals (Thm.cprop_of eq)))) eq in Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of eq')) eq' end; val equals_cong = store_standard_thm_open (Binding.make ("equals_cong", \<^here>)) (Thm.reflexive (read_prop "x::'a \ y::'a")); val imp_cong = let val ABC = read_prop "A \ B::prop \ C::prop" val AB = read_prop "A \ B" val AC = read_prop "A \ C" val A = read_prop "A" in store_standard_thm_open (Binding.make ("imp_cong", \<^here>)) (Thm.implies_intr ABC (Thm.equal_intr (Thm.implies_intr AB (Thm.implies_intr A (Thm.equal_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.implies_elim (Thm.assume AB) (Thm.assume A))))) (Thm.implies_intr AC (Thm.implies_intr A (Thm.equal_elim (Thm.symmetric (Thm.implies_elim (Thm.assume ABC) (Thm.assume A))) (Thm.implies_elim (Thm.assume AC) (Thm.assume A))))))) end; val swap_prems_eq = let val ABC = read_prop "A \ B \ C" val BAC = read_prop "B \ A \ C" val A = read_prop "A" val B = read_prop "B" in store_standard_thm_open (Binding.make ("swap_prems_eq", \<^here>)) (Thm.equal_intr (Thm.implies_intr ABC (Thm.implies_intr B (Thm.implies_intr A (Thm.implies_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.assume B))))) (Thm.implies_intr BAC (Thm.implies_intr A (Thm.implies_intr B (Thm.implies_elim (Thm.implies_elim (Thm.assume BAC) (Thm.assume B)) (Thm.assume A)))))) end; val imp_cong_rule = Thm.combination o Thm.combination (Thm.reflexive implies); fun arg_cong_rule ct th = Thm.combination (Thm.reflexive ct) th; (*AP_TERM in LCF/HOL*) fun fun_cong_rule th ct = Thm.combination th (Thm.reflexive ct); (*AP_THM in LCF/HOL*) fun binop_cong_rule ct th1 th2 = Thm.combination (arg_cong_rule ct th1) th2; fun beta_eta_conversion ct = let val thm = Thm.beta_conversion true ct in Thm.transitive thm (Thm.eta_conversion (Thm.rhs_of thm)) end; (*Contract all eta-redexes in the theorem, lest they give rise to needless abstractions*) fun eta_contraction_rule th = Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of th)) th; (* abs_def *) (* f ?x1 ... ?xn \ u -------------------- f \ \x1 ... xn. u *) local fun contract_lhs th = Thm.transitive (Thm.symmetric (beta_eta_conversion (#1 (Thm.dest_equals (Thm.cprop_of th))))) th; fun var_args ct = (case try Thm.dest_comb ct of SOME (f, arg) => (case Thm.term_of arg of Var ((x, _), _) => update (eq_snd (op aconvc)) (x, arg) (var_args f) | _ => []) | NONE => []); in fun abs_def th = let val th' = contract_lhs th; val args = var_args (Thm.lhs_of th'); in contract_lhs (fold (uncurry Thm.abstract_rule) args th') end; end; (*** Some useful meta-theorems ***) (*The rule V/V, obtains assumption solving for eresolve_tac*) val asm_rl = store_standard_thm_open (Binding.make ("asm_rl", \<^here>)) (Thm.trivial (read_prop "?psi")); (*Meta-level cut rule: \V \ W; V\ \ W *) val cut_rl = store_standard_thm_open (Binding.make ("cut_rl", \<^here>)) (Thm.trivial (read_prop "?psi \ ?theta")); (*Generalized elim rule for one conclusion; cut_rl with reversed premises: \PROP V; PROP V \ PROP W\ \ PROP W *) val revcut_rl = let val V = read_prop "V"; val VW = read_prop "V \ W"; in store_standard_thm_open (Binding.make ("revcut_rl", \<^here>)) (Thm.implies_intr V (Thm.implies_intr VW (Thm.implies_elim (Thm.assume VW) (Thm.assume V)))) end; (*for deleting an unwanted assumption*) val thin_rl = let val V = read_prop "V"; val W = read_prop "W"; val thm = Thm.implies_intr V (Thm.implies_intr W (Thm.assume W)); in store_standard_thm_open (Binding.make ("thin_rl", \<^here>)) thm end; (* (\x. PROP ?V) \ PROP ?V Allows removal of redundant parameters*) val triv_forall_equality = let val V = read_prop "V"; val QV = read_prop "\x::'a. V"; val x = certify (Free ("x", Term.aT [])); in store_standard_thm_open (Binding.make ("triv_forall_equality", \<^here>)) (Thm.equal_intr (Thm.implies_intr QV (Thm.forall_elim x (Thm.assume QV))) (Thm.implies_intr V (Thm.forall_intr x (Thm.assume V)))) end; (* (PROP ?Phi \ PROP ?Phi \ PROP ?Psi) \ (PROP ?Phi \ PROP ?Psi) *) val distinct_prems_rl = let val AAB = read_prop "Phi \ Phi \ Psi"; val A = read_prop "Phi"; in store_standard_thm_open (Binding.make ("distinct_prems_rl", \<^here>)) (implies_intr_list [AAB, A] (implies_elim_list (Thm.assume AAB) [Thm.assume A, Thm.assume A])) end; (* \PROP ?phi \ PROP ?psi; PROP ?psi \ PROP ?phi\ \ PROP ?phi \ PROP ?psi Introduction rule for \ as a meta-theorem. *) val equal_intr_rule = let val PQ = read_prop "phi \ psi"; val QP = read_prop "psi \ phi"; in store_standard_thm_open (Binding.make ("equal_intr_rule", \<^here>)) (Thm.implies_intr PQ (Thm.implies_intr QP (Thm.equal_intr (Thm.assume PQ) (Thm.assume QP)))) end; (* PROP ?phi \ PROP ?psi \ PROP ?phi \ PROP ?psi *) val equal_elim_rule1 = let val eq = read_prop "phi::prop \ psi::prop"; val P = read_prop "phi"; in store_standard_thm_open (Binding.make ("equal_elim_rule1", \<^here>)) (Thm.equal_elim (Thm.assume eq) (Thm.assume P) |> implies_intr_list [eq, P]) end; (* PROP ?psi \ PROP ?phi \ PROP ?phi \ PROP ?psi *) val equal_elim_rule2 = store_standard_thm_open (Binding.make ("equal_elim_rule2", \<^here>)) (symmetric_thm RS equal_elim_rule1); (* PROP ?phi \ PROP ?phi \ PROP ?psi \ PROP ?psi *) val remdups_rl = let val P = read_prop "phi"; val Q = read_prop "psi"; val thm = implies_intr_list [P, P, Q] (Thm.assume Q); in store_standard_thm_open (Binding.make ("remdups_rl", \<^here>)) thm end; (** embedded terms and types **) local val A = certify (Free ("A", propT)); val axiom = Thm.unvarify_axiom (Context.the_global_context ()); val prop_def = axiom "Pure.prop_def"; val term_def = axiom "Pure.term_def"; val sort_constraint_def = axiom "Pure.sort_constraint_def"; val C = Thm.lhs_of sort_constraint_def; val T = Thm.dest_arg C; val CA = mk_implies (C, A); in (* protect *) val protect = Thm.apply (certify Logic.protectC); val protectI = store_standard_thm (Binding.concealed (Binding.make ("protectI", \<^here>))) (Thm.equal_elim (Thm.symmetric prop_def) (Thm.assume A)); val protectD = store_standard_thm (Binding.concealed (Binding.make ("protectD", \<^here>))) (Thm.equal_elim prop_def (Thm.assume (protect A))); val protect_cong = store_standard_thm_open (Binding.make ("protect_cong", \<^here>)) (Thm.reflexive (protect A)); fun implies_intr_protected asms th = let val asms' = map protect asms in implies_elim_list (implies_intr_list asms th) (map (fn asm' => Thm.assume asm' RS protectD) asms') |> implies_intr_list asms' end; (* term *) val termI = store_standard_thm (Binding.concealed (Binding.make ("termI", \<^here>))) (Thm.equal_elim (Thm.symmetric term_def) (Thm.forall_intr A (Thm.trivial A))); fun mk_term ct = let val cT = Thm.ctyp_of_cterm ct; val T = Thm.typ_of cT; in Thm.instantiate ([((("'a", 0), []), cT)], [((("x", 0), T), ct)]) termI end; fun dest_term th = let val cprop = strip_imp_concl (Thm.cprop_of th) in if can Logic.dest_term (Thm.term_of cprop) then Thm.dest_arg cprop else raise THM ("dest_term", 0, [th]) end; fun cterm_rule f = dest_term o f o mk_term; -val cterm_add_frees = Thm.add_frees o mk_term; -val cterm_add_vars = Thm.add_vars o mk_term; - val dummy_thm = mk_term (certify Term.dummy_prop); val free_dummy_thm = Thm.tag_free_dummy dummy_thm; (* sort_constraint *) fun is_sort_constraint (Const ("Pure.sort_constraint", _) $ Const ("Pure.type", _)) = true | is_sort_constraint _ = false; val sort_constraintI = store_standard_thm (Binding.concealed (Binding.make ("sort_constraintI", \<^here>))) (Thm.equal_elim (Thm.symmetric sort_constraint_def) (mk_term T)); val sort_constraint_eq = store_standard_thm (Binding.concealed (Binding.make ("sort_constraint_eq", \<^here>))) (Thm.equal_intr (Thm.implies_intr CA (Thm.implies_elim (Thm.assume CA) (Thm.unvarify_global (Context.the_global_context ()) sort_constraintI))) (implies_intr_list [A, C] (Thm.assume A))); end; (* HHF normalization *) (* (PROP ?phi \ (\x. PROP ?psi x)) \ (\x. PROP ?phi \ PROP ?psi x) *) val norm_hhf_eq = let val aT = TFree ("'a", []); val x = Free ("x", aT); val phi = Free ("phi", propT); val psi = Free ("psi", aT --> propT); val cx = certify x; val cphi = certify phi; val lhs = certify (Logic.mk_implies (phi, Logic.all x (psi $ x))); val rhs = certify (Logic.all x (Logic.mk_implies (phi, psi $ x))); in Thm.equal_intr (Thm.implies_elim (Thm.assume lhs) (Thm.assume cphi) |> Thm.forall_elim cx |> Thm.implies_intr cphi |> Thm.forall_intr cx |> Thm.implies_intr lhs) (Thm.implies_elim (Thm.assume rhs |> Thm.forall_elim cx) (Thm.assume cphi) |> Thm.forall_intr cx |> Thm.implies_intr cphi |> Thm.implies_intr rhs) |> store_standard_thm_open (Binding.make ("norm_hhf_eq", \<^here>)) end; val norm_hhf_prop = Logic.dest_equals (Thm.prop_of norm_hhf_eq); val norm_hhf_eqs = [norm_hhf_eq, sort_constraint_eq]; fun is_norm_hhf (Const ("Pure.sort_constraint", _)) = false | is_norm_hhf (Const ("Pure.imp", _) $ _ $ (Const ("Pure.all", _) $ _)) = false | is_norm_hhf (Abs _ $ _) = false | is_norm_hhf (t $ u) = is_norm_hhf t andalso is_norm_hhf u | is_norm_hhf (Abs (_, _, t)) = is_norm_hhf t | is_norm_hhf _ = true; fun norm_hhf thy t = if is_norm_hhf t then t else Pattern.rewrite_term thy [norm_hhf_prop] [] t; fun norm_hhf_cterm ctxt raw_ct = let val thy = Proof_Context.theory_of ctxt; val ct = Thm.transfer_cterm thy raw_ct; val t = Thm.term_of ct; in if is_norm_hhf t then ct else Thm.cterm_of ctxt (norm_hhf thy t) end; (* var indexes *) fun incr_indexes th = Thm.incr_indexes (Thm.maxidx_of th + 1); fun incr_indexes2 th1 th2 = Thm.incr_indexes (Int.max (Thm.maxidx_of th1, Thm.maxidx_of th2) + 1); local (*compose Q and \Q1,Q2,...,Qk\ \ R to \Q2,...,Qk\ \ R getting unique result*) fun comp incremented th1 th2 = Thm.bicompose NONE {flatten = true, match = false, incremented = incremented} (false, th1, 0) 1 th2 |> Seq.list_of |> distinct Thm.eq_thm |> (fn [th] => Thm.solve_constraints th | _ => raise THM ("COMP", 1, [th1, th2])); in fun th1 COMP th2 = comp false th1 th2; fun th1 INCR_COMP th2 = comp true (incr_indexes th2 th1) th2; fun th1 COMP_INCR th2 = comp true th1 (incr_indexes th1 th2); end; fun comp_no_flatten (th, n) i rule = (case distinct Thm.eq_thm (Seq.list_of (Thm.bicompose NONE {flatten = false, match = false, incremented = true} (false, th, n) i (incr_indexes th rule))) of [th'] => Thm.solve_constraints th' | [] => raise THM ("comp_no_flatten", i, [th, rule]) | _ => raise THM ("comp_no_flatten: unique result expected", i, [th, rule])); (** variations on Thm.instantiate **) fun instantiate_normalize instpair th = Thm.adjust_maxidx_thm ~1 (Thm.instantiate instpair th COMP_INCR asm_rl); fun instantiate'_normalize Ts ts th = Thm.adjust_maxidx_thm ~1 (Thm.instantiate' Ts ts th COMP_INCR asm_rl); (*instantiation with type-inference for variables*) fun infer_instantiate_types _ [] th = th | infer_instantiate_types ctxt args raw_th = let val thy = Proof_Context.theory_of ctxt; val th = Thm.transfer thy raw_th; fun infer ((xi, T), cu) (tyenv, maxidx) = let val _ = Thm.ctyp_of ctxt T; val _ = Thm.transfer_cterm thy cu; val U = Thm.typ_of_cterm cu; val maxidx' = maxidx |> Integer.max (#2 xi) |> Term.maxidx_typ T |> Integer.max (Thm.maxidx_of_cterm cu); val (tyenv', maxidx'') = Sign.typ_unify thy (T, U) (tyenv, maxidx') handle Type.TUNIFY => let val t = Var (xi, T); val u = Thm.term_of cu; in raise THM ("infer_instantiate_types: type " ^ Syntax.string_of_typ ctxt (Envir.norm_type tyenv T) ^ " of variable " ^ Syntax.string_of_term ctxt (Term.map_types (Envir.norm_type tyenv) t) ^ "\ncannot be unified with type " ^ Syntax.string_of_typ ctxt (Envir.norm_type tyenv U) ^ " of term " ^ Syntax.string_of_term ctxt (Term.map_types (Envir.norm_type tyenv) u), 0, [th]) end; in (tyenv', maxidx'') end; val (tyenv, _) = fold infer args (Vartab.empty, 0); val instT = Vartab.fold (fn (xi, (S, T)) => cons ((xi, S), Thm.ctyp_of ctxt (Envir.norm_type tyenv T))) tyenv []; val inst = args |> map (fn ((xi, T), cu) => ((xi, Envir.norm_type tyenv T), Thm.instantiate_cterm (instT, []) (Thm.transfer_cterm thy cu))); in instantiate_normalize (instT, inst) th end handle CTERM (msg, _) => raise THM (msg, 0, [raw_th]) | TERM (msg, _) => raise THM (msg, 0, [raw_th]) | TYPE (msg, _, _) => raise THM (msg, 0, [raw_th]); fun infer_instantiate _ [] th = th | infer_instantiate ctxt args th = let val vars = Term.add_vars (Thm.full_prop_of th) []; val dups = duplicates (eq_fst op =) vars; val _ = null dups orelse raise THM ("infer_instantiate: inconsistent types for variables " ^ commas_quote (map (Syntax.string_of_term (Config.put show_types true ctxt) o Var) dups), 0, [th]); val args' = args |> map_filter (fn (xi, cu) => AList.lookup (op =) vars xi |> Option.map (fn T => ((xi, T), cu))); in infer_instantiate_types ctxt args' th end; fun infer_instantiate' ctxt args th = let - val vars = rev (Term.add_vars (Thm.full_prop_of th) []); + val vars = build_rev (Term.add_vars (Thm.full_prop_of th)); val args' = zip_options vars args handle ListPair.UnequalLengths => raise THM ("infer_instantiate': more instantiations than variables in thm", 0, [th]); in infer_instantiate_types ctxt args' th end; (** renaming of bound variables **) (* replace bound variables x_i in thm by y_i *) (* where vs = [(x_1, y_1), ..., (x_n, y_n)] *) fun rename_bvars [] thm = thm | rename_bvars vs thm = let fun rename (Abs (x, T, t)) = Abs (AList.lookup (op =) vs x |> the_default x, T, rename t) | rename (t $ u) = rename t $ rename u | rename a = a; in Thm.renamed_prop (rename (Thm.prop_of thm)) thm end; (* renaming in left-to-right order *) fun rename_bvars' xs thm = let fun rename [] t = ([], t) | rename (x' :: xs) (Abs (x, T, t)) = let val (xs', t') = rename xs t in (xs', Abs (the_default x x', T, t')) end | rename xs (t $ u) = let val (xs', t') = rename xs t; val (xs'', u') = rename xs' u; in (xs'', t' $ u') end | rename xs a = (xs, a); in (case rename xs (Thm.prop_of thm) of ([], prop') => Thm.renamed_prop prop' thm | _ => error "More names than abstractions in theorem") end; end; structure Basic_Drule: BASIC_DRULE = Drule; open Basic_Drule; diff --git a/src/Pure/envir.ML b/src/Pure/envir.ML --- a/src/Pure/envir.ML +++ b/src/Pure/envir.ML @@ -1,424 +1,424 @@ (* Title: Pure/envir.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Free-form environments. The type of a term variable / sort of a type variable is part of its name. The lookup function must apply type substitutions, since they may change the identity of a variable. *) signature ENVIR = sig type tenv = (typ * term) Vartab.table datatype env = Envir of {maxidx: int, tenv: tenv, tyenv: Type.tyenv} val maxidx_of: env -> int val term_env: env -> tenv val type_env: env -> Type.tyenv val is_empty: env -> bool val empty: int -> env val init: env val merge: env * env -> env val insert_sorts: env -> sort list -> sort list val genvars: string -> env * typ list -> env * term list val genvar: string -> env * typ -> env * term val lookup1: tenv -> indexname * typ -> term option val lookup: env -> indexname * typ -> term option val update: (indexname * typ) * term -> env -> env val above: env -> int -> bool val vupdate: (indexname * typ) * term -> env -> env val norm_type_same: Type.tyenv -> typ Same.operation val norm_types_same: Type.tyenv -> typ list Same.operation val norm_type: Type.tyenv -> typ -> typ val norm_term_same: env -> term Same.operation val norm_term: env -> term -> term val beta_norm: term -> term val head_norm: env -> term -> term val eta_long: typ list -> term -> term val eta_contract: term -> term val beta_eta_contract: term -> term val aeconv: term * term -> bool val body_type: env -> typ -> typ val binder_types: env -> typ -> typ list val strip_type: env -> typ -> typ list * typ val fastype: env -> typ list -> term -> typ val subst_type_same: Type.tyenv -> typ Same.operation val subst_term_types_same: Type.tyenv -> term Same.operation val subst_term_same: Type.tyenv * tenv -> term Same.operation val subst_type: Type.tyenv -> typ -> typ val subst_term_types: Type.tyenv -> term -> term val subst_term: Type.tyenv * tenv -> term -> term val expand_atom: typ -> typ * term -> term val expand_term: (term -> (typ * term) option) -> term -> term val expand_term_frees: ((string * typ) * term) list -> term -> term end; structure Envir: ENVIR = struct (** datatype env **) (*Updating can destroy environment in 2 ways! (1) variables out of range (2) circular assignments *) type tenv = (typ * term) Vartab.table; datatype env = Envir of {maxidx: int, (*upper bound of maximum index of vars*) tenv: tenv, (*assignments to Vars*) tyenv: Type.tyenv}; (*assignments to TVars*) fun make_env (maxidx, tenv, tyenv) = Envir {maxidx = maxidx, tenv = tenv, tyenv = tyenv}; fun maxidx_of (Envir {maxidx, ...}) = maxidx; fun term_env (Envir {tenv, ...}) = tenv; fun type_env (Envir {tyenv, ...}) = tyenv; fun is_empty env = Vartab.is_empty (term_env env) andalso Vartab.is_empty (type_env env); (* build env *) fun empty maxidx = make_env (maxidx, Vartab.empty, Vartab.empty); val init = empty ~1; fun merge (Envir {maxidx = maxidx1, tenv = tenv1, tyenv = tyenv1}, Envir {maxidx = maxidx2, tenv = tenv2, tyenv = tyenv2}) = make_env (Int.max (maxidx1, maxidx2), Vartab.merge (op =) (tenv1, tenv2), Vartab.merge (op =) (tyenv1, tyenv2)); (*NB: type unification may invent new sorts*) (* FIXME tenv!? *) val insert_sorts = Vartab.fold (fn (_, (_, T)) => Sorts.insert_typ T) o type_env; (*Generate a list of distinct variables. Increments index to make them distinct from ALL present variables. *) fun genvars name (Envir {maxidx, tenv, tyenv}, Ts) : env * term list = let fun genvs (_, [] : typ list) : term list = [] | genvs (_, [T]) = [Var ((name, maxidx + 1), T)] | genvs (n, T :: Ts) = Var ((name ^ radixstring (26, "a" , n), maxidx + 1), T) :: genvs (n + 1, Ts); in (Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}, genvs (0, Ts)) end; (*Generate a variable.*) fun genvar name (env, T) : env * term = let val (env', [v]) = genvars name (env, [T]) in (env', v) end; fun var_clash xi T T' = raise TYPE ("Variable has two distinct types", [], [Var (xi, T'), Var (xi, T)]); fun lookup_check check tenv (xi, T) = (case Vartab.lookup tenv xi of NONE => NONE | SOME (U, t) => if check (T, U) then SOME t else var_clash xi T U); (*When dealing with environments produced by matching instead of unification, there is no need to chase assigned TVars. In this case, we can simply ignore the type substitution and use = instead of eq_type.*) fun lookup1 tenv = lookup_check (op =) tenv; fun lookup (Envir {tenv, tyenv, ...}) = lookup_check (Type.unified tyenv) tenv; fun update ((xi, T), t) (Envir {maxidx, tenv, tyenv}) = Envir {maxidx = maxidx, tenv = Vartab.update_new (xi, (T, t)) tenv, tyenv = tyenv}; (*Determine if the least index updated exceeds lim*) fun above (Envir {tenv, tyenv, ...}) lim = (case Vartab.min tenv of SOME ((_, i), _) => i > lim | NONE => true) andalso (case Vartab.min tyenv of SOME ((_, i), _) => i > lim | NONE => true); (*Update, checking Var-Var assignments: try to suppress higher indexes*) fun vupdate ((a, U), t) env = (case t of Var (nT as (name', T)) => if a = name' then env (*cycle!*) else if Term_Ord.indexname_ord (a, name') = LESS then (case lookup env nT of (*if already assigned, chase*) NONE => update (nT, Var (a, T)) env | SOME u => vupdate ((a, U), u) env) else update ((a, U), t) env | _ => update ((a, U), t) env); (** beta normalization wrt. environment **) (*Chases variables in env. Does not exploit sharing of variable bindings Does not check types, so could loop.*) local fun norm_type0 tyenv : typ Same.operation = let fun norm (Type (a, Ts)) = Type (a, Same.map norm Ts) | norm (TFree _) = raise Same.SAME | norm (TVar v) = (case Type.lookup tyenv v of SOME U => Same.commit norm U | NONE => raise Same.SAME); in norm end; fun norm_term1 tenv : term Same.operation = let fun norm (Var v) = (case lookup1 tenv v of SOME u => Same.commit norm u | NONE => raise Same.SAME) | norm (Abs (a, T, body)) = Abs (a, T, norm body) | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) | norm (f $ t) = ((case norm f of Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) | nf => nf $ Same.commit norm t) handle Same.SAME => f $ norm t) | norm _ = raise Same.SAME; in norm end; fun norm_term2 (envir as Envir {tyenv, ...}) : term Same.operation = let val normT = norm_type0 tyenv; fun norm (Const (a, T)) = Const (a, normT T) | norm (Free (a, T)) = Free (a, normT T) | norm (Var (xi, T)) = (case lookup envir (xi, T) of SOME u => Same.commit norm u | NONE => Var (xi, normT T)) | norm (Abs (a, T, body)) = (Abs (a, normT T, Same.commit norm body) handle Same.SAME => Abs (a, T, norm body)) | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) | norm (f $ t) = ((case norm f of Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) | nf => nf $ Same.commit norm t) handle Same.SAME => f $ norm t) | norm _ = raise Same.SAME; in norm end; in fun norm_type_same tyenv T = if Vartab.is_empty tyenv then raise Same.SAME else norm_type0 tyenv T; fun norm_types_same tyenv Ts = if Vartab.is_empty tyenv then raise Same.SAME else Same.map (norm_type0 tyenv) Ts; fun norm_type tyenv T = norm_type_same tyenv T handle Same.SAME => T; fun norm_term_same (envir as Envir {tenv, tyenv, ...}) = if Vartab.is_empty tyenv then norm_term1 tenv else norm_term2 envir; fun norm_term envir t = norm_term_same envir t handle Same.SAME => t; fun beta_norm t = if Term.could_beta_contract t then norm_term init t else t; end; (* head normal form for unification *) fun head_norm env = let fun norm (Var v) = (case lookup env v of SOME u => head_norm env u | NONE => raise Same.SAME) | norm (Abs (a, T, body)) = Abs (a, T, norm body) | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) | norm (f $ t) = (case norm f of Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) | nf => nf $ t) | norm _ = raise Same.SAME; in Same.commit norm end; (* eta-long beta-normal form *) fun eta_long Ts (Abs (s, T, t)) = Abs (s, T, eta_long (T :: Ts) t) | eta_long Ts t = (case strip_comb t of (Abs _, _) => eta_long Ts (beta_norm t) | (u, ts) => let val Us = binder_types (fastype_of1 (Ts, t)); val i = length Us; val long = eta_long (rev Us @ Ts); in fold_rev (Term.abs o pair "x") Us (list_comb (incr_boundvars i u, map (long o incr_boundvars i) ts @ map (long o Bound) (i - 1 downto 0))) end); (* full eta contraction *) local fun decr lev (Bound i) = if i >= lev then Bound (i - 1) else raise Same.SAME | decr lev (Abs (a, T, body)) = Abs (a, T, decr (lev + 1) body) | decr lev (t $ u) = (decr lev t $ decrh lev u handle Same.SAME => t $ decr lev u) | decr _ _ = raise Same.SAME and decrh lev t = (decr lev t handle Same.SAME => t); fun eta (Abs (a, T, body)) = ((case eta body of body' as (f $ Bound 0) => if Term.is_dependent f then Abs (a, T, body') else decrh 0 f | body' => Abs (a, T, body')) handle Same.SAME => (case body of f $ Bound 0 => if Term.is_dependent f then raise Same.SAME else decrh 0 f | _ => raise Same.SAME)) | eta (t $ u) = (eta t $ Same.commit eta u handle Same.SAME => t $ eta u) | eta _ = raise Same.SAME; in fun eta_contract t = if Term.could_eta_contract t then Same.commit eta t else t; end; val beta_eta_contract = eta_contract o beta_norm; fun aeconv (t, u) = t aconv u orelse eta_contract t aconv eta_contract u; fun body_type env (Type ("fun", [_, T])) = body_type env T | body_type env (T as TVar v) = (case Type.lookup (type_env env) v of NONE => T | SOME T' => body_type env T') | body_type _ T = T; fun binder_types env (Type ("fun", [T, U])) = T :: binder_types env U | binder_types env (TVar v) = (case Type.lookup (type_env env) v of NONE => [] | SOME T' => binder_types env T') | binder_types _ _ = []; fun strip_type env T = (binder_types env T, body_type env T); (*finds type of term without checking that combinations are consistent Ts holds types of bound variables*) fun fastype (Envir {tyenv, ...}) = let val funerr = "fastype: expected function type"; fun fast Ts (f $ u) = (case Type.devar tyenv (fast Ts f) of Type ("fun", [_, T]) => T | TVar _ => raise TERM (funerr, [f $ u]) | _ => raise TERM (funerr, [f $ u])) | fast _ (Const (_, T)) = T | fast _ (Free (_, T)) = T | fast Ts (Bound i) = (nth Ts i handle General.Subscript => raise TERM ("fastype: Bound", [Bound i])) | fast _ (Var (_, T)) = T | fast Ts (Abs (_, T, u)) = T --> fast (T :: Ts) u; in fast end; (** plain substitution -- without variable chasing **) local fun subst_type0 tyenv = Term_Subst.map_atypsT_same (fn TVar v => (case Type.lookup tyenv v of SOME U => U | NONE => raise Same.SAME) | _ => raise Same.SAME); fun subst_term1 tenv = Term_Subst.map_aterms_same (fn Var v => (case lookup1 tenv v of SOME u => u | NONE => raise Same.SAME) | _ => raise Same.SAME); fun subst_term2 tenv tyenv : term Same.operation = let val substT = subst_type0 tyenv; fun subst (Const (a, T)) = Const (a, substT T) | subst (Free (a, T)) = Free (a, substT T) | subst (Var (xi, T)) = (case lookup1 tenv (xi, T) of SOME u => u | NONE => Var (xi, substT T)) | subst (Bound _) = raise Same.SAME | subst (Abs (a, T, t)) = (Abs (a, substT T, Same.commit subst t) handle Same.SAME => Abs (a, T, subst t)) | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); in subst end; in fun subst_type_same tyenv T = if Vartab.is_empty tyenv then raise Same.SAME else subst_type0 tyenv T; fun subst_term_types_same tyenv t = if Vartab.is_empty tyenv then raise Same.SAME else Term_Subst.map_types_same (subst_type0 tyenv) t; fun subst_term_same (tyenv, tenv) = if Vartab.is_empty tenv then subst_term_types_same tyenv else if Vartab.is_empty tyenv then subst_term1 tenv else subst_term2 tenv tyenv; fun subst_type tyenv T = subst_type_same tyenv T handle Same.SAME => T; fun subst_term_types tyenv t = subst_term_types_same tyenv t handle Same.SAME => t; fun subst_term envs t = subst_term_same envs t handle Same.SAME => t; end; (** expand defined atoms -- with local beta reduction **) fun expand_atom T (U, u) = - subst_term_types (Type.raw_match (U, T) Vartab.empty) u + subst_term_types (Vartab.build (Type.raw_match (U, T))) u handle Type.TYPE_MATCH => raise TYPE ("expand_atom: ill-typed replacement", [T, U], [u]); fun expand_term get = let fun expand tm = let val (head, args) = Term.strip_comb tm; val args' = map expand args; fun comb head' = Term.list_comb (head', args'); in (case head of Abs (x, T, t) => comb (Abs (x, T, expand t)) | _ => (case get head of SOME def => Term.betapplys (expand_atom (Term.fastype_of head) def, args') | NONE => comb head)) end; in expand end; fun expand_term_frees defs = let val eqs = map (fn ((x, U), u) => (x, (U, u))) defs; val get = fn Free (x, _) => AList.lookup (op =) eqs x | _ => NONE; in expand_term get end; end; diff --git a/src/Pure/facts.ML b/src/Pure/facts.ML --- a/src/Pure/facts.ML +++ b/src/Pure/facts.ML @@ -1,299 +1,300 @@ (* Title: Pure/facts.ML Author: Makarius Environment of named facts, optionally indexed by proposition. *) signature FACTS = sig val err_selection: string * Position.T -> int -> thm list -> 'a val err_single: string * Position.T -> thm list -> 'a val the_single: string * Position.T -> thm list -> thm datatype interval = FromTo of int * int | From of int | Single of int datatype ref = Named of (string * Position.T) * interval list option | Fact of string val named: string -> ref val ref_name: ref -> string val ref_pos: ref -> Position.T val map_ref_name: (string -> string) -> ref -> ref val string_of_selection: interval list option -> string val string_of_ref: ref -> string val select: ref -> thm list -> thm list val selections: string * thm list -> (ref * thm) list type T val empty: T val space_of: T -> Name_Space.T val alias: Name_Space.naming -> binding -> string -> T -> T val is_concealed: T -> string -> bool val check: Context.generic -> T -> xstring * Position.T -> string val intern: T -> xstring -> string val extern: Proof.context -> T -> string -> xstring val markup_extern: Proof.context -> T -> string -> Markup.T * xstring val defined: T -> string -> bool val is_dynamic: T -> string -> bool val lookup: Context.generic -> T -> string -> {dynamic: bool, thms: thm list} option val retrieve: Context.generic -> T -> xstring * Position.T -> {name: string, dynamic: bool, thms: thm list} val fold_static: (string * thm list -> 'a -> 'a) -> T -> 'a -> 'a val dest_static: bool -> T list -> T -> (string * thm list) list val dest_all: Context.generic -> bool -> T list -> T -> (string * thm list) list val props: T -> (thm * Position.T) list val could_unify: T -> term -> (thm * Position.T) list val merge: T * T -> T val add_static: Context.generic -> {strict: bool, index: bool} -> binding * thm list lazy -> T -> string * T val add_dynamic: Context.generic -> binding * (Context.generic -> thm list) -> T -> string * T val del: string -> T -> T val hide: bool -> string -> T -> T end; structure Facts: FACTS = struct (** fact references **) fun length_msg (thms: thm list) pos = " (length " ^ string_of_int (length thms) ^ ")" ^ Position.here pos; fun err_selection (name, pos) i thms = error ("Bad fact selection " ^ quote (name ^ enclose "(" ")" (string_of_int i)) ^ length_msg thms pos); fun err_single (name, pos) thms = error ("Expected singleton fact " ^ quote name ^ length_msg thms pos); fun the_single _ [thm] = thm | the_single name_pos thms = err_single name_pos thms; (* datatype interval *) datatype interval = FromTo of int * int | From of int | Single of int; fun string_of_interval (FromTo (i, j)) = string_of_int i ^ "-" ^ string_of_int j | string_of_interval (From i) = string_of_int i ^ "-" | string_of_interval (Single i) = string_of_int i; fun interval n iv = let fun err () = raise Fail ("Bad interval specification " ^ string_of_interval iv) in (case iv of FromTo (i, j) => if i <= j then i upto j else err () | From i => if i <= n then i upto n else err () | Single i => [i]) end; (* datatype ref *) datatype ref = Named of (string * Position.T) * interval list option | Fact of string; fun named name = Named ((name, Position.none), NONE); fun ref_name (Named ((name, _), _)) = name | ref_name (Fact _) = raise Fail "Illegal literal fact"; fun ref_pos (Named ((_, pos), _)) = pos | ref_pos (Fact _) = Position.none; fun map_ref_name f (Named ((name, pos), is)) = Named ((f name, pos), is) | map_ref_name _ r = r; fun string_of_selection NONE = "" | string_of_selection (SOME is) = enclose "(" ")" (commas (map string_of_interval is)); fun string_of_ref (Named ((name, _), sel)) = name ^ string_of_selection sel | string_of_ref (Fact _) = raise Fail "Illegal literal fact"; (* select *) fun select (Fact _) ths = ths | select (Named (_, NONE)) ths = ths | select (Named ((name, pos), SOME ivs)) ths = let val n = length ths; fun err msg = error (msg ^ " for fact " ^ quote name ^ length_msg ths pos); fun sel i = if i > 0 andalso i <= n then nth ths (i - 1) else err_selection (name, pos) i ths; val is = maps (interval n) ivs handle Fail msg => err msg; in map sel is end; (* selections *) fun selections (name, [th]) = [(Named ((name, Position.none), NONE), th)] | selections (name, ths) = map2 (fn i => fn th => (Named ((name, Position.none), SOME [Single i]), th)) (1 upto length ths) ths; (** fact environment **) (* datatypes *) datatype fact = Static of thm list lazy | Dynamic of Context.generic -> thm list; datatype T = Facts of {facts: fact Name_Space.table, props: (thm * Position.T) Net.net}; fun make_facts facts props = Facts {facts = facts, props = props}; val empty = make_facts (Name_Space.empty_table Markup.factN) Net.empty; (* named facts *) fun facts_of (Facts {facts, ...}) = facts; val space_of = Name_Space.space_of_table o facts_of; fun alias naming binding name (Facts {facts, props}) = make_facts (Name_Space.alias_table naming binding name facts) props; val is_concealed = Name_Space.is_concealed o space_of; fun check context facts (xname, pos) = let val (name, fact) = Name_Space.check context (facts_of facts) (xname, pos); val _ = (case fact of Static _ => () | Dynamic _ => Context_Position.report_generic context pos (Markup.dynamic_fact name)); in name end; val intern = Name_Space.intern o space_of; fun extern ctxt = Name_Space.extern ctxt o space_of; fun markup_extern ctxt = Name_Space.markup_extern ctxt o space_of; (* retrieve *) val defined = Name_Space.defined o facts_of; fun is_dynamic facts name = (case Name_Space.lookup (facts_of facts) name of SOME (Dynamic _) => true | _ => false); fun lookup context facts name = (case Name_Space.lookup (facts_of facts) name of NONE => NONE | SOME (Static ths) => SOME {dynamic = false, thms = Lazy.force ths} | SOME (Dynamic f) => SOME {dynamic = true, thms = f context}); fun retrieve context facts (xname, pos) = let val name = check context facts (xname, pos); val {dynamic, thms} = (case lookup context facts name of SOME res => (if #dynamic res then Context_Position.report_generic context pos (Markup.dynamic_fact name) else (); res) | NONE => error ("Unknown fact " ^ quote name ^ Position.here pos)); in {name = name, dynamic = dynamic, thms = map (Thm.transfer'' context) thms} end; (* content *) local fun fold_static_lazy f = Name_Space.fold_table (fn (a, Static ths) => f (a, ths) | _ => I) o facts_of; fun consolidate facts = let val unfinished = - (facts, []) |-> fold_static_lazy (fn (_, ths) => if Lazy.is_finished ths then I else cons ths); + build (facts |> fold_static_lazy (fn (_, ths) => + if Lazy.is_finished ths then I else cons ths)); val _ = Lazy.consolidate unfinished; in facts end; fun included verbose prev_facts facts name = not (exists (fn prev => defined prev name) prev_facts orelse not verbose andalso is_concealed facts name); in fun fold_static f facts = fold_static_lazy (f o apsnd Lazy.force) (consolidate facts); fun dest_static verbose prev_facts facts = fold_static (fn (a, ths) => included verbose prev_facts facts a ? cons (a, ths)) facts [] |> sort_by #1; fun dest_all context verbose prev_facts facts = (facts_of (consolidate facts), []) |-> Name_Space.fold_table (fn (a, fact) => let val ths = (case fact of Static ths => Lazy.force ths | Dynamic f => f context) in included verbose prev_facts facts a ? cons (a, ths) end) |> sort_by #1; end; (* indexed props *) val prop_ord = Term_Ord.term_ord o apply2 (Thm.full_prop_of o fst); fun props (Facts {props, ...}) = sort_distinct prop_ord (Net.content props); fun could_unify (Facts {props, ...}) = Net.unify_term props; (* merge facts *) fun merge (Facts {facts = facts1, props = props1}, Facts {facts = facts2, props = props2}) = let val facts' = Name_Space.merge_tables (facts1, facts2); val props' = if Net.is_empty props2 then props1 else if Net.is_empty props1 then props2 else Net.merge (is_equal o prop_ord) (props1, props2); (*beware of non-canonical merge*) in make_facts facts' props' end; (* add entries *) fun add_prop pos th = Net.insert_term (K false) (Thm.full_prop_of th, (th, pos)); fun add_static context {strict, index} (b, ths) (Facts {facts, props}) = let val ths' = ths |> index ? Lazy.force_value |> Lazy.map_finished (map Thm.trim_context); val (name, facts') = if Binding.is_empty b then ("", facts) else Name_Space.define context strict (b, Static ths') facts; val props' = if index then props |> fold (add_prop (Binding.pos_of (Binding.default_pos b))) (Lazy.force ths') else props; in (name, make_facts facts' props') end; fun add_dynamic context (b, f) (Facts {facts, props}) = let val (name, facts') = Name_Space.define context true (b, Dynamic f) facts; in (name, make_facts facts' props) end; (* remove entries *) fun del name (Facts {facts, props}) = make_facts (Name_Space.del_table name facts) props; fun hide fully name (Facts {facts, props}) = make_facts (Name_Space.hide_table fully name facts) props; end; diff --git a/src/Pure/global_theory.ML b/src/Pure/global_theory.ML --- a/src/Pure/global_theory.ML +++ b/src/Pure/global_theory.ML @@ -1,371 +1,371 @@ (* 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 get_thm_names: theory -> Thm_Name.T Inttab.table 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); val extend = I; 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 lazy_thm_names thy = (case thm_names_of thy of NONE => Lazy.lazy (fn () => make_thm_names thy) | SOME lazy_tab => lazy_tab); val get_thm_names = Lazy.force o lazy_thm_names; 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); in Inttab.fold_rev (fn (i, thm_name) => cons (thm_id i, thm_name)) (get_thm_names thy) [] end; fun lookup_thm_id thy = let val theories = - fold (fn thy' => Symtab.update (Context.theory_long_name thy', lazy_thm_names thy')) - (Theory.nodes_of thy) Symtab.empty; + Symtab.build (Theory.nodes_of thy |> fold (fn thy' => + Symtab.update (Context.theory_long_name thy', lazy_thm_names thy'))); 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_begin (fn thy => if is_none (thm_names_of thy) then NONE else SOME (map_thm_names (K NONE) thy)) #> 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; + Symtab.build (Theory.nodes_of thy |> fold (fn thy' => + Symtab.update (Context.theory_name thy', thy'))); 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 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/goal.ML b/src/Pure/goal.ML --- a/src/Pure/goal.ML +++ b/src/Pure/goal.ML @@ -1,338 +1,342 @@ (* Title: Pure/goal.ML Author: Makarius Goals in tactical theorem proving, with support for forked proofs. *) signature BASIC_GOAL = sig val quick_and_dirty: bool Config.T val SELECT_GOAL: tactic -> int -> tactic val PREFER_GOAL: tactic -> int -> tactic val CONJUNCTS: tactic -> int -> tactic val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic end; signature GOAL = sig include BASIC_GOAL val init: cterm -> thm val protect: int -> thm -> thm val conclude: thm -> thm val check_finished: Proof.context -> thm -> thm val finish: Proof.context -> thm -> thm val norm_result: Proof.context -> thm -> thm val skip_proofs_enabled: unit -> bool val future_result: Proof.context -> thm future -> term -> thm val prove_internal: Proof.context -> cterm list -> cterm -> (thm list -> tactic) -> thm val is_schematic: term -> bool val prove_common: Proof.context -> int option -> string list -> term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) -> thm list val prove_future: Proof.context -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove: Proof.context -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove_global_future: theory -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove_global: theory -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove_sorry: Proof.context -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove_sorry_global: theory -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val restrict: int -> int -> thm -> thm val unrestrict: int -> thm -> thm val conjunction_tac: int -> tactic val precise_conjunction_tac: int -> int -> tactic val recover_conjunction_tac: tactic val norm_hhf_tac: Proof.context -> int -> tactic val assume_rule_tac: Proof.context -> int -> tactic end; structure Goal: GOAL = struct (** goals **) (* -------- (init) C \ #C *) fun init C = Thm.instantiate ([], [((("A", 0), propT), C)]) Drule.protectI; (* A1 \ ... \ An \ C ------------------------ (protect n) A1 \ ... \ An \ #C *) fun protect n th = Drule.comp_no_flatten (th, n) 1 Drule.protectI; (* A \ ... \ #C ---------------- (conclude) A \ ... \ C *) fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD; (* #C --- (finish) C *) fun check_finished ctxt th = if Thm.no_prems th then th else raise THM ("Proof failed.\n" ^ Pretty.string_of (Goal_Display.pretty_goal ctxt th), 0, [th]); fun finish ctxt = check_finished ctxt #> conclude; (** results **) (* normal form *) fun norm_result ctxt = Drule.flexflex_unique (SOME ctxt) #> Raw_Simplifier.norm_hhf_protect ctxt #> Thm.strip_shyps #> Drule.zero_var_indexes; (* scheduling parameters *) fun skip_proofs_enabled () = let val skip = Options.default_bool "skip_proofs" in if Proofterm.proofs_enabled () andalso skip then (warning "Proof terms enabled -- cannot skip proofs"; false) else skip end; (* future_result *) fun future_result ctxt result prop = let val assms = Assumption.all_assms_of ctxt; val As = map Thm.term_of assms; - val xs = map Free (fold Term.add_frees (prop :: As) []); - val fixes = map (Thm.cterm_of ctxt) xs; + val frees = Term_Subst.Frees.build (fold Term_Subst.add_frees (prop :: As)); + val xs = Term_Subst.Frees.fold_rev (cons o Thm.cterm_of ctxt o Free o #1) frees []; - val tfrees = fold Term.add_tfrees (prop :: As) []; - val tfrees_set = fold (Symtab.insert_set o #1) tfrees Symtab.empty; - val instT = map (fn (a, S) => (((a, 0), S), Thm.ctyp_of ctxt (TFree (a, S)))) tfrees; + val tfrees = Term_Subst.TFrees.build (fold Term_Subst.add_tfrees (prop :: As)); + val Ts = Symtab.build (Term_Subst.TFrees.fold (Symtab.insert_set o #1 o #1) tfrees); + val instT = + build (tfrees |> Term_Subst.TFrees.fold (fn ((a, S), ()) => + cons (((a, 0), S), Thm.ctyp_of ctxt (TFree (a, S))))); val global_prop = - Logic.varify_types_global (fold_rev Logic.all xs (Logic.list_implies (As, prop))) + Logic.list_implies (As, prop) + |> Term_Subst.Frees.fold_rev (Logic.all o Free o #1) frees + |> Logic.varify_types_global |> Thm.cterm_of ctxt |> Thm.weaken_sorts' ctxt; val global_result = result |> Future.map (Drule.flexflex_unique (SOME ctxt) #> Drule.implies_intr_list assms #> - Drule.forall_intr_list fixes #> + Drule.forall_intr_list xs #> Thm.adjust_maxidx_thm ~1 #> - Thm.generalize (tfrees_set, Symtab.empty) 0 #> + Thm.generalize (Ts, Symtab.empty) 0 #> Thm.strip_shyps #> Thm.solve_constraints); val local_result = Thm.future global_result global_prop |> Thm.close_derivation \<^here> |> Thm.instantiate (instT, []) - |> Drule.forall_elim_list fixes + |> Drule.forall_elim_list xs |> fold (Thm.elim_implies o Thm.assume) assms |> Thm.solve_constraints; in local_result end; (** tactical theorem proving **) (* prove_internal -- minimal checks, no normalization of result! *) fun prove_internal ctxt casms cprop tac = (case SINGLE (tac (map (Assumption.assume ctxt) casms)) (init cprop) of SOME th => Drule.implies_intr_list casms (finish ctxt th) | NONE => error "Tactic failed"); (* prove variations *) fun is_schematic t = Term.exists_subterm Term.is_Var t orelse Term.exists_type (Term.exists_subtype Term.is_TVar) t; fun prove_common ctxt fork_pri xs asms props tac = let val thy = Proof_Context.theory_of ctxt; val schematic = exists is_schematic props; val immediate = is_none fork_pri; val future = Future.proofs_enabled 1 andalso not (Proofterm.proofs_enabled ()); val skip = not immediate andalso not schematic andalso future andalso skip_proofs_enabled (); val pos = Position.thread_data (); fun err msg = cat_error msg ("The error(s) above occurred for the goal statement" ^ Position.here pos ^ ":\n" ^ Syntax.string_of_term ctxt (Logic.list_implies (asms, Logic.mk_conjunction_list props))); fun cert_safe t = Thm.cterm_of ctxt (Envir.beta_norm (Term.no_dummy_patterns t)) handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg; val casms = map cert_safe asms; val cprops = map cert_safe props; val (prems, ctxt') = ctxt |> Variable.add_fixes_direct xs |> fold Variable.declare_term (asms @ props) |> Assumption.add_assumes casms ||> Variable.set_body true; val stmt = Thm.weaken_sorts' ctxt' (Conjunction.mk_conjunction_balanced cprops); fun tac' args st = if skip then ALLGOALS (Skip_Proof.cheat_tac ctxt) st before Skip_Proof.report ctxt else tac args st; fun result () = (case SINGLE (tac' {prems = prems, context = ctxt'}) (init stmt) of NONE => err "Tactic failed" | SOME st => let val _ = Context.subthy_id (Thm.theory_id st, Context.theory_id thy) orelse err "Bad background theory of goal state"; val res = (finish ctxt' st |> Drule.flexflex_unique (SOME ctxt') |> Thm.check_shyps ctxt' |> Thm.check_hyps (Context.Proof ctxt')) handle THM (msg, _, _) => err msg | ERROR msg => err msg; in if Unify.matches_list (Context.Proof ctxt') [Thm.term_of stmt] [Thm.prop_of res] then res else err ("Proved a different theorem: " ^ Syntax.string_of_term ctxt' (Thm.prop_of res)) end); val res = if immediate orelse schematic orelse not future orelse skip then result () else future_result ctxt' (Execution.fork {name = "Goal.prove", pos = Position.thread_data (), pri = the fork_pri} result) (Thm.term_of stmt); in res |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length props) |> map (Assumption.export false ctxt' ctxt) |> Variable.export ctxt' ctxt |> map Drule.zero_var_indexes end; fun prove_future_pri ctxt pri xs asms prop tac = hd (prove_common ctxt (SOME pri) xs asms [prop] tac); fun prove_future ctxt = prove_future_pri ctxt ~1; fun prove ctxt xs asms prop tac = hd (prove_common ctxt NONE xs asms [prop] tac); fun prove_global_future thy xs asms prop tac = Drule.export_without_context (prove_future (Proof_Context.init_global thy) xs asms prop tac); fun prove_global thy xs asms prop tac = Drule.export_without_context (prove (Proof_Context.init_global thy) xs asms prop tac); (* skip proofs *) val quick_and_dirty = Config.declare_option_bool ("quick_and_dirty", \<^here>); fun prove_sorry ctxt xs asms prop tac = if Config.get ctxt quick_and_dirty then prove ctxt xs asms prop (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt)) else (if Future.proofs_enabled 1 then prove_future_pri ctxt ~2 else prove ctxt) xs asms prop tac; fun prove_sorry_global thy xs asms prop tac = Drule.export_without_context (prove_sorry (Proof_Context.init_global thy) xs asms prop tac); (** goal structure **) (* rearrange subgoals *) fun restrict i n st = if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then raise THM ("Goal.restrict", i, [st]) else rotate_prems (i - 1) st |> protect n; fun unrestrict i = conclude #> rotate_prems (1 - i); (*with structural marker*) fun SELECT_GOAL tac i st = if Thm.nprems_of st = 1 andalso i = 1 then tac st else (PRIMITIVE (restrict i 1) THEN tac THEN PRIMITIVE (unrestrict i)) st; (*without structural marker*) fun PREFER_GOAL tac i st = if i < 1 orelse i > Thm.nprems_of st then Seq.empty else (PRIMITIVE (rotate_prems (i - 1)) THEN tac THEN PRIMITIVE (rotate_prems (1 - i))) st; (* multiple goals *) fun precise_conjunction_tac 0 i = eq_assume_tac i | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i | precise_conjunction_tac n i = PRIMITIVE (Drule.with_subgoal i (Conjunction.curry_balanced n)); val adhoc_conjunction_tac = REPEAT_ALL_NEW (SUBGOAL (fn (goal, i) => if can Logic.dest_conjunction goal then resolve0_tac [Conjunction.conjunctionI] i else no_tac)); val conjunction_tac = SUBGOAL (fn (goal, i) => precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE TRY (adhoc_conjunction_tac i)); val recover_conjunction_tac = PRIMITIVE (fn th => Conjunction.uncurry_balanced (Thm.nprems_of th) th); fun PRECISE_CONJUNCTS n tac = SELECT_GOAL (precise_conjunction_tac n 1 THEN tac THEN recover_conjunction_tac); fun CONJUNCTS tac = SELECT_GOAL (conjunction_tac 1 THEN tac THEN recover_conjunction_tac); (* hhf normal form *) fun norm_hhf_tac ctxt = resolve_tac ctxt [Drule.asm_rl] (*cheap approximation -- thanks to builtin Logic.flatten_params*) THEN' SUBGOAL (fn (t, i) => if Drule.is_norm_hhf t then all_tac else rewrite_goal_tac ctxt Drule.norm_hhf_eqs i); (* non-atomic goal assumptions *) fun non_atomic (Const ("Pure.imp", _) $ _ $ _) = true | non_atomic (Const ("Pure.all", _) $ _) = true | non_atomic _ = false; fun assume_rule_tac ctxt = norm_hhf_tac ctxt THEN' CSUBGOAL (fn (goal, i) => let val ((_, goal'), ctxt') = Variable.focus_cterm NONE goal ctxt; val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal'; val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal''); val tacs = Rs |> map (fn R => eresolve_tac ctxt [Raw_Simplifier.norm_hhf ctxt (Thm.trivial R)] THEN_ALL_NEW assume_tac ctxt); in fold_rev (curry op APPEND') tacs (K no_tac) i end); end; structure Basic_Goal: BASIC_GOAL = Goal; open Basic_Goal; diff --git a/src/Pure/library.ML b/src/Pure/library.ML --- a/src/Pure/library.ML +++ b/src/Pure/library.ML @@ -1,1101 +1,1106 @@ (* Title: Pure/library.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Markus Wenzel, TU Muenchen Basic library: functions, pairs, booleans, lists, integers, strings, lists as sets, orders, current directory, misc. See also General/basics.ML for the most fundamental concepts. *) infixr 0 ||| infix 2 ? infix 3 o oo ooo oooo infix 4 ~~ upto downto infix orf andf signature BASIC_LIBRARY = sig (*functions*) val undefined: 'a -> 'b val I: 'a -> 'a val K: 'a -> 'b -> 'a val curry: ('a * 'b -> 'c) -> 'a -> 'b -> 'c val uncurry: ('a -> 'b -> 'c) -> 'a * 'b -> 'c val ? : bool * ('a -> 'a) -> 'a -> 'a val oo: ('a -> 'b) * ('c -> 'd -> 'a) -> 'c -> 'd -> 'b val ooo: ('a -> 'b) * ('c -> 'd -> 'e -> 'a) -> 'c -> 'd -> 'e -> 'b val oooo: ('a -> 'b) * ('c -> 'd -> 'e -> 'f -> 'a) -> 'c -> 'd -> 'e -> 'f -> 'b val funpow: int -> ('a -> 'a) -> 'a -> 'a val funpow_yield: int -> ('a -> 'b * 'a) -> 'a -> 'b list * 'a (*pairs*) val pair: 'a -> 'b -> 'a * 'b val rpair: 'a -> 'b -> 'b * 'a val fst: 'a * 'b -> 'a val snd: 'a * 'b -> 'b val eq_fst: ('a * 'c -> bool) -> ('a * 'b) * ('c * 'd) -> bool val eq_snd: ('b * 'd -> bool) -> ('a * 'b) * ('c * 'd) -> bool val eq_pair: ('a * 'c -> bool) -> ('b * 'd -> bool) -> ('a * 'b) * ('c * 'd) -> bool val swap: 'a * 'b -> 'b * 'a val apfst: ('a -> 'b) -> 'a * 'c -> 'b * 'c val apsnd: ('a -> 'b) -> 'c * 'a -> 'c * 'b val apply2: ('a -> 'b) -> 'a * 'a -> 'b * 'b (*booleans*) val equal: ''a -> ''a -> bool val not_equal: ''a -> ''a -> bool val orf: ('a -> bool) * ('a -> bool) -> 'a -> bool val andf: ('a -> bool) * ('a -> bool) -> 'a -> bool val exists: ('a -> bool) -> 'a list -> bool val forall: ('a -> bool) -> 'a list -> bool (*lists*) + val build: ('a list -> 'a list) -> 'a list + val build_rev: ('a list -> 'a list) -> 'a list val single: 'a -> 'a list val the_single: 'a list -> 'a val singleton: ('a list -> 'b list) -> 'a -> 'b val yield_singleton: ('a list -> 'c -> 'b list * 'c) -> 'a -> 'c -> 'b * 'c val perhaps_apply: ('a -> 'a option) list -> 'a -> 'a option val perhaps_loop: ('a -> 'a option) -> 'a -> 'a option val foldl1: ('a * 'a -> 'a) -> 'a list -> 'a val foldr1: ('a * 'a -> 'a) -> 'a list -> 'a val eq_list: ('a * 'a -> bool) -> 'a list * 'a list -> bool val maps: ('a -> 'b list) -> 'a list -> 'b list val filter: ('a -> bool) -> 'a list -> 'a list val filter_out: ('a -> bool) -> 'a list -> 'a list val map_filter: ('a -> 'b option) -> 'a list -> 'b list val take: int -> 'a list -> 'a list val drop: int -> 'a list -> 'a list val chop: int -> 'a list -> 'a list * 'a list val chop_groups: int -> 'a list -> 'a list list val nth: 'a list -> int -> 'a val nth_list: 'a list list -> int -> 'a list val nth_map: int -> ('a -> 'a) -> 'a list -> 'a list val nth_drop: int -> 'a list -> 'a list val map_index: (int * 'a -> 'b) -> 'a list -> 'b list val fold_index: (int * 'a -> 'b -> 'b) -> 'a list -> 'b -> 'b val map_range: (int -> 'a) -> int -> 'a list val fold_range: (int -> 'a -> 'a) -> int -> 'a -> 'a val split_last: 'a list -> 'a list * 'a val find_first: ('a -> bool) -> 'a list -> 'a option val find_index: ('a -> bool) -> 'a list -> int val get_first: ('a -> 'b option) -> 'a list -> 'b option val get_index: ('a -> 'b option) -> 'a list -> (int * 'b) option val flat: 'a list list -> 'a list val unflat: 'a list list -> 'b list -> 'b list list val grouped: int -> (('a list -> 'b list) -> 'c list list -> 'd list list) -> ('a -> 'b) -> 'c list -> 'd list val burrow: ('a list -> 'b list) -> 'a list list -> 'b list list val burrow_options: ('a list -> 'b list) -> 'a option list -> 'b option list val fold_burrow: ('a list -> 'c -> 'b list * 'd) -> 'a list list -> 'c -> 'b list list * 'd val separate: 'a -> 'a list -> 'a list val surround: 'a -> 'a list -> 'a list val replicate: int -> 'a -> 'a list val map_product: ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list val fold_product: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c val map2: ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list val fold2: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c val map_split: ('a -> 'b * 'c) -> 'a list -> 'b list * 'c list val zip_options: 'a list -> 'b option list -> ('a * 'b) list val ~~ : 'a list * 'b list -> ('a * 'b) list val split_list: ('a * 'b) list -> 'a list * 'b list val burrow_fst: ('a list -> 'b list) -> ('a * 'c) list -> ('b * 'c) list val take_prefix: ('a -> bool) -> 'a list -> 'a list val drop_prefix: ('a -> bool) -> 'a list -> 'a list val chop_prefix: ('a -> bool) -> 'a list -> 'a list * 'a list val take_suffix: ('a -> bool) -> 'a list -> 'a list val drop_suffix: ('a -> bool) -> 'a list -> 'a list val chop_suffix: ('a -> bool) -> 'a list -> 'a list * 'a list val is_prefix: ('a * 'a -> bool) -> 'a list -> 'a list -> bool val chop_common_prefix: ('a * 'b -> bool) -> 'a list * 'b list -> 'a list * ('a list * 'b list) val prefixes1: 'a list -> 'a list list val prefixes: 'a list -> 'a list list val suffixes1: 'a list -> 'a list list val suffixes: 'a list -> 'a list list val trim: ('a -> bool) -> 'a list -> 'a list (*integers*) val upto: int * int -> int list val downto: int * int -> int list val hex_digit: int -> string val radixpand: int * int -> int list val radixstring: int * string * int -> string val string_of_int: int -> string val signed_string_of_int: int -> string val string_of_indexname: string * int -> string val read_radix_int: int -> string list -> int * string list val read_int: string list -> int * string list val oct_char: string -> string (*strings*) val nth_string: string -> int -> string val fold_string: (string -> 'a -> 'a) -> string -> 'a -> 'a val exists_string: (string -> bool) -> string -> bool val forall_string: (string -> bool) -> string -> bool val member_string: string -> string -> bool val first_field: string -> string -> (string * string) option val enclose: string -> string -> string -> string val unenclose: string -> string val quote: string -> string val cartouche: string -> string val space_implode: string -> string list -> string val commas: string list -> string val commas_quote: string list -> string val cat_lines: string list -> string val space_explode: string -> string -> string list val split_lines: string -> string list val plain_words: string -> string val prefix_lines: string -> string -> string val prefix: string -> string -> string val suffix: string -> string -> string val unprefix: string -> string -> string val unsuffix: string -> string -> string val trim_line: string -> string val trim_split_lines: string -> string list val normalize_lines: string -> string val replicate_string: int -> string -> string val translate_string: (string -> string) -> string -> string val encode_lines: string -> string val decode_lines: string -> string val align_right: string -> int -> string -> string val match_string: string -> string -> bool (*reals*) val string_of_real: real -> string val signed_string_of_real: real -> string (*lists as sets -- see also Pure/General/ord_list.ML*) val member: ('b * 'a -> bool) -> 'a list -> 'b -> bool val insert: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list val remove: ('b * 'a -> bool) -> 'b -> 'a list -> 'a list val update: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list val union: ('a * 'a -> bool) -> 'a list -> 'a list -> 'a list val subtract: ('b * 'a -> bool) -> 'b list -> 'a list -> 'a list val inter: ('a * 'b -> bool) -> 'b list -> 'a list -> 'a list val merge: ('a * 'a -> bool) -> 'a list * 'a list -> 'a list val subset: ('a * 'b -> bool) -> 'a list * 'b list -> bool val eq_set: ('a * 'a -> bool) -> 'a list * 'a list -> bool val distinct: ('a * 'a -> bool) -> 'a list -> 'a list val duplicates: ('a * 'a -> bool) -> 'a list -> 'a list val has_duplicates: ('a * 'a -> bool) -> 'a list -> bool val map_transpose: ('a list -> 'b) -> 'a list list -> 'b list (*lists as multisets*) val remove1: ('b * 'a -> bool) -> 'b -> 'a list -> 'a list val combine: ('a * 'a -> bool) -> 'a list -> 'a list -> 'a list val submultiset: ('a * 'b -> bool) -> 'a list * 'b list -> bool (*orders*) type 'a ord = 'a * 'a -> order val is_equal: order -> bool val is_less: order -> bool val is_less_equal: order -> bool val is_greater: order -> bool val is_greater_equal: order -> bool val rev_order: order -> order val make_ord: ('a * 'a -> bool) -> 'a ord val pointer_eq_ord: ('a * 'a -> order) -> 'a * 'a -> order val bool_ord: bool ord val int_ord: int ord val string_ord: string ord val size_ord: string ord val fast_string_ord: string ord val option_ord: ('a * 'b -> order) -> 'a option * 'b option -> order val ||| : ('a -> order) * ('a -> order) -> 'a -> order val prod_ord: ('a * 'b -> order) -> ('c * 'd -> order) -> ('a * 'c) * ('b * 'd) -> order val dict_ord: ('a * 'b -> order) -> 'a list * 'b list -> order val length_ord: 'a list * 'b list -> order val list_ord: ('a * 'b -> order) -> 'a list * 'b list -> order val sort: 'a ord -> 'a list -> 'a list val sort_distinct: 'a ord -> 'a list -> 'a list val sort_strings: string list -> string list val sort_by: ('a -> string) -> 'a list -> 'a list val tag_list: int -> 'a list -> (int * 'a) list val untag_list: (int * 'a) list -> 'a list val order_list: (int * 'a) list -> 'a list (*misc*) val divide_and_conquer: ('a -> 'a list * ('b list -> 'b)) -> 'a -> 'b val divide_and_conquer': ('a -> 'b -> ('a list * ('c list * 'b -> 'c * 'b)) * 'b) -> 'a -> 'b -> 'c * 'b val partition_eq: ('a * 'a -> bool) -> 'a list -> 'a list list val partition_list: (int -> 'a -> bool) -> int -> int -> 'a list -> 'a list list type serial = int val serial: unit -> serial val serial_string: unit -> string eqtype stamp val stamp: unit -> stamp structure Any: sig type T = exn end val getenv: string -> string val getenv_strict: string -> string end; signature LIBRARY = sig include BASIC_LIBRARY val foldl: ('a * 'b -> 'a) -> 'a * 'b list -> 'a val foldr: ('a * 'b -> 'b) -> 'a list * 'b -> 'b end; structure Library: LIBRARY = struct (* functions *) fun undefined _ = raise Match; fun I x = x; fun K x = fn _ => x; fun curry f x y = f (x, y); fun uncurry f (x, y) = f x y; (*conditional application*) fun b ? f = fn x => if b then f x else x; (*composition with multiple args*) fun (f oo g) x y = f (g x y); fun (f ooo g) x y z = f (g x y z); fun (f oooo g) x y z w = f (g x y z w); (*function exponentiation: f (... (f x) ...) with n applications of f*) fun funpow (0: int) _ x = x | funpow n f x = funpow (n - 1) f (f x); fun funpow_yield (0 : int) _ x = ([], x) | funpow_yield n f x = x |> f ||>> funpow_yield (n - 1) f |>> op ::; (* pairs *) fun pair x y = (x, y); fun rpair x y = (y, x); fun fst (x, y) = x; fun snd (x, y) = y; fun eq_fst eq ((x1, _), (x2, _)) = eq (x1, x2); fun eq_snd eq ((_, y1), (_, y2)) = eq (y1, y2); fun eq_pair eqx eqy ((x1, y1), (x2, y2)) = eqx (x1, x2) andalso eqy (y1, y2); fun swap (x, y) = (y, x); fun apfst f (x, y) = (f x, y); fun apsnd f (x, y) = (x, f y); fun apply2 f (x, y) = (f x, f y); (* booleans *) (*polymorphic equality*) fun equal x y = x = y; fun not_equal x y = x <> y; (*combining predicates*) fun p orf q = fn x => p x orelse q x; fun p andf q = fn x => p x andalso q x; val exists = List.exists; val forall = List.all; (** lists **) +fun build (f: 'a list -> 'a list) = f []; +fun build_rev f = build f |> rev; + fun single x = [x]; fun the_single [x] = x | the_single _ = raise List.Empty; fun singleton f x = the_single (f [x]); fun yield_singleton f x = f [x] #>> the_single; fun perhaps_apply funs arg = let fun app [] res = res | app (f :: fs) (changed, x) = (case f x of NONE => app fs (changed, x) | SOME x' => app fs (true, x')); in (case app funs (false, arg) of (false, _) => NONE | (true, arg') => SOME arg') end; fun perhaps_loop f arg = let fun loop (changed, x) = (case f x of NONE => (changed, x) | SOME x' => loop (true, x')); in (case loop (false, arg) of (false, _) => NONE | (true, arg') => SOME arg') end; (* fold -- old versions *) (*the following versions of fold are designed to fit nicely with infixes*) (* (op @) (e, [x1, ..., xn]) ===> ((e @ x1) @ x2) ... @ xn for operators that associate to the left (TAIL RECURSIVE)*) fun foldl (f: 'a * 'b -> 'a) : 'a * 'b list -> 'a = let fun itl (e, []) = e | itl (e, a::l) = itl (f(e, a), l) in itl end; (* (op @) ([x1, ..., xn], e) ===> x1 @ (x2 ... @ (xn @ e)) for operators that associate to the right (not tail recursive)*) fun foldr f (l, e) = let fun itr [] = e | itr (a::l) = f(a, itr l) in itr l end; (* (op @) [x1, ..., xn] ===> ((x1 @ x2) @ x3) ... @ xn for operators that associate to the left (TAIL RECURSIVE)*) fun foldl1 f [] = raise List.Empty | foldl1 f (x :: xs) = foldl f (x, xs); (* (op @) [x1, ..., xn] ===> x1 @ (x2 ... @ (x[n-1] @ xn)) for n > 0, operators that associate to the right (not tail recursive)*) fun foldr1 f [] = raise List.Empty | foldr1 f l = let fun itr [x] = x | itr (x::l) = f(x, itr l) in itr l end; (* basic list functions *) fun eq_list eq (list1, list2) = pointer_eq (list1, list2) orelse let fun eq_lst (x :: xs, y :: ys) = eq (x, y) andalso eq_lst (xs, ys) | eq_lst _ = true; in length list1 = length list2 andalso eq_lst (list1, list2) end; fun maps f [] = [] | maps f (x :: xs) = f x @ maps f xs; val filter = List.filter; fun filter_out f = filter (not o f); val map_filter = List.mapPartial; fun take (0: int) xs = [] | take _ [] = [] | take n (x :: xs) = x :: take (n - 1) xs; fun drop (0: int) xs = xs | drop _ [] = [] | drop n (x :: xs) = drop (n - 1) xs; fun chop (0: int) xs = ([], xs) | chop _ [] = ([], []) | chop n (x :: xs) = chop (n - 1) xs |>> cons x; fun chop_groups n list = (case chop (Int.max (n, 1)) list of ([], _) => [] | (g, rest) => g :: chop_groups n rest); (*return nth element of a list, where 0 designates the first element; raise Subscript if list too short*) fun nth xs i = List.nth (xs, i); fun nth_list xss i = nth xss i handle General.Subscript => []; fun nth_map 0 f (x :: xs) = f x :: xs | nth_map n f (x :: xs) = x :: nth_map (n - 1) f xs | nth_map (_: int) _ [] = raise Subscript; fun nth_drop n xs = List.take (xs, n) @ List.drop (xs, n + 1); fun map_index f = let fun map_aux (_: int) [] = [] | map_aux i (x :: xs) = f (i, x) :: map_aux (i + 1) xs in map_aux 0 end; fun fold_index f = let fun fold_aux (_: int) [] y = y | fold_aux i (x :: xs) y = fold_aux (i + 1) xs (f (i, x) y) in fold_aux 0 end; fun map_range f i = let fun map_aux (k: int) = if k < i then f k :: map_aux (k + 1) else [] in map_aux 0 end; fun fold_range f i = let fun fold_aux (k: int) y = if k < i then fold_aux (k + 1) (f k y) else y in fold_aux 0 end; (*rear decomposition*) fun split_last [] = raise List.Empty | split_last [x] = ([], x) | split_last (x :: xs) = apfst (cons x) (split_last xs); (*find first element satisfying predicate*) val find_first = List.find; (*find position of first element satisfying a predicate*) fun find_index pred = let fun find (_: int) [] = ~1 | find n (x :: xs) = if pred x then n else find (n + 1) xs; in find 0 end; (*get first element by lookup function*) fun get_first _ [] = NONE | get_first f (x :: xs) = (case f x of NONE => get_first f xs | some => some); fun get_index f = let fun get_aux (_: int) [] = NONE | get_aux i (x :: xs) = (case f x of NONE => get_aux (i + 1) xs | SOME y => SOME (i, y)) in get_aux 0 end; val flat = List.concat; fun unflat (xs :: xss) ys = let val (ps, qs) = chop (length xs) ys in ps :: unflat xss qs end | unflat [] [] = [] | unflat _ _ = raise ListPair.UnequalLengths; fun grouped n comb f = chop_groups n #> comb (map f) #> flat; fun burrow f xss = unflat xss (f (flat xss)); fun burrow_options f os = map (try hd) (burrow f (map the_list os)); fun fold_burrow f xss s = apfst (unflat xss) (f (flat xss) s); (*separate s [x1, x2, ..., xn] ===> [x1, s, x2, s, ..., s, xn]*) fun separate s (x :: (xs as _ :: _)) = x :: s :: separate s xs | separate _ xs = xs; fun surround s (x :: xs) = s :: x :: surround s xs | surround s [] = [s]; (*make the list [x, x, ..., x] of length n*) fun replicate (n: int) x = let fun rep (0, xs) = xs | rep (n, xs) = rep (n - 1, x :: xs) in if n < 0 then raise Subscript else rep (n, []) end; (* direct product *) fun map_product f _ [] = [] | map_product f [] _ = [] | map_product f (x :: xs) ys = map (f x) ys @ map_product f xs ys; fun fold_product f _ [] z = z | fold_product f [] _ z = z | fold_product f (x :: xs) ys z = z |> fold (f x) ys |> fold_product f xs ys; (* lists of pairs *) fun map2 _ [] [] = [] | map2 f (x :: xs) (y :: ys) = f x y :: map2 f xs ys | map2 _ _ _ = raise ListPair.UnequalLengths; fun fold2 _ [] [] z = z | fold2 f (x :: xs) (y :: ys) z = fold2 f xs ys (f x y z) | fold2 _ _ _ _ = raise ListPair.UnequalLengths; fun map_split _ [] = ([], []) | map_split f (x :: xs) = let val (y, w) = f x; val (ys, ws) = map_split f xs; in (y :: ys, w :: ws) end; fun zip_options (x :: xs) (SOME y :: ys) = (x, y) :: zip_options xs ys | zip_options (_ :: xs) (NONE :: ys) = zip_options xs ys | zip_options _ [] = [] | zip_options [] _ = raise ListPair.UnequalLengths; (*combine two lists forming a list of pairs: [x1, ..., xn] ~~ [y1, ..., yn] ===> [(x1, y1), ..., (xn, yn)]*) fun [] ~~ [] = [] | (x :: xs) ~~ (y :: ys) = (x, y) :: (xs ~~ ys) | _ ~~ _ = raise ListPair.UnequalLengths; (*inverse of ~~; the old 'split': [(x1, y1), ..., (xn, yn)] ===> ([x1, ..., xn], [y1, ..., yn])*) val split_list = ListPair.unzip; fun burrow_fst f xs = split_list xs |>> f |> op ~~; (* take, drop, chop, trim according to predicate *) fun take_prefix pred list = let fun take res (x :: xs) = if pred x then take (x :: res) xs else rev res | take res [] = rev res; in take [] list end; fun drop_prefix pred list = let fun drop (x :: xs) = if pred x then drop xs else x :: xs | drop [] = []; in drop list end; fun chop_prefix pred list = let val prfx = take_prefix pred list; val sffx = drop (length prfx) list; in (prfx, sffx) end; fun take_suffix pred list = let fun take res (x :: xs) = if pred x then take (x :: res) xs else res | take res [] = res; in take [] (rev list) end; fun drop_suffix pred list = let fun drop (x :: xs) = if pred x then drop xs else rev (x :: xs) | drop [] = []; in drop (rev list) end; fun chop_suffix pred list = let val prfx = drop_suffix pred list; val sffx = drop (length prfx) list; in (prfx, sffx) end; fun trim pred = drop_prefix pred #> drop_suffix pred; (* prefixes, suffixes *) fun is_prefix _ [] _ = true | is_prefix eq (x :: xs) (y :: ys) = eq (x, y) andalso is_prefix eq xs ys | is_prefix eq _ _ = false; fun chop_common_prefix eq ([], ys) = ([], ([], ys)) | chop_common_prefix eq (xs, []) = ([], (xs, [])) | chop_common_prefix eq (xs as x :: xs', ys as y :: ys') = if eq (x, y) then let val (ps', xys'') = chop_common_prefix eq (xs', ys') in (x :: ps', xys'') end else ([], (xs, ys)); fun prefixes1 [] = [] | prefixes1 (x :: xs) = map (cons x) ([] :: prefixes1 xs); fun prefixes xs = [] :: prefixes1 xs; fun suffixes1 xs = map rev (prefixes1 (rev xs)); fun suffixes xs = [] :: suffixes1 xs; (** integers **) (* lists of integers *) (*make the list [from, from + 1, ..., to]*) fun ((i: int) upto j) = if i > j then [] else i :: (i + 1 upto j); (*make the list [from, from - 1, ..., to]*) fun ((i: int) downto j) = if i < j then [] else i :: (i - 1 downto j); (* convert integers to strings *) (*hexadecimal*) fun hex_digit i = if i < 10 then chr (Char.ord #"0" + i) else chr (Char.ord #"a" + i - 10); (*expand the number in the given base; example: radixpand (2, 8) gives [1, 0, 0, 0]*) fun radixpand (base, num) : int list = let fun radix (n, tail) = if n < base then n :: tail else radix (n div base, (n mod base) :: tail) in radix (num, []) end; (*expands a number into a string of characters starting from "zerochar"; example: radixstring (2, "0", 8) gives "1000"*) fun radixstring (base, zerochar, num) = let val offset = ord zerochar; fun chrof n = chr (offset + n) in implode (map chrof (radixpand (base, num))) end; local val zero = Char.ord #"0"; val small_int = 10000: int; val small_int_table = Vector.tabulate (small_int, Int.toString); in fun string_of_int i = if i < 0 then Int.toString i else if i < 10 then chr (zero + i) else if i < small_int then Vector.sub (small_int_table, i) else Int.toString i; end; fun signed_string_of_int i = if i < 0 then "-" ^ string_of_int (~ i) else string_of_int i; fun string_of_indexname (a, 0) = a | string_of_indexname (a, i) = a ^ "_" ^ string_of_int i; (* read integers *) fun read_radix_int radix cs = let val zero = Char.ord #"0"; val limit = zero + radix; fun scan (num, []) = (num, []) | scan (num, c :: cs) = if zero <= ord c andalso ord c < limit then scan (radix * num + (ord c - zero), cs) else (num, c :: cs); in scan (0, cs) end; val read_int = read_radix_int 10; fun oct_char s = chr (#1 (read_radix_int 8 (raw_explode s))); (** strings **) (* functions tuned for strings, avoiding explode *) fun nth_string str i = (case try String.substring (str, i, 1) of SOME s => s | NONE => raise Subscript); fun fold_string f str x0 = let val n = size str; fun iter (x, i) = if i < n then iter (f (String.substring (str, i, 1)) x, i + 1) else x; in iter (x0, 0) end; fun exists_string pred str = let val n = size str; fun ex i = i < n andalso (pred (String.substring (str, i, 1)) orelse ex (i + 1)); in ex 0 end; fun forall_string pred = not o exists_string (not o pred); fun member_string str s = exists_string (fn s' => s = s') str; fun first_field sep str = let val n = size sep; val len = size str; fun find i = if i + n > len then NONE else if String.substring (str, i, n) = sep then SOME i else find (i + 1); in (case find 0 of NONE => NONE | SOME i => SOME (String.substring (str, 0, i), String.extract (str, i + n, NONE))) end; (*enclose in brackets*) fun enclose lpar rpar str = lpar ^ str ^ rpar; fun unenclose str = String.substring (str, 1, size str - 2); (*simple quoting (does not escape special chars)*) val quote = enclose "\"" "\""; val cartouche = enclose "\" "\"; val space_implode = String.concatWith; val commas = space_implode ", "; val commas_quote = commas o map quote; val cat_lines = space_implode "\n"; (*space_explode "." "h.e..l.lo" = ["h", "e", "", "l", "lo"]*) fun space_explode _ "" = [] | space_explode sep s = String.fields (fn c => str c = sep) s; val split_lines = space_explode "\n"; fun plain_words s = space_explode "_" s |> space_implode " "; fun prefix_lines "" txt = txt | prefix_lines prfx txt = txt |> split_lines |> map (fn s => prfx ^ s) |> cat_lines; fun prefix prfx s = prfx ^ s; fun suffix sffx s = s ^ sffx; fun unprefix prfx s = if String.isPrefix prfx s then String.substring (s, size prfx, size s - size prfx) else raise Fail "unprefix"; fun unsuffix sffx s = if String.isSuffix sffx s then String.substring (s, 0, size s - size sffx) else raise Fail "unsuffix"; fun trim_line s = if String.isSuffix "\r\n" s then String.substring (s, 0, size s - 2) else if String.isSuffix "\r" s orelse String.isSuffix "\n" s then String.substring (s, 0, size s - 1) else s; val trim_split_lines = trim_line #> split_lines #> map trim_line; fun normalize_lines str = if exists_string (fn s => s = "\r") str then split_lines str |> map trim_line |> cat_lines else str; fun replicate_string (0: int) _ = "" | replicate_string 1 a = a | replicate_string k a = if k mod 2 = 0 then replicate_string (k div 2) (a ^ a) else replicate_string (k div 2) (a ^ a) ^ a; fun translate_string f = String.translate (f o String.str); val encode_lines = translate_string (fn "\n" => "\v" | c => c); val decode_lines = translate_string (fn "\v" => "\n" | c => c); fun align_right c k s = let val _ = if size c <> 1 orelse size s > k then raise Fail "align_right" else () in replicate_string (k - size s) c ^ s end; (*crude matching of str against simple glob pat*) fun match_string pat str = let fun match [] _ = true | match (p :: ps) s = size p <= size s andalso (case try (unprefix p) s of SOME s' => match ps s' | NONE => match (p :: ps) (String.substring (s, 1, size s - 1))); in match (space_explode "*" pat) str end; (** reals **) val string_of_real = Real.fmt (StringCvt.GEN NONE); fun signed_string_of_real x = if x < 0.0 then "-" ^ string_of_real (~ x) else string_of_real x; (** lists as sets -- see also Pure/General/ord_list.ML **) (* canonical operations *) fun member eq list x = let fun memb [] = false | memb (y :: ys) = eq (x, y) orelse memb ys; in memb list end; fun insert eq x xs = if member eq xs x then xs else x :: xs; fun remove eq x xs = if member eq xs x then filter_out (fn y => eq (x, y)) xs else xs; fun update eq x xs = cons x (remove eq x xs); fun inter eq xs = filter (member eq xs); fun union eq = fold (insert eq); fun subtract eq = fold (remove eq); fun merge eq (xs, ys) = if pointer_eq (xs, ys) then xs else if null xs then ys else fold_rev (insert eq) ys xs; (* subset and set equality *) fun subset eq (xs, ys) = forall (member eq ys) xs; fun eq_set eq (xs, ys) = eq_list eq (xs, ys) orelse (subset eq (xs, ys) andalso subset (eq o swap) (ys, xs)); (*makes a list of the distinct members of the input; preserves order, takes first of equal elements*) fun distinct eq lst = let fun dist (rev_seen, []) = rev rev_seen | dist (rev_seen, x :: xs) = if member eq rev_seen x then dist (rev_seen, xs) else dist (x :: rev_seen, xs); in dist ([], lst) end; (*returns a list containing all repeated elements exactly once; preserves order, takes first of equal elements*) fun duplicates eq lst = let fun dups (rev_dups, []) = rev rev_dups | dups (rev_dups, x :: xs) = if member eq rev_dups x orelse not (member eq xs x) then dups (rev_dups, xs) else dups (x :: rev_dups, xs); in dups ([], lst) end; fun has_duplicates eq = let fun dups [] = false | dups (x :: xs) = member eq xs x orelse dups xs; in dups end; (* matrices *) fun map_transpose f xss = let val n = (case distinct (op =) (map length xss) of [] => 0 | [n] => n | _ => raise ListPair.UnequalLengths); in map_range (fn m => f (map (fn xs => nth xs m) xss)) n end; (** lists as multisets **) fun remove1 eq x [] = [] | remove1 eq x (y :: ys) = if eq (x, y) then ys else y :: remove1 eq x ys; fun combine eq xs ys = fold (remove1 eq) ys xs @ ys; fun submultiset _ ([], _) = true | submultiset eq (x :: xs, ys) = member eq ys x andalso submultiset eq (xs, remove1 eq x ys); (** orders **) type 'a ord = 'a * 'a -> order; fun is_equal ord = ord = EQUAL; fun is_less ord = ord = LESS; fun is_less_equal ord = ord = LESS orelse ord = EQUAL; fun is_greater ord = ord = GREATER; fun is_greater_equal ord = ord = GREATER orelse ord = EQUAL; fun rev_order LESS = GREATER | rev_order EQUAL = EQUAL | rev_order GREATER = LESS; (*compose orders*) fun (a_ord ||| b_ord) p = (case a_ord p of EQUAL => b_ord p | ord => ord); (*assume rel is a linear strict order*) fun make_ord rel (x, y) = if rel (x, y) then LESS else if rel (y, x) then GREATER else EQUAL; fun pointer_eq_ord ord (x, y) = if pointer_eq (x, y) then EQUAL else ord (x, y); fun bool_ord (false, true) = LESS | bool_ord (true, false) = GREATER | bool_ord _ = EQUAL; val int_ord = Int.compare; val string_ord = String.compare; val size_ord = int_ord o apply2 size; val fast_string_ord = pointer_eq_ord (size_ord ||| string_ord); fun option_ord ord (SOME x, SOME y) = ord (x, y) | option_ord _ (NONE, NONE) = EQUAL | option_ord _ (NONE, SOME _) = LESS | option_ord _ (SOME _, NONE) = GREATER; (*lexicographic product*) fun prod_ord a_ord b_ord ((x, y), (x', y')) = (case a_ord (x, x') of EQUAL => b_ord (y, y') | ord => ord); (*dictionary order -- in general NOT well-founded!*) fun dict_ord elem_ord (x :: xs, y :: ys) = (case elem_ord (x, y) of EQUAL => dict_ord elem_ord (xs, ys) | ord => ord) | dict_ord _ ([], []) = EQUAL | dict_ord _ ([], _ :: _) = LESS | dict_ord _ (_ :: _, []) = GREATER; (*lexicographic product of lists*) fun length_ord (xs, ys) = int_ord (length xs, length ys); fun list_ord elem_ord = length_ord ||| dict_ord elem_ord; (* sorting *) (*stable mergesort -- preserves order of equal elements*) fun mergesort unique ord = let fun merge (xs as x :: xs') (ys as y :: ys') = (case ord (x, y) of LESS => x :: merge xs' ys | EQUAL => if unique then merge xs ys' else x :: merge xs' ys | GREATER => y :: merge xs ys') | merge [] ys = ys | merge xs [] = xs; fun merge_all [xs] = xs | merge_all xss = merge_all (merge_pairs xss) and merge_pairs (xs :: ys :: xss) = merge xs ys :: merge_pairs xss | merge_pairs xss = xss; fun runs (x :: y :: xs) = (case ord (x, y) of LESS => ascending y [x] xs | EQUAL => if unique then runs (x :: xs) else ascending y [x] xs | GREATER => descending y [x] xs) | runs xs = [xs] and ascending x xs (zs as y :: ys) = (case ord (x, y) of LESS => ascending y (x :: xs) ys | EQUAL => if unique then ascending x xs ys else ascending y (x :: xs) ys | GREATER => rev (x :: xs) :: runs zs) | ascending x xs [] = [rev (x :: xs)] and descending x xs (zs as y :: ys) = (case ord (x, y) of GREATER => descending y (x :: xs) ys | EQUAL => if unique then descending x xs ys else (x :: xs) :: runs zs | LESS => (x :: xs) :: runs zs) | descending x xs [] = [x :: xs]; in merge_all o runs end; fun sort ord = mergesort false ord; fun sort_distinct ord = mergesort true ord; val sort_strings = sort string_ord; fun sort_by key xs = sort (string_ord o apply2 key) xs; (* items tagged by integer index *) (*insert tags*) fun tag_list k [] = [] | tag_list k (x :: xs) = (k:int, x) :: tag_list (k + 1) xs; (*remove tags and suppress duplicates -- list is assumed sorted!*) fun untag_list [] = [] | untag_list [(k: int, x)] = [x] | untag_list ((k, x) :: (rest as (k', x') :: _)) = if k = k' then untag_list rest else x :: untag_list rest; (*return list elements in original order*) fun order_list list = untag_list (sort (int_ord o apply2 fst) list); (** misc **) fun divide_and_conquer decomp x = let val (ys, recomb) = decomp x in recomb (map (divide_and_conquer decomp) ys) end; fun divide_and_conquer' decomp x s = let val ((ys, recomb), s') = decomp x s in recomb (fold_map (divide_and_conquer' decomp) ys s') end; (*Partition a list into buckets [ bi, b(i+1), ..., bj ] putting x in bk if p(k)(x) holds. Preserve order of elements if possible.*) fun partition_list p i j = let fun part (k: int) xs = if k > j then (case xs of [] => [] | _ => raise Fail "partition_list") else let val (ns, rest) = List.partition (p k) xs in ns :: part (k + 1) rest end; in part (i: int) end; fun partition_eq (eq: 'a * 'a -> bool) = let fun part [] = [] | part (x :: ys) = let val (xs, xs') = List.partition (fn y => eq (x, y)) ys in (x :: xs) :: part xs' end; in part end; (* serial numbers and abstract stamps *) type serial = int; val serial = Counter.make (); val serial_string = string_of_int o serial; datatype stamp = Stamp of serial; fun stamp () = Stamp (serial ()); (* values of any type *) (*note that the builtin exception datatype may be extended by new constructors at any time*) structure Any = struct type T = exn end; (* getenv *) fun getenv x = (case OS.Process.getEnv x of NONE => "" | SOME y => y); fun getenv_strict x = (case getenv x of "" => error ("Undefined Isabelle environment variable: " ^ quote x) | y => y); end; structure Basic_Library: BASIC_LIBRARY = Library; open Basic_Library; diff --git a/src/Pure/more_thm.ML b/src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML +++ b/src/Pure/more_thm.ML @@ -1,746 +1,783 @@ (* 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 add_tvars: thm -> ctyp Term_Subst.TVars.table -> ctyp Term_Subst.TVars.table + val add_vars: thm -> cterm Term_Subst.Vars.table -> cterm Term_Subst.Vars.table 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 item_net: thm Item_Net.T val item_net_intro: thm Item_Net.T val item_net_elim: 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 forall_intr_vars: 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 register_proofs: thm list lazy -> theory -> theory val consolidate_theory: theory -> unit val expose_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); +val add_tvars = + Thm.fold_atomic_ctyps {hyps = false} Term.is_TVar (fn cT => fn tab => + let val v = Term.dest_TVar (Thm.typ_of cT) + in tab |> not (Term_Subst.TVars.defined tab v) ? Term_Subst.TVars.update (v, cT) end); + +val add_vars = + Thm.fold_atomic_cterms {hyps = false} Term.is_Var (fn ct => fn tab => + let val v = Term.dest_Var (Thm.term_of ct) + in tab |> not (Term_Subst.Vars.defined tab v) ? Term_Subst.Vars.update (v, ct) end); (* 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 item_net = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); val item_net_intro = Item_Net.init eq_thm_prop (single o Thm.concl_of); val item_net_elim = 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) + (Thm.fold_terms {hyps = false} 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 tfrees = fold (Symtab.insert_set o #1 o #1) instT Symtab.empty; - val frees = fold (Symtab.insert_set o #1 o #1) inst Symtab.empty; + val tfrees = Symtab.build (fold (Symtab.insert_set o #1 o #1) instT); + val frees = Symtab.build (fold (Symtab.insert_set o #1 o #1) inst); val hyps = Thm.chyps_of th; val inst_cterm = Thm.generalize_cterm (tfrees, frees) idx #> Thm.instantiate_cterm insts; in th |> fold_rev Thm.implies_intr hyps |> Thm.generalize (tfrees, 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; + val instT = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_tvars thm)) cTs; + val thm' = Thm.instantiate (instT, []) thm; + val inst = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_vars thm')) cts; + in Thm.instantiate ([], inst) thm' end; -(* forall_intr_frees: generalization over all suitable Free variables *) +(* implicit generalization over variables -- canonical order *) + +fun forall_intr_vars th = + let + val (_, vars) = + (th, (Term_Subst.Vars.empty, [])) |-> Thm.fold_atomic_cterms {hyps = false} Term.is_Var + (fn ct => fn (seen, vars) => + let val v = Term.dest_Var (Thm.term_of ct) in + if not (Term_Subst.Vars.defined seen v) + then (Term_Subst.Vars.add (v, ()) seen, ct :: vars) + else (seen, vars) + end); + in fold Thm.forall_intr vars th end; 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 []; + Term_Subst.Frees.build + (fold Term_Subst.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th)) #> + fold Term_Subst.add_frees (Thm.hyps_of th)); + val (_, frees) = + (th, (fixed, [])) |-> Thm.fold_atomic_cterms {hyps = false} Term.is_Free + (fn ct => fn (seen, frees) => + let val v = Term.dest_Free (Thm.term_of ct) in + if not (Term_Subst.Frees.defined seen v) + then (Term_Subst.Frees.add (v, ()) seen, ct :: frees) + else (seen, frees) + end); 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 instantiateT = Term_Subst.instantiateT (Term_Subst.TVars.table instT); - val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) => - let val T' = instantiateT 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; + val cert = Thm.global_cterm_of thy; + val certT = Thm.global_ctyp_of thy; + + val instT = + Term_Subst.TVars.build (prop |> (Term.fold_types o Term.fold_atyps) + (fn T => fn instT => + (case T of + TVar (v as ((a, _), S)) => + if Term_Subst.TVars.defined instT v then instT + else Term_Subst.TVars.update (v, TFree (a, S)) instT + | _ => instT))); + val cinstT = Term_Subst.TVars.map (K certT) instT; + val cinst = + Term_Subst.Vars.build (prop |> Term.fold_aterms + (fn t => fn inst => + (case t of + Var ((x, i), T) => + let val T' = Term_Subst.instantiateT instT T + in Term_Subst.Vars.update (((x, i), T'), cert (Free ((x, T')))) inst end + | _ => inst))); + in Thm.instantiate (Term_Subst.TVars.dest cinstT, Term_Subst.Vars.dest cinst) 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 = build_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)); (** forked proofs **) structure Proofs = Theory_Data ( type T = thm list lazy Inttab.table; val empty = Inttab.empty; val extend = I; val merge = Inttab.merge (K true); ); fun reset_proofs thy = if Inttab.is_empty (Proofs.get thy) then NONE else SOME (Proofs.put Inttab.empty thy); val _ = Theory.setup (Theory.at_begin reset_proofs); fun register_proofs ths thy = let val entry = (serial (), Lazy.map_finished (map Thm.trim_context) ths) in (Proofs.map o Inttab.update) entry thy end; fun force_proofs thy = Proofs.get thy |> Inttab.dest |> maps (map (Thm.transfer thy) o Lazy.force o #2); val consolidate_theory = Thm.consolidate o force_proofs; fun expose_theory thy = if Proofterm.export_enabled () then Thm.expose_proofs thy (force_proofs thy) else (); (** 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/morphism.ML b/src/Pure/morphism.ML --- a/src/Pure/morphism.ML +++ b/src/Pure/morphism.ML @@ -1,178 +1,178 @@ (* Title: Pure/morphism.ML Author: Makarius Abstract morphisms on formal entities. *) infix 1 $> signature BASIC_MORPHISM = sig type morphism type declaration = morphism -> Context.generic -> Context.generic val $> : morphism * morphism -> morphism end signature MORPHISM = sig include BASIC_MORPHISM exception MORPHISM of string * exn val morphism: string -> {binding: (binding -> binding) list, typ: (typ -> typ) list, term: (term -> term) list, fact: (thm list -> thm list) list} -> morphism val pretty: morphism -> Pretty.T val binding: morphism -> binding -> binding val binding_prefix: morphism -> (string * bool) list val typ: morphism -> typ -> typ val term: morphism -> term -> term val fact: morphism -> thm list -> thm list val thm: morphism -> thm -> thm val cterm: morphism -> cterm -> cterm val identity: morphism val compose: morphism -> morphism -> morphism val transform: morphism -> (morphism -> 'a) -> morphism -> 'a val form: (morphism -> 'a) -> 'a val binding_morphism: string -> (binding -> binding) -> morphism val typ_morphism: string -> (typ -> typ) -> morphism val term_morphism: string -> (term -> term) -> morphism val fact_morphism: string -> (thm list -> thm list) -> morphism val thm_morphism: string -> (thm -> thm) -> morphism val transfer_morphism: theory -> morphism val transfer_morphism': Proof.context -> morphism val transfer_morphism'': Context.generic -> morphism val trim_context_morphism: morphism val instantiate_frees_morphism: ((string * sort) * ctyp) list * ((string * typ) * cterm) list -> morphism val instantiate_morphism: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> morphism end; structure Morphism: MORPHISM = struct (* named functions *) type 'a funs = (string * ('a -> 'a)) list; exception MORPHISM of string * exn; fun app (name, f) x = f x handle exn => if Exn.is_interrupt exn then Exn.reraise exn else raise MORPHISM (name, exn); fun apply fs = fold_rev app fs; (* type morphism *) datatype morphism = Morphism of {names: string list, binding: binding funs, typ: typ funs, term: term funs, fact: thm list funs}; type declaration = morphism -> Context.generic -> Context.generic; fun morphism a {binding, typ, term, fact} = Morphism { names = if a = "" then [] else [a], binding = map (pair a) binding, typ = map (pair a) typ, term = map (pair a) term, fact = map (pair a) fact}; fun pretty (Morphism {names, ...}) = Pretty.enum ";" "{" "}" (map Pretty.str (rev names)); val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty); fun binding (Morphism {binding, ...}) = apply binding; fun binding_prefix morph = Binding.name "x" |> binding morph |> Binding.prefix_of; fun typ (Morphism {typ, ...}) = apply typ; fun term (Morphism {term, ...}) = apply term; fun fact (Morphism {fact, ...}) = apply fact; val thm = singleton o fact; val cterm = Drule.cterm_rule o thm; (* morphism combinators *) val identity = morphism "" {binding = [], typ = [], term = [], fact = []}; fun compose (Morphism {names = names1, binding = binding1, typ = typ1, term = term1, fact = fact1}) (Morphism {names = names2, binding = binding2, typ = typ2, term = term2, fact = fact2}) = Morphism { names = names1 @ names2, binding = binding1 @ binding2, typ = typ1 @ typ2, term = term1 @ term2, fact = fact1 @ fact2}; fun phi1 $> phi2 = compose phi2 phi1; fun transform phi f = fn psi => f (phi $> psi); fun form f = f identity; (* concrete morphisms *) fun binding_morphism a binding = morphism a {binding = [binding], typ = [], term = [], fact = []}; fun typ_morphism a typ = morphism a {binding = [], typ = [typ], term = [], fact = []}; fun term_morphism a term = morphism a {binding = [], typ = [], term = [term], fact = []}; fun fact_morphism a fact = morphism a {binding = [], typ = [], term = [], fact = [fact]}; fun thm_morphism a thm = morphism a {binding = [], typ = [], term = [], fact = [map thm]}; val transfer_morphism = thm_morphism "transfer" o Thm.join_transfer; val transfer_morphism' = transfer_morphism o Proof_Context.theory_of; val transfer_morphism'' = transfer_morphism o Context.theory_of; val trim_context_morphism = thm_morphism "trim_context" Thm.trim_context; (* instantiate *) fun instantiate_frees_morphism ([], []) = identity | instantiate_frees_morphism (cinstT, cinst) = let val instT = - fold (fn (v, cT) => Term_Subst.TFrees.add (v, Thm.typ_of cT)) - cinstT Term_Subst.TFrees.empty; + Term_Subst.TFrees.build + (cinstT |> fold (fn (v, cT) => Term_Subst.TFrees.add (v, Thm.typ_of cT))); val inst = - fold (fn (v, ct) => Term_Subst.Frees.add (v, Thm.term_of ct)) - cinst Term_Subst.Frees.empty; + Term_Subst.Frees.build + (cinst |> fold (fn (v, ct) => Term_Subst.Frees.add (v, Thm.term_of ct))); in morphism "instantiate_frees" {binding = [], typ = if Term_Subst.TFrees.is_empty instT then [] else [Term_Subst.instantiateT_frees instT], term = [Term_Subst.instantiate_frees (instT, inst)], fact = [map (Thm.instantiate_frees (cinstT, cinst))]} end; fun instantiate_morphism ([], []) = identity | instantiate_morphism (cinstT, cinst) = let val instT = - fold (fn (v, cT) => Term_Subst.TVars.add (v, Thm.typ_of cT)) - cinstT Term_Subst.TVars.empty; + Term_Subst.TVars.build (cinstT |> fold (fn (v, cT) => + Term_Subst.TVars.add (v, Thm.typ_of cT))); val inst = - fold (fn (v, ct) => Term_Subst.Vars.add (v, Thm.term_of ct)) - cinst Term_Subst.Vars.empty; + Term_Subst.Vars.build (cinst |> fold (fn (v, ct) => + Term_Subst.Vars.add (v, Thm.term_of ct))); in morphism "instantiate" {binding = [], typ = if Term_Subst.TVars.is_empty instT then [] else [Term_Subst.instantiateT instT], term = [Term_Subst.instantiate (instT, inst)], fact = [map (Thm.instantiate (cinstT, cinst))]} end; end; structure Basic_Morphism: BASIC_MORPHISM = Morphism; open Basic_Morphism; diff --git a/src/Pure/primitive_defs.ML b/src/Pure/primitive_defs.ML --- a/src/Pure/primitive_defs.ML +++ b/src/Pure/primitive_defs.ML @@ -1,85 +1,85 @@ (* Title: Pure/primitive_defs.ML Author: Makarius Primitive definition forms. *) signature PRIMITIVE_DEFS = sig val dest_def: Proof.context -> {check_head: term -> bool, check_free_lhs: string -> bool, check_free_rhs: string -> bool, check_tfree: string -> bool} -> term -> (term * term) * term list * term val abs_def: term -> term * term end; structure Primitive_Defs: PRIMITIVE_DEFS = struct fun term_kind (Const _) = "existing constant " | term_kind (Free _) = "free variable " | term_kind (Bound _) = "bound variable " | term_kind _ = ""; (*c x \ t[x] to \x. c x \ t[x]*) fun dest_def ctxt {check_head, check_free_lhs, check_free_rhs, check_tfree} eq = let fun err msg = raise TERM (msg, [eq]); val eq_vars = Term.strip_all_vars eq; val eq_body = Term.strip_all_body eq; val display_terms = commas_quote o map (Syntax.string_of_term ctxt o Syntax_Trans.bound_vars eq_vars); val display_types = commas_quote o map (Syntax.string_of_typ ctxt); val (raw_lhs, rhs) = Logic.dest_equals eq_body handle TERM _ => err "Not a meta-equality (\)"; val lhs = Envir.beta_eta_contract raw_lhs; val (head, args) = Term.strip_comb lhs; - val head_tfrees = Term.add_tfrees head []; + val head_tfrees = Term_Subst.TFrees.build (Term_Subst.add_tfrees head); fun check_arg (Bound _) = true | check_arg (Free (x, _)) = check_free_lhs x | check_arg (Const ("Pure.type", Type ("itself", [TFree _]))) = true | check_arg _ = false; fun close_arg (Bound _) t = t | close_arg x t = Logic.all x t; val lhs_bads = filter_out check_arg args; val lhs_dups = duplicates (op aconv) args; val rhs_extras = Term.fold_aterms (fn v as Free (x, _) => if check_free_rhs x orelse member (op aconv) args v then I else insert (op aconv) v | _ => I) rhs []; val rhs_extrasT = Term.fold_aterms (Term.fold_types (fn v as TFree (a, S) => - if check_tfree a orelse member (op =) head_tfrees (a, S) then I + if check_tfree a orelse Term_Subst.TFrees.defined head_tfrees (a, S) then I else insert (op =) v | _ => I)) rhs []; in if not (check_head head) then err ("Bad head of lhs: " ^ term_kind head ^ display_terms [head]) else if not (null lhs_bads) then err ("Bad arguments on lhs: " ^ display_terms lhs_bads) else if not (null lhs_dups) then err ("Duplicate arguments on lhs: " ^ display_terms lhs_dups) else if not (null rhs_extras) then err ("Extra variables on rhs: " ^ display_terms rhs_extras) else if not (null rhs_extrasT) then err ("Extra type variables on rhs: " ^ display_types rhs_extrasT) else if exists_subterm (fn t => t aconv head) rhs then err "Entity to be defined occurs on rhs" else ((lhs, rhs), args, fold_rev close_arg args (Logic.list_all (eq_vars, (Logic.mk_equals (lhs, rhs))))) end; (*\x. c x \ t[x] to c \ \x. t[x]*) fun abs_def eq = let val body = Term.strip_all_body eq; val vars = map Free (Term.rename_wrt_term body (Term.strip_all_vars eq)); val (lhs, rhs) = Logic.dest_equals (Term.subst_bounds (vars, body)); val (lhs', args) = Term.strip_comb lhs; val rhs' = fold_rev (absfree o dest_Free) args rhs; in (lhs', rhs') end; end; diff --git a/src/Pure/proofterm.ML b/src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML +++ b/src/Pure/proofterm.ML @@ -1,2326 +1,2341 @@ (* Title: Pure/proofterm.ML Author: Stefan Berghofer, TU Muenchen LF style proof terms. *) infix 8 % %% %>; signature PROOFTERM = sig type thm_header = {serial: serial, pos: Position.T list, theory_name: string, name: string, prop: term, types: typ list option} type thm_body type thm_node datatype proof = MinProof | PBound of int | Abst of string * typ option * proof | AbsP of string * term option * proof | % of proof * term option | %% of proof * proof | Hyp of term | PAxm of string * term * typ list option | PClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} type oracle = (string * Position.T) * term option type thm = serial * thm_node exception MIN_PROOF of unit val proof_of: proof_body -> proof val join_proof: proof_body future -> proof val map_proof_of: (proof -> proof) -> proof_body -> proof_body val thm_header: serial -> Position.T list -> string -> string -> term -> typ list option -> thm_header val thm_body: proof_body -> thm_body val thm_body_proof_raw: thm_body -> proof val thm_body_proof_open: thm_body -> proof val thm_node_theory_name: thm_node -> string val thm_node_name: thm_node -> string val thm_node_prop: thm_node -> term val thm_node_body: thm_node -> proof_body future val thm_node_thms: thm_node -> thm list val join_thms: thm list -> proof_body list val make_thm: thm_header -> thm_body -> thm val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a val fold_body_thms: ({serial: serial, name: string, prop: term, body: proof_body} -> 'a -> 'a) -> proof_body list -> 'a -> 'a val oracle_ord: oracle ord val thm_ord: thm ord val unions_oracles: oracle Ord_List.T list -> oracle Ord_List.T val unions_thms: thm Ord_List.T list -> thm Ord_List.T val no_proof_body: proof -> proof_body val no_thm_names: proof -> proof val no_thm_proofs: proof -> proof val no_body_proofs: proof -> proof val encode: Consts.T -> proof XML.Encode.T val encode_body: Consts.T -> proof_body XML.Encode.T val encode_standard_term: Consts.T -> term XML.Encode.T val encode_standard_proof: Consts.T -> proof XML.Encode.T val decode: Consts.T -> proof XML.Decode.T val decode_body: Consts.T -> proof_body XML.Decode.T val %> : proof * term -> proof (*primitive operations*) val proofs: int Unsynchronized.ref val proofs_enabled: unit -> bool val atomic_proof: proof -> bool val compact_proof: proof -> bool val proof_combt: proof * term list -> proof val proof_combt': proof * term option list -> proof val proof_combP: proof * proof list -> proof val strip_combt: proof -> proof * term option list val strip_combP: proof -> proof * proof list val strip_thm_body: proof_body -> proof_body val map_proof_same: term Same.operation -> typ Same.operation -> (typ * class -> proof) -> proof Same.operation val map_proof_terms_same: term Same.operation -> typ Same.operation -> proof Same.operation val map_proof_types_same: typ Same.operation -> proof Same.operation val map_proof_terms: (term -> term) -> (typ -> typ) -> proof -> proof val map_proof_types: (typ -> typ) -> proof -> proof val fold_proof_terms: (term -> 'a -> 'a) -> proof -> 'a -> 'a val fold_proof_terms_types: (term -> 'a -> 'a) -> (typ -> 'a -> 'a) -> proof -> 'a -> 'a val maxidx_proof: proof -> int -> int val size_of_proof: proof -> int val change_types: typ list option -> proof -> proof val prf_abstract_over: term -> proof -> proof val prf_incr_bv: int -> int -> int -> int -> proof -> proof val incr_pboundvars: int -> int -> proof -> proof val prf_loose_bvar1: proof -> int -> bool val prf_loose_Pbvar1: proof -> int -> bool val prf_add_loose_bnos: int -> int -> proof -> int list * int list -> int list * int list val norm_proof: Envir.env -> proof -> proof val norm_proof': Envir.env -> proof -> proof val prf_subst_bounds: term list -> proof -> proof val prf_subst_pbounds: proof list -> proof -> proof val freeze_thaw_prf: proof -> proof * (proof -> proof) (*proof terms for specific inference rules*) val trivial_proof: proof val implies_intr_proof: term -> proof -> proof val implies_intr_proof': term -> proof -> proof val forall_intr_proof: string * term -> typ option -> proof -> proof val forall_intr_proof': term -> proof -> proof - val varify_proof: term -> (string * sort) list -> proof -> proof + val varify_proof: term -> Term_Subst.TFrees.set -> proof -> proof val legacy_freezeT: term -> proof -> proof val rotate_proof: term list -> term -> (string * typ) list -> term list -> int -> proof -> proof val permute_prems_proof: term list -> int -> int -> proof -> proof val generalize_proof: Symtab.set * Symtab.set -> int -> term -> proof -> proof val instantiate: typ Term_Subst.TVars.table * term Term_Subst.Vars.table -> proof -> proof val lift_proof: term -> int -> term -> proof -> proof val incr_indexes: int -> proof -> proof val assumption_proof: term list -> term -> int -> proof -> proof val bicompose_proof: bool -> term list -> term list -> term option -> term list -> int -> int -> proof -> proof -> proof val equality_axms: (string * term) list val reflexive_axm: proof val symmetric_axm: proof val transitive_axm: proof val equal_intr_axm: proof val equal_elim_axm: proof val abstract_rule_axm: proof val combination_axm: proof val reflexive_proof: proof val symmetric_proof: proof -> proof val transitive_proof: typ -> term -> proof -> proof -> proof val equal_intr_proof: term -> term -> proof -> proof -> proof val equal_elim_proof: term -> term -> proof -> proof -> proof val abstract_rule_proof: string * term -> proof -> proof val combination_proof: term -> term -> term -> term -> proof -> proof -> proof val strip_shyps_proof: Sorts.algebra -> (typ * sort) list -> (typ * sort) list -> sort list -> proof -> proof val of_sort_proof: Sorts.algebra -> (class * class -> proof) -> (string * class list list * class -> proof) -> (typ * class -> proof) -> typ * sort -> proof list val axm_proof: string -> term -> proof val oracle_proof: string -> term -> proof val shrink_proof: proof -> proof (*rewriting on proof terms*) val add_prf_rrule: proof * proof -> theory -> theory val add_prf_rproc: (typ list -> term option list -> proof -> (proof * proof) option) -> theory -> theory val set_preproc: (theory -> proof -> proof) -> theory -> theory val apply_preproc: theory -> proof -> proof val forall_intr_variables_term: term -> term val forall_intr_variables: term -> proof -> proof val no_skel: proof val normal_skel: proof val rewrite_proof: theory -> (proof * proof) list * (typ list -> term option list -> proof -> (proof * proof) option) list -> proof -> proof val rewrite_proof_notypes: (proof * proof) list * (typ list -> term option list -> proof -> (proof * proof) option) list -> proof -> proof val rew_proof: theory -> proof -> proof val reconstruct_proof: theory -> term -> proof -> proof val prop_of': term list -> proof -> term val prop_of: proof -> term val expand_name_empty: thm_header -> string option val expand_proof: theory -> (thm_header -> string option) -> proof -> proof val standard_vars: Name.context -> term * proof option -> term * proof option val standard_vars_term: Name.context -> term -> term val add_standard_vars: proof -> (string * typ) list -> (string * typ) list val add_standard_vars_term: term -> (string * typ) list -> (string * typ) list val export_enabled: unit -> bool val export_standard_enabled: unit -> bool val export_proof_boxes_required: theory -> bool val export_proof_boxes: proof_body list -> unit val fulfill_norm_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body val thm_proof: theory -> (class * class -> proof) -> (string * class list list * class -> proof) -> string * Position.T -> sort list -> term list -> term -> (serial * proof_body future) list -> proof_body -> thm * proof val unconstrain_thm_proof: theory -> (class * class -> proof) -> (string * class list list * class -> proof) -> sort list -> term -> (serial * proof_body future) list -> proof_body -> thm * proof val get_identity: sort list -> term list -> term -> proof -> {serial: serial, theory_name: string, name: string} option val get_approximative_name: sort list -> term list -> term -> proof -> string type thm_id = {serial: serial, theory_name: string} val make_thm_id: serial * string -> thm_id val thm_header_id: thm_header -> thm_id val thm_id: thm -> thm_id val get_id: sort list -> term list -> term -> proof -> thm_id option val this_id: thm_id option -> thm_id -> bool val proof_boxes: {excluded: thm_id -> bool, included: thm_id -> bool} -> proof list -> (thm_header * proof) list (*exception MIN_PROOF*) end structure Proofterm : PROOFTERM = struct (** datatype proof **) type thm_header = {serial: serial, pos: Position.T list, theory_name: string, name: string, prop: term, types: typ list option}; datatype proof = MinProof | PBound of int | Abst of string * typ option * proof | AbsP of string * term option * proof | op % of proof * term option | op %% of proof * proof | Hyp of term | PAxm of string * term * typ list option | PClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} and thm_body = Thm_Body of {open_proof: proof -> proof, body: proof_body future} and thm_node = Thm_Node of {theory_name: string, name: string, prop: term, body: proof_body future, export: unit lazy, consolidate: unit lazy}; type oracle = (string * Position.T) * term option; val oracle_ord: oracle ord = prod_ord (prod_ord fast_string_ord Position.ord) (option_ord Term_Ord.fast_term_ord); type thm = serial * thm_node; val thm_ord: thm ord = fn ((i, _), (j, _)) => int_ord (j, i); exception MIN_PROOF of unit; fun proof_of (PBody {proof, ...}) = proof; val join_proof = Future.join #> proof_of; fun map_proof_of f (PBody {oracles, thms, proof}) = PBody {oracles = oracles, thms = thms, proof = f proof}; fun thm_header serial pos theory_name name prop types : thm_header = {serial = serial, pos = pos, theory_name = theory_name, name = name, prop = prop, types = types}; fun thm_body body = Thm_Body {open_proof = I, body = Future.value body}; fun thm_body_proof_raw (Thm_Body {body, ...}) = join_proof body; fun thm_body_proof_open (Thm_Body {open_proof, body, ...}) = open_proof (join_proof body); fun rep_thm_node (Thm_Node args) = args; val thm_node_theory_name = #theory_name o rep_thm_node; val thm_node_name = #name o rep_thm_node; val thm_node_prop = #prop o rep_thm_node; val thm_node_body = #body o rep_thm_node; val thm_node_thms = thm_node_body #> Future.join #> (fn PBody {thms, ...} => thms); val thm_node_export = #export o rep_thm_node; val thm_node_consolidate = #consolidate o rep_thm_node; fun join_thms (thms: thm list) = Future.joins (map (thm_node_body o #2) thms); val consolidate_bodies = maps (fn PBody {thms, ...} => map (thm_node_consolidate o #2) thms) #> Lazy.consolidate #> map Lazy.force #> ignore; fun make_thm_node theory_name name prop body export = let val consolidate = Lazy.lazy_name "Proofterm.make_thm_node" (fn () => let val PBody {thms, ...} = Future.join body in consolidate_bodies (join_thms thms) end); in Thm_Node {theory_name = theory_name, name = name, prop = prop, body = body, export = export, consolidate = consolidate} end; val no_export = Lazy.value (); fun make_thm ({serial, theory_name, name, prop, ...}: thm_header) (Thm_Body {body, ...}) = (serial, make_thm_node theory_name name prop body no_export); (* proof atoms *) fun fold_proof_atoms all f = let fun app (Abst (_, _, prf)) = app prf | app (AbsP (_, _, prf)) = app prf | app (prf % _) = app prf | app (prf1 %% prf2) = app prf1 #> app prf2 | app (prf as PThm ({serial = i, ...}, Thm_Body {body, ...})) = (fn (x, seen) => if Inttab.defined seen i then (x, seen) else let val (x', seen') = (if all then app (join_proof body) else I) (x, Inttab.update (i, ()) seen) in (f prf x', seen') end) | app prf = (fn (x, seen) => (f prf x, seen)); in fn prfs => fn x => #1 (fold app prfs (x, Inttab.empty)) end; fun fold_body_thms f = let fun app (PBody {thms, ...}) = tap join_thms thms |> fold (fn (i, thm_node) => fn (x, seen) => if Inttab.defined seen i then (x, seen) else let val name = thm_node_name thm_node; val prop = thm_node_prop thm_node; val body = Future.join (thm_node_body thm_node); val (x', seen') = app body (x, Inttab.update (i, ()) seen); in (f {serial = i, name = name, prop = prop, body = body} x', seen') end); in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end; (* proof body *) val unions_oracles = Ord_List.unions oracle_ord; val unions_thms = Ord_List.unions thm_ord; fun no_proof_body proof = PBody {oracles = [], thms = [], proof = proof}; val no_thm_body = thm_body (no_proof_body MinProof); fun no_thm_names (Abst (x, T, prf)) = Abst (x, T, no_thm_names prf) | no_thm_names (AbsP (x, t, prf)) = AbsP (x, t, no_thm_names prf) | no_thm_names (prf % t) = no_thm_names prf % t | no_thm_names (prf1 %% prf2) = no_thm_names prf1 %% no_thm_names prf2 | no_thm_names (PThm ({serial, pos, theory_name, name = _, prop, types}, thm_body)) = PThm (thm_header serial pos theory_name "" prop types, thm_body) | no_thm_names a = a; fun no_thm_proofs (Abst (x, T, prf)) = Abst (x, T, no_thm_proofs prf) | no_thm_proofs (AbsP (x, t, prf)) = AbsP (x, t, no_thm_proofs prf) | no_thm_proofs (prf % t) = no_thm_proofs prf % t | no_thm_proofs (prf1 %% prf2) = no_thm_proofs prf1 %% no_thm_proofs prf2 | no_thm_proofs (PThm (header, _)) = PThm (header, no_thm_body) | no_thm_proofs a = a; fun no_body_proofs (Abst (x, T, prf)) = Abst (x, T, no_body_proofs prf) | no_body_proofs (AbsP (x, t, prf)) = AbsP (x, t, no_body_proofs prf) | no_body_proofs (prf % t) = no_body_proofs prf % t | no_body_proofs (prf1 %% prf2) = no_body_proofs prf1 %% no_body_proofs prf2 | no_body_proofs (PThm (header, Thm_Body {open_proof, body})) = let val body' = Future.value (no_proof_body (join_proof body)); val thm_body' = Thm_Body {open_proof = open_proof, body = body'}; in PThm (header, thm_body') end | no_body_proofs a = a; (** XML data representation **) (* encode *) local open XML.Encode Term_XML.Encode; fun proof consts prf = prf |> variant [fn MinProof => ([], []), fn PBound a => ([], int a), fn Abst (a, b, c) => ([a], pair (option typ) (proof consts) (b, c)), fn AbsP (a, b, c) => ([a], pair (option (term consts)) (proof consts) (b, c)), fn a % b => ([], pair (proof consts) (option (term consts)) (a, b)), fn a %% b => ([], pair (proof consts) (proof consts) (a, b)), fn Hyp a => ([], term consts a), fn PAxm (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PClass (a, b) => ([b], typ a), fn Oracle (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PThm ({serial, pos, theory_name, name, prop, types}, Thm_Body {open_proof, body, ...}) => ([int_atom serial, theory_name, name], pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) (map Position.properties_of pos, (prop, (types, map_proof_of open_proof (Future.join body)))))] and proof_body consts (PBody {oracles, thms, proof = prf}) = triple (list (pair (pair string (properties o Position.properties_of)) (option (term consts)))) (list (thm consts)) (proof consts) (oracles, thms, prf) and thm consts (a, thm_node) = pair int (pair string (pair string (pair (term consts) (proof_body consts)))) (a, (thm_node_theory_name thm_node, (thm_node_name thm_node, (thm_node_prop thm_node, (Future.join (thm_node_body thm_node)))))); fun standard_term consts t = t |> variant [fn Const (a, b) => ([a], list typ (Consts.typargs consts (a, b))), fn Free (a, _) => ([a], []), fn Var (a, _) => (indexname a, []), fn Bound a => ([], int a), fn Abs (a, b, c) => ([a], pair typ (standard_term consts) (b, c)), fn op $ a => ([], pair (standard_term consts) (standard_term consts) a)]; fun standard_proof consts prf = prf |> variant [fn MinProof => ([], []), fn PBound a => ([], int a), fn Abst (a, SOME b, c) => ([a], pair typ (standard_proof consts) (b, c)), fn AbsP (a, SOME b, c) => ([a], pair (standard_term consts) (standard_proof consts) (b, c)), fn a % SOME b => ([], pair (standard_proof consts) (standard_term consts) (a, b)), fn a %% b => ([], pair (standard_proof consts) (standard_proof consts) (a, b)), fn Hyp a => ([], standard_term consts a), fn PAxm (name, _, SOME Ts) => ([name], list typ Ts), fn PClass (T, c) => ([c], typ T), fn Oracle (name, prop, SOME Ts) => ([name], pair (standard_term consts) (list typ) (prop, Ts)), fn PThm ({serial, theory_name, name, types = SOME Ts, ...}, _) => ([int_atom serial, theory_name, name], list typ Ts)]; in val encode = proof; val encode_body = proof_body; val encode_standard_term = standard_term; val encode_standard_proof = standard_proof; end; (* decode *) local open XML.Decode Term_XML.Decode; fun proof consts prf = prf |> variant [fn ([], []) => MinProof, fn ([], a) => PBound (int a), fn ([a], b) => let val (c, d) = pair (option typ) (proof consts) b in Abst (a, c, d) end, fn ([a], b) => let val (c, d) = pair (option (term consts)) (proof consts) b in AbsP (a, c, d) end, fn ([], a) => op % (pair (proof consts) (option (term consts)) a), fn ([], a) => op %% (pair (proof consts) (proof consts) a), fn ([], a) => Hyp (term consts a), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in PAxm (a, c, d) end, fn ([b], a) => PClass (typ a, b), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in Oracle (a, c, d) end, fn ([a, b, c], d) => let val ((e, (f, (g, h)))) = pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) d; val header = thm_header (int_atom a) (map Position.of_properties e) b c f g; in PThm (header, thm_body h) end] and proof_body consts x = let val (a, b, c) = triple (list (pair (pair string (Position.of_properties o properties)) (option (term consts)))) (list (thm consts)) (proof consts) x; in PBody {oracles = a, thms = b, proof = c} end and thm consts x = let val (a, (b, (c, (d, e)))) = pair int (pair string (pair string (pair (term consts) (proof_body consts)))) x in (a, make_thm_node b c d (Future.value e) no_export) end; in val decode = proof; val decode_body = proof_body; end; (** proof objects with different levels of detail **) val proofs = Unsynchronized.ref 2; fun proofs_enabled () = ! proofs >= 2; fun atomic_proof prf = (case prf of Abst _ => false | AbsP _ => false | op % _ => false | op %% _ => false | MinProof => false | _ => true); fun compact_proof (prf % _) = compact_proof prf | compact_proof (prf1 %% prf2) = atomic_proof prf2 andalso compact_proof prf1 | compact_proof prf = atomic_proof prf; fun (prf %> t) = prf % SOME t; val proof_combt = Library.foldl (op %>); val proof_combt' = Library.foldl (op %); val proof_combP = Library.foldl (op %%); fun strip_combt prf = let fun stripc (prf % t, ts) = stripc (prf, t::ts) | stripc x = x in stripc (prf, []) end; fun strip_combP prf = let fun stripc (prf %% prf', prfs) = stripc (prf, prf'::prfs) | stripc x = x in stripc (prf, []) end; fun strip_thm_body (body as PBody {proof, ...}) = (case fst (strip_combt (fst (strip_combP proof))) of PThm (_, Thm_Body {body = body', ...}) => Future.join body' | _ => body); val mk_Abst = fold_rev (fn (x, _: typ) => fn prf => Abst (x, NONE, prf)); val mk_AbsP = fold_rev (fn _: term => fn prf => AbsP ("H", NONE, prf)); fun map_proof_same term typ ofclass = let val typs = Same.map typ; fun proof (Abst (s, T, prf)) = (Abst (s, Same.map_option typ T, Same.commit proof prf) handle Same.SAME => Abst (s, T, proof prf)) | proof (AbsP (s, t, prf)) = (AbsP (s, Same.map_option term t, Same.commit proof prf) handle Same.SAME => AbsP (s, t, proof prf)) | proof (prf % t) = (proof prf % Same.commit (Same.map_option term) t handle Same.SAME => prf % Same.map_option term t) | proof (prf1 %% prf2) = (proof prf1 %% Same.commit proof prf2 handle Same.SAME => prf1 %% proof prf2) | proof (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (typs Ts)) | proof (PClass T_c) = ofclass T_c | proof (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (typs Ts)) | proof (PThm ({serial, pos, theory_name, name, prop, types = SOME Ts}, thm_body)) = PThm (thm_header serial pos theory_name name prop (SOME (typs Ts)), thm_body) | proof _ = raise Same.SAME; in proof end; fun map_proof_terms_same term typ = map_proof_same term typ (fn (T, c) => PClass (typ T, c)); fun map_proof_types_same typ = map_proof_terms_same (Term_Subst.map_types_same typ) typ; fun same eq f x = let val x' = f x in if eq (x, x') then raise Same.SAME else x' end; fun map_proof_terms f g = Same.commit (map_proof_terms_same (same (op =) f) (same (op =) g)); fun map_proof_types f = Same.commit (map_proof_types_same (same (op =) f)); fun fold_proof_terms f (Abst (_, _, prf)) = fold_proof_terms f prf | fold_proof_terms f (AbsP (_, SOME t, prf)) = f t #> fold_proof_terms f prf | fold_proof_terms f (AbsP (_, NONE, prf)) = fold_proof_terms f prf | fold_proof_terms f (prf % SOME t) = fold_proof_terms f prf #> f t | fold_proof_terms f (prf % NONE) = fold_proof_terms f prf | fold_proof_terms f (prf1 %% prf2) = fold_proof_terms f prf1 #> fold_proof_terms f prf2 | fold_proof_terms _ _ = I; fun fold_proof_terms_types f g (Abst (_, SOME T, prf)) = g T #> fold_proof_terms_types f g prf | fold_proof_terms_types f g (Abst (_, NONE, prf)) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (AbsP (_, SOME t, prf)) = f t #> fold_proof_terms_types f g prf | fold_proof_terms_types f g (AbsP (_, NONE, prf)) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (prf % SOME t) = fold_proof_terms_types f g prf #> f t | fold_proof_terms_types f g (prf % NONE) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (prf1 %% prf2) = fold_proof_terms_types f g prf1 #> fold_proof_terms_types f g prf2 | fold_proof_terms_types _ g (PAxm (_, _, SOME Ts)) = fold g Ts | fold_proof_terms_types _ g (PClass (T, _)) = g T | fold_proof_terms_types _ g (Oracle (_, _, SOME Ts)) = fold g Ts | fold_proof_terms_types _ g (PThm ({types = SOME Ts, ...}, _)) = fold g Ts | fold_proof_terms_types _ _ _ = I; fun maxidx_proof prf = fold_proof_terms_types Term.maxidx_term Term.maxidx_typ prf; fun size_of_proof (Abst (_, _, prf)) = 1 + size_of_proof prf | size_of_proof (AbsP (_, _, prf)) = 1 + size_of_proof prf | size_of_proof (prf % _) = 1 + size_of_proof prf | size_of_proof (prf1 %% prf2) = size_of_proof prf1 + size_of_proof prf2 | size_of_proof _ = 1; fun change_types types (PAxm (name, prop, _)) = PAxm (name, prop, types) | change_types (SOME [T]) (PClass (_, c)) = PClass (T, c) | change_types types (Oracle (name, prop, _)) = Oracle (name, prop, types) | change_types types (PThm ({serial, pos, theory_name, name, prop, types = _}, thm_body)) = PThm (thm_header serial pos theory_name name prop types, thm_body) | change_types _ prf = prf; (* utilities *) fun strip_abs (_::Ts) (Abs (_, _, t)) = strip_abs Ts t | strip_abs _ t = t; fun mk_abs Ts t = Library.foldl (fn (t', T) => Abs ("", T, t')) (t, Ts); (*Abstraction of a proof term over its occurrences of v, which must contain no loose bound variables. The resulting proof term is ready to become the body of an Abst.*) fun prf_abstract_over v = let fun abst' lev u = if v aconv u then Bound lev else (case u of Abs (a, T, t) => Abs (a, T, abst' (lev + 1) t) | f $ t => (abst' lev f $ absth' lev t handle Same.SAME => f $ abst' lev t) | _ => raise Same.SAME) and absth' lev t = (abst' lev t handle Same.SAME => t); fun abst lev (AbsP (a, t, prf)) = (AbsP (a, Same.map_option (abst' lev) t, absth lev prf) handle Same.SAME => AbsP (a, t, abst lev prf)) | abst lev (Abst (a, T, prf)) = Abst (a, T, abst (lev + 1) prf) | abst lev (prf1 %% prf2) = (abst lev prf1 %% absth lev prf2 handle Same.SAME => prf1 %% abst lev prf2) | abst lev (prf % t) = (abst lev prf % Option.map (absth' lev) t handle Same.SAME => prf % Same.map_option (abst' lev) t) | abst _ _ = raise Same.SAME and absth lev prf = (abst lev prf handle Same.SAME => prf); in absth 0 end; (*increments a proof term's non-local bound variables required when moving a proof term within abstractions inc is increment for bound variables lev is level at which a bound variable is considered 'loose'*) fun incr_bv' inct tlev t = incr_bv (inct, tlev, t); fun prf_incr_bv' incP _ Plev _ (PBound i) = if i >= Plev then PBound (i+incP) else raise Same.SAME | prf_incr_bv' incP inct Plev tlev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (same (op =) (incr_bv' inct tlev)) t, prf_incr_bv incP inct (Plev+1) tlev body) handle Same.SAME => AbsP (a, t, prf_incr_bv' incP inct (Plev+1) tlev body)) | prf_incr_bv' incP inct Plev tlev (Abst (a, T, body)) = Abst (a, T, prf_incr_bv' incP inct Plev (tlev+1) body) | prf_incr_bv' incP inct Plev tlev (prf %% prf') = (prf_incr_bv' incP inct Plev tlev prf %% prf_incr_bv incP inct Plev tlev prf' handle Same.SAME => prf %% prf_incr_bv' incP inct Plev tlev prf') | prf_incr_bv' incP inct Plev tlev (prf % t) = (prf_incr_bv' incP inct Plev tlev prf % Option.map (incr_bv' inct tlev) t handle Same.SAME => prf % Same.map_option (same (op =) (incr_bv' inct tlev)) t) | prf_incr_bv' _ _ _ _ _ = raise Same.SAME and prf_incr_bv incP inct Plev tlev prf = (prf_incr_bv' incP inct Plev tlev prf handle Same.SAME => prf); fun incr_pboundvars 0 0 prf = prf | incr_pboundvars incP inct prf = prf_incr_bv incP inct 0 0 prf; fun prf_loose_bvar1 (prf1 %% prf2) k = prf_loose_bvar1 prf1 k orelse prf_loose_bvar1 prf2 k | prf_loose_bvar1 (prf % SOME t) k = prf_loose_bvar1 prf k orelse loose_bvar1 (t, k) | prf_loose_bvar1 (_ % NONE) _ = true | prf_loose_bvar1 (AbsP (_, SOME t, prf)) k = loose_bvar1 (t, k) orelse prf_loose_bvar1 prf k | prf_loose_bvar1 (AbsP (_, NONE, _)) _ = true | prf_loose_bvar1 (Abst (_, _, prf)) k = prf_loose_bvar1 prf (k+1) | prf_loose_bvar1 _ _ = false; fun prf_loose_Pbvar1 (PBound i) k = i = k | prf_loose_Pbvar1 (prf1 %% prf2) k = prf_loose_Pbvar1 prf1 k orelse prf_loose_Pbvar1 prf2 k | prf_loose_Pbvar1 (prf % _) k = prf_loose_Pbvar1 prf k | prf_loose_Pbvar1 (AbsP (_, _, prf)) k = prf_loose_Pbvar1 prf (k+1) | prf_loose_Pbvar1 (Abst (_, _, prf)) k = prf_loose_Pbvar1 prf k | prf_loose_Pbvar1 _ _ = false; fun prf_add_loose_bnos plev _ (PBound i) (is, js) = if i < plev then (is, js) else (insert (op =) (i-plev) is, js) | prf_add_loose_bnos plev tlev (prf1 %% prf2) p = prf_add_loose_bnos plev tlev prf2 (prf_add_loose_bnos plev tlev prf1 p) | prf_add_loose_bnos plev tlev (prf % opt) (is, js) = prf_add_loose_bnos plev tlev prf (case opt of NONE => (is, insert (op =) ~1 js) | SOME t => (is, add_loose_bnos (t, tlev, js))) | prf_add_loose_bnos plev tlev (AbsP (_, opt, prf)) (is, js) = prf_add_loose_bnos (plev+1) tlev prf (case opt of NONE => (is, insert (op =) ~1 js) | SOME t => (is, add_loose_bnos (t, tlev, js))) | prf_add_loose_bnos plev tlev (Abst (_, _, prf)) p = prf_add_loose_bnos plev (tlev+1) prf p | prf_add_loose_bnos _ _ _ _ = ([], []); (* substitutions *) -fun del_conflicting_tvars envT T = +local + +fun conflicting_tvarsT envT = + Term.fold_atyps + (fn T => fn instT => + (case T of + TVar (v as (_, S)) => + if Term_Subst.TVars.defined instT v orelse can (Type.lookup envT) v then instT + else Term_Subst.TVars.update (v, Logic.dummy_tfree S) instT + | _ => instT)); + +fun conflicting_tvars env = + Term.fold_aterms + (fn t => fn inst => + (case t of + Var (v as (_, T)) => + if Term_Subst.Vars.defined inst v orelse can (Envir.lookup env) v then inst + else Term_Subst.Vars.update (v, Free ("dummy", T)) inst + | _ => inst)); + +fun del_conflicting_tvars envT ty = + Term_Subst.instantiateT (Term_Subst.TVars.build (conflicting_tvarsT envT ty)) ty; + +fun del_conflicting_vars env tm = let val instT = - map_filter (fn ixnS as (_, S) => - (Type.lookup envT ixnS; NONE) handle TYPE _ => - SOME (ixnS, Logic.dummy_tfree S)) (Term.add_tvarsT T []) - in Term_Subst.instantiateT (Term_Subst.TVars.table instT) T end; + Term_Subst.TVars.build (tm |> Term.fold_types (conflicting_tvarsT (Envir.type_env env))); + val inst = Term_Subst.Vars.build (tm |> conflicting_tvars env); + in Term_Subst.instantiate (instT, inst) tm end; -fun del_conflicting_vars env t = - let - val instT = - 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 []); - val inst = - map_filter (fn (ixnT as (_, T)) => - (Envir.lookup env ixnT; NONE) handle TYPE _ => - SOME (ixnT, Free ("dummy", T))) (Term.add_vars t []); - in Term_Subst.instantiate (Term_Subst.TVars.make instT, Term_Subst.Vars.table inst) t end; +in fun norm_proof env = let val envT = Envir.type_env env; fun msg s = warning ("type conflict in norm_proof:\n" ^ s); fun htype f t = f env t handle TYPE (s, _, _) => (msg s; f env (del_conflicting_vars env t)); fun htypeT f T = f envT T handle TYPE (s, _, _) => (msg s; f envT (del_conflicting_tvars envT T)); fun htypeTs f Ts = f envT Ts handle TYPE (s, _, _) => (msg s; f envT (map (del_conflicting_tvars envT) Ts)); fun norm (Abst (s, T, prf)) = (Abst (s, Same.map_option (htypeT Envir.norm_type_same) T, Same.commit norm prf) handle Same.SAME => Abst (s, T, norm prf)) | norm (AbsP (s, t, prf)) = (AbsP (s, Same.map_option (htype Envir.norm_term_same) t, Same.commit norm prf) handle Same.SAME => AbsP (s, t, norm prf)) | norm (prf % t) = (norm prf % Option.map (htype Envir.norm_term) t handle Same.SAME => prf % Same.map_option (htype Envir.norm_term_same) t) | norm (prf1 %% prf2) = (norm prf1 %% Same.commit norm prf2 handle Same.SAME => prf1 %% norm prf2) | norm (PAxm (s, prop, Ts)) = PAxm (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) | norm (PClass (T, c)) = PClass (htypeT Envir.norm_type_same T, c) | norm (Oracle (s, prop, Ts)) = Oracle (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) | norm (PThm ({serial = i, pos = p, theory_name, name = a, prop = t, types = Ts}, thm_body)) = PThm (thm_header i p theory_name a t (Same.map_option (htypeTs Envir.norm_types_same) Ts), thm_body) | norm _ = raise Same.SAME; in Same.commit norm end; +end; + (* remove some types in proof term (to save space) *) fun remove_types (Abs (s, _, t)) = Abs (s, dummyT, remove_types t) | remove_types (t $ u) = remove_types t $ remove_types u | remove_types (Const (s, _)) = Const (s, dummyT) | remove_types t = t; fun remove_types_env (Envir.Envir {maxidx, tenv, tyenv}) = Envir.Envir {maxidx = maxidx, tenv = Vartab.map (K (apsnd remove_types)) tenv, tyenv = tyenv}; fun norm_proof' env prf = norm_proof (remove_types_env env) prf; (* substitution of bound variables *) fun prf_subst_bounds args prf = let val n = length args; fun subst' lev (Bound i) = (if i Bound (i-n)) (*loose: change it*) | subst' lev (Abs (a, T, body)) = Abs (a, T, subst' (lev+1) body) | subst' lev (f $ t) = (subst' lev f $ substh' lev t handle Same.SAME => f $ subst' lev t) | subst' _ _ = raise Same.SAME and substh' lev t = (subst' lev t handle Same.SAME => t); fun subst lev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (subst' lev) t, substh lev body) handle Same.SAME => AbsP (a, t, subst lev body)) | subst lev (Abst (a, T, body)) = Abst (a, T, subst (lev+1) body) | subst lev (prf %% prf') = (subst lev prf %% substh lev prf' handle Same.SAME => prf %% subst lev prf') | subst lev (prf % t) = (subst lev prf % Option.map (substh' lev) t handle Same.SAME => prf % Same.map_option (subst' lev) t) | subst _ _ = raise Same.SAME and substh lev prf = (subst lev prf handle Same.SAME => prf); in (case args of [] => prf | _ => substh 0 prf) end; fun prf_subst_pbounds args prf = let val n = length args; fun subst (PBound i) Plev tlev = (if i < Plev then raise Same.SAME (*var is locally bound*) else incr_pboundvars Plev tlev (nth args (i-Plev)) handle General.Subscript => PBound (i-n) (*loose: change it*)) | subst (AbsP (a, t, body)) Plev tlev = AbsP (a, t, subst body (Plev+1) tlev) | subst (Abst (a, T, body)) Plev tlev = Abst (a, T, subst body Plev (tlev+1)) | subst (prf %% prf') Plev tlev = (subst prf Plev tlev %% substh prf' Plev tlev handle Same.SAME => prf %% subst prf' Plev tlev) | subst (prf % t) Plev tlev = subst prf Plev tlev % t | subst _ _ _ = raise Same.SAME and substh prf Plev tlev = (subst prf Plev tlev handle Same.SAME => prf) in (case args of [] => prf | _ => substh prf 0 0) end; (* freezing and thawing of variables in proof terms *) local fun frzT names = map_type_tvar (fn (ixn, S) => TFree (the (AList.lookup (op =) names ixn), S)); fun thawT names = map_type_tfree (fn (a, S) => (case AList.lookup (op =) names a of NONE => TFree (a, S) | SOME ixn => TVar (ixn, S))); fun freeze names names' (t $ u) = freeze names names' t $ freeze names names' u | freeze names names' (Abs (s, T, t)) = Abs (s, frzT names' T, freeze names names' t) | freeze _ names' (Const (s, T)) = Const (s, frzT names' T) | freeze _ names' (Free (s, T)) = Free (s, frzT names' T) | freeze names names' (Var (ixn, T)) = Free (the (AList.lookup (op =) names ixn), frzT names' T) | freeze _ _ t = t; fun thaw names names' (t $ u) = thaw names names' t $ thaw names names' u | thaw names names' (Abs (s, T, t)) = Abs (s, thawT names' T, thaw names names' t) | thaw _ names' (Const (s, T)) = Const (s, thawT names' T) | thaw names names' (Free (s, T)) = let val T' = thawT names' T in (case AList.lookup (op =) names s of NONE => Free (s, T') | SOME ixn => Var (ixn, T')) end | thaw _ names' (Var (ixn, T)) = Var (ixn, thawT names' T) | thaw _ _ t = t; in fun freeze_thaw_prf prf = let val (fs, Tfs, vs, Tvs) = fold_proof_terms_types (fn t => fn (fs, Tfs, vs, Tvs) => (Term.add_free_names t fs, Term.add_tfree_names t Tfs, Term.add_var_names t vs, Term.add_tvar_names t Tvs)) (fn T => fn (fs, Tfs, vs, Tvs) => (fs, Term.add_tfree_namesT T Tfs, vs, Term.add_tvar_namesT T Tvs)) prf ([], [], [], []); val names = vs ~~ Name.variant_list fs (map fst vs); val names' = Tvs ~~ Name.variant_list Tfs (map fst Tvs); val rnames = map swap names; val rnames' = map swap names'; in (map_proof_terms (freeze names names') (frzT names') prf, map_proof_terms (thaw rnames rnames') (thawT rnames')) end; end; (** inference rules **) (* trivial implication *) val trivial_proof = AbsP ("H", NONE, PBound 0); (* implication introduction *) fun gen_implies_intr_proof f h prf = let fun abshyp i (Hyp t) = if h aconv t then PBound i else raise Same.SAME | abshyp i (Abst (s, T, prf)) = Abst (s, T, abshyp i prf) | abshyp i (AbsP (s, t, prf)) = AbsP (s, t, abshyp (i + 1) prf) | abshyp i (prf % t) = abshyp i prf % t | abshyp i (prf1 %% prf2) = (abshyp i prf1 %% abshyph i prf2 handle Same.SAME => prf1 %% abshyp i prf2) | abshyp _ _ = raise Same.SAME and abshyph i prf = (abshyp i prf handle Same.SAME => prf); in AbsP ("H", f h, abshyph 0 prf) end; val implies_intr_proof = gen_implies_intr_proof (K NONE); val implies_intr_proof' = gen_implies_intr_proof SOME; (* forall introduction *) fun forall_intr_proof (a, v) opt_T prf = Abst (a, opt_T, prf_abstract_over v prf); fun forall_intr_proof' v prf = let val (a, T) = (case v of Var ((a, _), T) => (a, T) | Free (a, T) => (a, T)) in forall_intr_proof (a, v) (SOME T) prf end; (* varify *) fun varify_proof t fixed prf = let - val fs = Term.fold_types (Term.fold_atyps - (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t []; + val fs = + build (t |> (Term.fold_types o Term.fold_atyps) + (fn TFree v => if Term_Subst.TFrees.defined fixed v then I else insert (op =) v | _ => I)); val used = Name.context |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t; val fmap = fs ~~ #1 (fold_map Name.variant (map fst fs) used); fun thaw (a, S) = (case AList.lookup (op =) fmap (a, S) of NONE => TFree (a, S) | SOME b => TVar ((b, 0), S)); in map_proof_terms (map_types (map_type_tfree thaw)) (map_type_tfree thaw) prf end; local fun new_name ix (pairs, used) = let val v = singleton (Name.variant_list used) (string_of_indexname ix) in ((ix, v) :: pairs, v :: used) end; fun freeze_one alist (ix, sort) = (case AList.lookup (op =) alist ix of NONE => TVar (ix, sort) | SOME name => TFree (name, sort)); in fun legacy_freezeT t prf = let val used = Term.add_tfree_names t []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvars t [])) ([], used); in (case alist of [] => prf (*nothing to do!*) | _ => let val frzT = map_type_tvar (freeze_one alist) in map_proof_terms (map_types frzT) frzT prf end) end; end; (* rotate assumptions *) fun rotate_proof Bs Bi' params asms m prf = let val i = length asms; val j = length Bs; in mk_AbsP (Bs @ [Bi']) (proof_combP (prf, map PBound (j downto 1) @ [mk_Abst params (mk_AbsP asms (proof_combP (proof_combt (PBound i, map Bound ((length params - 1) downto 0)), map PBound (((i-m-1) downto 0) @ ((i-1) downto (i-m))))))])) end; (* permute premises *) fun permute_prems_proof prems' j k prf = let val n = length prems' in mk_AbsP prems' (proof_combP (prf, map PBound ((n-1 downto n-j) @ (k-1 downto 0) @ (n-j-1 downto k)))) end; (* generalization *) fun generalize_proof (tfrees, frees) idx prop prf = let val gen = if Symtab.is_empty frees then [] else fold_aterms (fn Free (x, T) => Symtab.defined frees x ? insert (op =) (x, T) | _ => I) (Term_Subst.generalize (tfrees, Symtab.empty) idx prop) []; in prf |> Same.commit (map_proof_terms_same (Term_Subst.generalize_same (tfrees, Symtab.empty) idx) (Term_Subst.generalizeT_same tfrees idx)) |> fold (fn (x, T) => forall_intr_proof (x, Free (x, T)) NONE) gen |> fold_rev (fn (x, T) => fn prf' => prf' %> Var (Name.clean_index (x, idx), T)) gen end; (* instantiation *) fun instantiate (instT, inst) = Same.commit (map_proof_terms_same (Term_Subst.instantiate_same (instT, Term_Subst.Vars.map (K remove_types) inst)) (Term_Subst.instantiateT_same instT)); (* lifting *) fun lift_proof Bi inc prop prf = let fun lift'' Us Ts t = strip_abs Ts (Logic.incr_indexes ([], Us, inc) (mk_abs Ts t)); fun lift' Us Ts (Abst (s, T, prf)) = (Abst (s, Same.map_option (Logic.incr_tvar_same inc) T, lifth' Us (dummyT::Ts) prf) handle Same.SAME => Abst (s, T, lift' Us (dummyT::Ts) prf)) | lift' Us Ts (AbsP (s, t, prf)) = (AbsP (s, Same.map_option (same (op =) (lift'' Us Ts)) t, lifth' Us Ts prf) handle Same.SAME => AbsP (s, t, lift' Us Ts prf)) | lift' Us Ts (prf % t) = (lift' Us Ts prf % Option.map (lift'' Us Ts) t handle Same.SAME => prf % Same.map_option (same (op =) (lift'' Us Ts)) t) | lift' Us Ts (prf1 %% prf2) = (lift' Us Ts prf1 %% lifth' Us Ts prf2 handle Same.SAME => prf1 %% lift' Us Ts prf2) | lift' _ _ (PAxm (s, prop, Ts)) = PAxm (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PClass (T, c)) = PClass (Logic.incr_tvar_same inc T, c) | lift' _ _ (Oracle (s, prop, Ts)) = Oracle (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PThm ({serial = i, pos = p, theory_name, name = s, prop, types = Ts}, thm_body)) = PThm (thm_header i p theory_name s prop ((Same.map_option o Same.map) (Logic.incr_tvar inc) Ts), thm_body) | lift' _ _ _ = raise Same.SAME and lifth' Us Ts prf = (lift' Us Ts prf handle Same.SAME => prf); val ps = map (Logic.lift_all inc Bi) (Logic.strip_imp_prems prop); val k = length ps; fun mk_app b (i, j, prf) = if b then (i-1, j, prf %% PBound i) else (i, j-1, prf %> Bound j); fun lift Us bs i j (Const ("Pure.imp", _) $ A $ B) = AbsP ("H", NONE (*A*), lift Us (true::bs) (i+1) j B) | lift Us bs i j (Const ("Pure.all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), lift (T::Us) (false::bs) i (j+1) t) | lift Us bs i j _ = proof_combP (lifth' (rev Us) [] prf, map (fn k => (#3 (fold_rev mk_app bs (i-1, j-1, PBound k)))) (i + k - 1 downto i)); in mk_AbsP ps (lift [] [] 0 0 Bi) end; fun incr_indexes i = Same.commit (map_proof_terms_same (Logic.incr_indexes_same ([], [], i)) (Logic.incr_tvar_same i)); (* proof by assumption *) fun mk_asm_prf t i m = let fun imp_prf _ i 0 = PBound i | imp_prf (Const ("Pure.imp", _) $ A $ B) i m = AbsP ("H", NONE (*A*), imp_prf B (i+1) (m-1)) | imp_prf _ i _ = PBound i; fun all_prf (Const ("Pure.all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), all_prf t) | all_prf t = imp_prf t (~i) m in all_prf t end; fun assumption_proof Bs Bi n prf = mk_AbsP Bs (proof_combP (prf, map PBound (length Bs - 1 downto 0) @ [mk_asm_prf Bi n ~1])); (* composition of object rule with proof state *) fun flatten_params_proof i j n (Const ("Pure.imp", _) $ A $ B, k) = AbsP ("H", NONE (*A*), flatten_params_proof (i+1) j n (B, k)) | flatten_params_proof i j n (Const ("Pure.all", _) $ Abs (a, T, t), k) = Abst (a, NONE (*T*), flatten_params_proof i (j+1) n (t, k)) | flatten_params_proof i j n (_, k) = proof_combP (proof_combt (PBound (k+i), map Bound (j-1 downto 0)), map PBound (remove (op =) (i-n) (i-1 downto 0))); fun bicompose_proof flatten Bs As A oldAs n m rprf sprf = let val lb = length Bs; val la = length As; in mk_AbsP (Bs @ As) (proof_combP (sprf, map PBound (lb + la - 1 downto la)) %% proof_combP (rprf, (if n>0 then [mk_asm_prf (the A) n m] else []) @ map (if flatten then flatten_params_proof 0 0 n else PBound o snd) (oldAs ~~ (la - 1 downto 0)))) end; (** type classes **) fun strip_shyps_proof algebra present witnessed extra prf = let val replacements = present @ witnessed @ map (`Logic.dummy_tfree) extra; fun get_replacement S = replacements |> get_first (fn (T', S') => if Sorts.sort_le algebra (S', S) then SOME T' else NONE); fun replace T = if exists (fn (T', _) => T' = T) present then raise Same.SAME else (case get_replacement (Type.sort_of_atyp T) of SOME T' => T' | NONE => raise Fail "strip_shyps_proof: bad type variable in proof term"); in Same.commit (map_proof_types_same (Term_Subst.map_atypsT_same replace)) prf end; fun of_sort_proof algebra classrel_proof arity_proof hyps = Sorts.of_sort_derivation algebra {class_relation = fn _ => fn _ => fn (prf, c1) => fn c2 => if c1 = c2 then prf else classrel_proof (c1, c2) %% prf, type_constructor = fn (a, _) => fn dom => fn c => let val Ss = map (map snd) dom and prfs = maps (map fst) dom in proof_combP (arity_proof (a, Ss, c), prfs) end, type_variable = fn typ => map (fn c => (hyps (typ, c), c)) (Type.sort_of_atyp typ)}; (** axioms and theorems **) val add_type_variables = (fold_types o fold_atyps) (insert (op =)); fun type_variables_of t = rev (add_type_variables t []); val add_variables = fold_aterms (fn a => (is_Var a orelse is_Free a) ? insert (op =) a); fun variables_of t = rev (add_variables t []); fun test_args _ [] = true | test_args is (Bound i :: ts) = not (member (op =) is i) andalso test_args (i :: is) ts | test_args _ _ = false; fun is_fun (Type ("fun", _)) = true | is_fun (TVar _) = true | is_fun _ = false; -fun vars_of t = map Var (rev (Term.add_vars t [])); +fun vars_of t = map Var (build_rev (Term.add_vars t)); fun add_funvars Ts (vs, t) = if is_fun (fastype_of1 (Ts, t)) then union (op =) vs (map_filter (fn Var (ixn, T) => if is_fun T then SOME ixn else NONE | _ => NONE) (vars_of t)) else vs; fun add_npvars q p Ts (vs, Const ("Pure.imp", _) $ t $ u) = add_npvars q p Ts (add_npvars q (not p) Ts (vs, t), u) | add_npvars q p Ts (vs, Const ("Pure.all", Type (_, [Type (_, [T, _]), _])) $ t) = add_npvars q p Ts (vs, if p andalso q then betapply (t, Var (("",0), T)) else t) | add_npvars q p Ts (vs, Abs (_, T, t)) = add_npvars q p (T::Ts) (vs, t) | add_npvars _ _ Ts (vs, t) = add_npvars' Ts (vs, t) and add_npvars' Ts (vs, t) = (case strip_comb t of (Var (ixn, _), ts) => if test_args [] ts then vs else Library.foldl (add_npvars' Ts) (AList.update (op =) (ixn, Library.foldl (add_funvars Ts) ((these ooo AList.lookup) (op =) vs ixn, ts)) vs, ts) | (Abs (_, T, u), ts) => Library.foldl (add_npvars' (T::Ts)) (vs, u :: ts) | (_, ts) => Library.foldl (add_npvars' Ts) (vs, ts)); fun prop_vars (Const ("Pure.imp", _) $ P $ Q) = union (op =) (prop_vars P) (prop_vars Q) | prop_vars (Const ("Pure.all", _) $ Abs (_, _, t)) = prop_vars t | prop_vars t = (case strip_comb t of (Var (ixn, _), _) => [ixn] | _ => []); fun is_proj t = let fun is_p i t = (case strip_comb t of (Bound _, []) => false | (Bound j, ts) => j >= i orelse exists (is_p i) ts | (Abs (_, _, u), _) => is_p (i+1) u | (_, ts) => exists (is_p i) ts) in (case strip_abs_body t of Bound _ => true | t' => is_p 0 t') end; fun prop_args prop = let val needed_vars = union (op =) (Library.foldl (uncurry (union (op =))) ([], map (uncurry (insert (op =))) (add_npvars true true [] ([], prop)))) (prop_vars prop); in variables_of prop |> map (fn var as Var (ixn, _) => if member (op =) needed_vars ixn then SOME var else NONE | free => SOME free) end; fun const_proof mk name prop = let val args = prop_args prop; val ({outer_constraints, ...}, prop1) = Logic.unconstrainT [] prop; val head = mk (name, prop1, NONE); in proof_combP (proof_combt' (head, args), map PClass outer_constraints) end; val axm_proof = const_proof PAxm; val oracle_proof = const_proof Oracle; val shrink_proof = let fun shrink ls lev (prf as Abst (a, T, body)) = let val (b, is, ch, body') = shrink ls (lev+1) body in (b, is, ch, if ch then Abst (a, T, body') else prf) end | shrink ls lev (prf as AbsP (a, t, body)) = let val (b, is, ch, body') = shrink (lev::ls) lev body in (b orelse member (op =) is 0, map_filter (fn 0 => NONE | i => SOME (i-1)) is, ch, if ch then AbsP (a, t, body') else prf) end | shrink ls lev prf = let val (is, ch, _, prf') = shrink' ls lev [] [] prf in (false, is, ch, prf') end and shrink' ls lev ts prfs (prf as prf1 %% prf2) = let val p as (_, is', ch', prf') = shrink ls lev prf2; val (is, ch, ts', prf'') = shrink' ls lev ts (p::prfs) prf1 in (union (op =) is is', ch orelse ch', ts', if ch orelse ch' then prf'' %% prf' else prf) end | shrink' ls lev ts prfs (prf as prf1 % t) = let val (is, ch, (ch', t')::ts', prf') = shrink' ls lev (t::ts) prfs prf1 in (is, ch orelse ch', ts', if ch orelse ch' then prf' % t' else prf) end | shrink' ls lev ts prfs (prf as PBound i) = (if exists (fn SOME (Bound j) => lev-j <= nth ls i | _ => true) ts orelse has_duplicates (op =) (Library.foldl (fn (js, SOME (Bound j)) => j :: js | (js, _) => js) ([], ts)) orelse exists #1 prfs then [i] else [], false, map (pair false) ts, prf) | shrink' _ _ ts _ (Hyp t) = ([], false, map (pair false) ts, Hyp t) | shrink' _ _ ts _ (prf as MinProof) = ([], false, map (pair false) ts, prf) | shrink' _ _ ts _ (prf as PClass _) = ([], false, map (pair false) ts, prf) | shrink' _ _ ts prfs prf = let val prop = (case prf of PAxm (_, prop, _) => prop | Oracle (_, prop, _) => prop | PThm ({prop, ...}, _) => prop | _ => raise Fail "shrink: proof not in normal form"); val vs = vars_of prop; val (ts', ts'') = chop (length vs) ts; val insts = take (length ts') (map (fst o dest_Var) vs) ~~ ts'; val nvs = Library.foldl (fn (ixns', (ixn, ixns)) => insert (op =) ixn (case AList.lookup (op =) insts ixn of SOME (SOME t) => if is_proj t then union (op =) ixns ixns' else ixns' | _ => union (op =) ixns ixns')) (needed prop ts'' prfs, add_npvars false true [] ([], prop)); val insts' = map (fn (ixn, x as SOME _) => if member (op =) nvs ixn then (false, x) else (true, NONE) | (_, x) => (false, x)) insts in ([], false, insts' @ map (pair false) ts'', prf) end and needed (Const ("Pure.imp", _) $ t $ u) ts ((b, _, _, _)::prfs) = union (op =) (if b then map (fst o dest_Var) (vars_of t) else []) (needed u ts prfs) | needed (Var (ixn, _)) (_::_) _ = [ixn] | needed _ _ _ = []; in fn prf => #4 (shrink [] 0 prf) end; (** axioms for equality **) val aT = TFree ("'a", []); val bT = TFree ("'b", []); val x = Free ("x", aT); val y = Free ("y", aT); val z = Free ("z", aT); val A = Free ("A", propT); val B = Free ("B", propT); val f = Free ("f", aT --> bT); val g = Free ("g", aT --> bT); val equality_axms = [("reflexive", Logic.mk_equals (x, x)), ("symmetric", Logic.mk_implies (Logic.mk_equals (x, y), Logic.mk_equals (y, x))), ("transitive", Logic.list_implies ([Logic.mk_equals (x, y), Logic.mk_equals (y, z)], Logic.mk_equals (x, z))), ("equal_intr", Logic.list_implies ([Logic.mk_implies (A, B), Logic.mk_implies (B, A)], Logic.mk_equals (A, B))), ("equal_elim", Logic.list_implies ([Logic.mk_equals (A, B), A], B)), ("abstract_rule", Logic.mk_implies (Logic.all x (Logic.mk_equals (f $ x, g $ x)), Logic.mk_equals (lambda x (f $ x), lambda x (g $ x)))), ("combination", Logic.list_implies ([Logic.mk_equals (f, g), Logic.mk_equals (x, y)], Logic.mk_equals (f $ x, g $ y)))]; val [reflexive_axm, symmetric_axm, transitive_axm, equal_intr_axm, equal_elim_axm, abstract_rule_axm, combination_axm] = map (fn (s, t) => PAxm ("Pure." ^ s, Logic.varify_global t, NONE)) equality_axms; val reflexive_proof = reflexive_axm % NONE; val is_reflexive_proof = fn PAxm ("Pure.reflexive", _, _) % _ => true | _ => false; fun symmetric_proof prf = if is_reflexive_proof prf then prf else symmetric_axm % NONE % NONE %% prf; fun transitive_proof U u prf1 prf2 = if is_reflexive_proof prf1 then prf2 else if is_reflexive_proof prf2 then prf1 else if U = propT then transitive_axm % NONE % SOME (remove_types u) % NONE %% prf1 %% prf2 else transitive_axm % NONE % NONE % NONE %% prf1 %% prf2; fun equal_intr_proof A B prf1 prf2 = equal_intr_axm %> remove_types A %> remove_types B %% prf1 %% prf2; fun equal_elim_proof A B prf1 prf2 = equal_elim_axm %> remove_types A %> remove_types B %% prf1 %% prf2; fun abstract_rule_proof (a, x) prf = abstract_rule_axm % NONE % NONE %% forall_intr_proof (a, x) NONE prf; fun check_comb (PAxm ("Pure.combination", _, _) % f % _ % _ % _ %% prf %% _) = is_some f orelse check_comb prf | check_comb (PAxm ("Pure.transitive", _, _) % _ % _ % _ %% prf1 %% prf2) = check_comb prf1 andalso check_comb prf2 | check_comb (PAxm ("Pure.symmetric", _, _) % _ % _ %% prf) = check_comb prf | check_comb _ = false; fun combination_proof f g t u prf1 prf2 = let val f = Envir.beta_norm f; val g = Envir.beta_norm g; val prf = if check_comb prf1 then combination_axm % NONE % NONE else (case prf1 of PAxm ("Pure.reflexive", _, _) % _ => combination_axm %> remove_types f % NONE | _ => combination_axm %> remove_types f %> remove_types g) in prf % (case head_of f of Abs _ => SOME (remove_types t) | Var _ => SOME (remove_types t) | _ => NONE) % (case head_of g of Abs _ => SOME (remove_types u) | Var _ => SOME (remove_types u) | _ => NONE) %% prf1 %% prf2 end; (** rewriting on proof terms **) (* simple first order matching functions for terms and proofs (see pattern.ML) *) exception PMatch; fun flt (i: int) = filter (fn n => n < i); fun fomatch Ts tymatch j instsp p = let fun mtch (instsp as (tyinsts, insts)) = fn (Var (ixn, T), t) => if j>0 andalso not (null (flt j (loose_bnos t))) then raise PMatch else (tymatch (tyinsts, fn () => (T, fastype_of1 (Ts, t))), (ixn, t) :: insts) | (Free (a, T), Free (b, U)) => if a=b then (tymatch (tyinsts, K (T, U)), insts) else raise PMatch | (Const (a, T), Const (b, U)) => if a=b then (tymatch (tyinsts, K (T, U)), insts) else raise PMatch | (f $ t, g $ u) => mtch (mtch instsp (f, g)) (t, u) | (Bound i, Bound j) => if i=j then instsp else raise PMatch | _ => raise PMatch in mtch instsp (apply2 Envir.beta_eta_contract p) end; fun match_proof Ts tymatch = let fun optmatch _ inst (NONE, _) = inst | optmatch _ _ (SOME _, NONE) = raise PMatch | optmatch mtch inst (SOME x, SOME y) = mtch inst (x, y) fun matcht Ts j (pinst, tinst) (t, u) = (pinst, fomatch Ts tymatch j tinst (t, Envir.beta_norm u)); fun matchT (pinst, (tyinsts, insts)) p = (pinst, (tymatch (tyinsts, K p), insts)); fun matchTs inst (Ts, Us) = Library.foldl (uncurry matchT) (inst, Ts ~~ Us); fun mtch Ts i j (pinst, tinst) (Hyp (Var (ixn, _)), prf) = if i = 0 andalso j = 0 then ((ixn, prf) :: pinst, tinst) else (case apfst (flt i) (apsnd (flt j) (prf_add_loose_bnos 0 0 prf ([], []))) of ([], []) => ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst) | ([], _) => if j = 0 then ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst) else raise PMatch | _ => raise PMatch) | mtch Ts i j inst (prf1 % opt1, prf2 % opt2) = optmatch (matcht Ts j) (mtch Ts i j inst (prf1, prf2)) (opt1, opt2) | mtch Ts i j inst (prf1 %% prf2, prf1' %% prf2') = mtch Ts i j (mtch Ts i j inst (prf1, prf1')) (prf2, prf2') | mtch Ts i j inst (Abst (_, opT, prf1), Abst (_, opU, prf2)) = mtch (the_default dummyT opU :: Ts) i (j+1) (optmatch matchT inst (opT, opU)) (prf1, prf2) | mtch Ts i j inst (prf1, Abst (_, opU, prf2)) = mtch (the_default dummyT opU :: Ts) i (j+1) inst (incr_pboundvars 0 1 prf1 %> Bound 0, prf2) | mtch Ts i j inst (AbsP (_, opt, prf1), AbsP (_, opu, prf2)) = mtch Ts (i+1) j (optmatch (matcht Ts j) inst (opt, opu)) (prf1, prf2) | mtch Ts i j inst (prf1, AbsP (_, _, prf2)) = mtch Ts (i+1) j inst (incr_pboundvars 1 0 prf1 %% PBound 0, prf2) | mtch Ts i j inst (PAxm (s1, _, opTs), PAxm (s2, _, opUs)) = if s1 = s2 then optmatch matchTs inst (opTs, opUs) else raise PMatch | mtch Ts i j inst (PClass (T1, c1), PClass (T2, c2)) = if c1 = c2 then matchT inst (T1, T2) else raise PMatch | mtch Ts i j inst (PThm ({name = name1, prop = prop1, types = types1, ...}, _), PThm ({name = name2, prop = prop2, types = types2, ...}, _)) = if name1 = name2 andalso prop1 = prop2 then optmatch matchTs inst (types1, types2) else raise PMatch | mtch _ _ _ inst (PBound i, PBound j) = if i = j then inst else raise PMatch | mtch _ _ _ _ _ = raise PMatch in mtch Ts 0 0 end; fun prf_subst (pinst, (tyinsts, insts)) = let val substT = Envir.subst_type_same tyinsts; val substTs = Same.map substT; fun subst' lev (Var (xi, _)) = (case AList.lookup (op =) insts xi of NONE => raise Same.SAME | SOME u => incr_boundvars lev u) | subst' _ (Const (s, T)) = Const (s, substT T) | subst' _ (Free (s, T)) = Free (s, substT T) | subst' lev (Abs (a, T, body)) = (Abs (a, substT T, Same.commit (subst' (lev + 1)) body) handle Same.SAME => Abs (a, T, subst' (lev + 1) body)) | subst' lev (f $ t) = (subst' lev f $ Same.commit (subst' lev) t handle Same.SAME => f $ subst' lev t) | subst' _ _ = raise Same.SAME; fun subst plev tlev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (subst' tlev) t, Same.commit (subst (plev + 1) tlev) body) handle Same.SAME => AbsP (a, t, subst (plev + 1) tlev body)) | subst plev tlev (Abst (a, T, body)) = (Abst (a, Same.map_option substT T, Same.commit (subst plev (tlev + 1)) body) handle Same.SAME => Abst (a, T, subst plev (tlev + 1) body)) | subst plev tlev (prf %% prf') = (subst plev tlev prf %% Same.commit (subst plev tlev) prf' handle Same.SAME => prf %% subst plev tlev prf') | subst plev tlev (prf % t) = (subst plev tlev prf % Same.commit (Same.map_option (subst' tlev)) t handle Same.SAME => prf % Same.map_option (subst' tlev) t) | subst plev tlev (Hyp (Var (xi, _))) = (case AList.lookup (op =) pinst xi of NONE => raise Same.SAME | SOME prf' => incr_pboundvars plev tlev prf') | subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Same.map_option substTs Ts) | subst _ _ (PClass (T, c)) = PClass (substT T, c) | subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Same.map_option substTs Ts) | subst _ _ (PThm ({serial = i, pos = p, theory_name, name = id, prop, types}, thm_body)) = PThm (thm_header i p theory_name id prop (Same.map_option substTs types), thm_body) | subst _ _ _ = raise Same.SAME; in fn t => subst 0 0 t handle Same.SAME => t end; (*A fast unification filter: true unless the two terms cannot be unified. Terms must be NORMAL. Treats all Vars as distinct. *) fun could_unify prf1 prf2 = let fun matchrands (prf1 %% prf2) (prf1' %% prf2') = could_unify prf2 prf2' andalso matchrands prf1 prf1' | matchrands (prf % SOME t) (prf' % SOME t') = Term.could_unify (t, t') andalso matchrands prf prf' | matchrands (prf % _) (prf' % _) = matchrands prf prf' | matchrands _ _ = true fun head_of (prf %% _) = head_of prf | head_of (prf % _) = head_of prf | head_of prf = prf in case (head_of prf1, head_of prf2) of (_, Hyp (Var _)) => true | (Hyp (Var _), _) => true | (PAxm (a, _, _), PAxm (b, _, _)) => a = b andalso matchrands prf1 prf2 | (PClass (_, c), PClass (_, d)) => c = d andalso matchrands prf1 prf2 | (PThm ({name = a, prop = propa, ...}, _), PThm ({name = b, prop = propb, ...}, _)) => a = b andalso propa = propb andalso matchrands prf1 prf2 | (PBound i, PBound j) => i = j andalso matchrands prf1 prf2 | (AbsP _, _) => true (*because of possible eta equality*) | (Abst _, _) => true | (_, AbsP _) => true | (_, Abst _) => true | _ => false end; (* rewrite proof *) val no_skel = PBound 0; val normal_skel = Hyp (Var ((Name.uu, 0), propT)); fun rewrite_prf tymatch (rules, procs) prf = let fun rew _ _ (Abst (_, _, body) % SOME t) = SOME (prf_subst_bounds [t] body, no_skel) | rew _ _ (AbsP (_, _, body) %% prf) = SOME (prf_subst_pbounds [prf] body, no_skel) | rew Ts hs prf = (case get_first (fn r => r Ts hs prf) procs of NONE => get_first (fn (prf1, prf2) => SOME (prf_subst (match_proof Ts tymatch ([], (Vartab.empty, [])) (prf1, prf)) prf2, prf2) handle PMatch => NONE) (filter (could_unify prf o fst) rules) | some => some); fun rew0 Ts hs (prf as AbsP (_, _, prf' %% PBound 0)) = if prf_loose_Pbvar1 prf' 0 then rew Ts hs prf else let val prf'' = incr_pboundvars (~1) 0 prf' in SOME (the_default (prf'', no_skel) (rew Ts hs prf'')) end | rew0 Ts hs (prf as Abst (_, _, prf' % SOME (Bound 0))) = if prf_loose_bvar1 prf' 0 then rew Ts hs prf else let val prf'' = incr_pboundvars 0 (~1) prf' in SOME (the_default (prf'', no_skel) (rew Ts hs prf'')) end | rew0 Ts hs prf = rew Ts hs prf; fun rew1 _ _ (Hyp (Var _)) _ = NONE | rew1 Ts hs skel prf = (case rew2 Ts hs skel prf of SOME prf1 => (case rew0 Ts hs prf1 of SOME (prf2, skel') => SOME (the_default prf2 (rew1 Ts hs skel' prf2)) | NONE => SOME prf1) | NONE => (case rew0 Ts hs prf of SOME (prf1, skel') => SOME (the_default prf1 (rew1 Ts hs skel' prf1)) | NONE => NONE)) and rew2 Ts hs skel (prf % SOME t) = (case prf of Abst (_, _, body) => let val prf' = prf_subst_bounds [t] body in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end | _ => (case rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf of SOME prf' => SOME (prf' % SOME t) | NONE => NONE)) | rew2 Ts hs skel (prf % NONE) = Option.map (fn prf' => prf' % NONE) (rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf) | rew2 Ts hs skel (prf1 %% prf2) = (case prf1 of AbsP (_, _, body) => let val prf' = prf_subst_pbounds [prf2] body in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end | _ => let val (skel1, skel2) = (case skel of skel1 %% skel2 => (skel1, skel2) | _ => (no_skel, no_skel)) in (case rew1 Ts hs skel1 prf1 of SOME prf1' => (case rew1 Ts hs skel2 prf2 of SOME prf2' => SOME (prf1' %% prf2') | NONE => SOME (prf1' %% prf2)) | NONE => (case rew1 Ts hs skel2 prf2 of SOME prf2' => SOME (prf1 %% prf2') | NONE => NONE)) end) | rew2 Ts hs skel (Abst (s, T, prf)) = (case rew1 (the_default dummyT T :: Ts) hs (case skel of Abst (_, _, skel') => skel' | _ => no_skel) prf of SOME prf' => SOME (Abst (s, T, prf')) | NONE => NONE) | rew2 Ts hs skel (AbsP (s, t, prf)) = (case rew1 Ts (t :: hs) (case skel of AbsP (_, _, skel') => skel' | _ => no_skel) prf of SOME prf' => SOME (AbsP (s, t, prf')) | NONE => NONE) | rew2 _ _ _ _ = NONE; in the_default prf (rew1 [] [] no_skel prf) end; fun rewrite_proof thy = rewrite_prf (fn (tyenv, f) => Sign.typ_match thy (f ()) tyenv handle Type.TYPE_MATCH => raise PMatch); fun rewrite_proof_notypes rews = rewrite_prf fst rews; (* theory data *) structure Data = Theory_Data ( type T = ((stamp * (proof * proof)) list * (stamp * (typ list -> term option list -> proof -> (proof * proof) option)) list) * (theory -> proof -> proof) option; val empty = (([], []), NONE); val extend = I; fun merge (((rules1, procs1), preproc1), ((rules2, procs2), preproc2)) : T = ((AList.merge (op =) (K true) (rules1, rules2), AList.merge (op =) (K true) (procs1, procs2)), merge_options (preproc1, preproc2)); ); fun get_rew_data thy = let val (rules, procs) = #1 (Data.get thy) in (map #2 rules, map #2 procs) end; fun rew_proof thy = rewrite_prf fst (get_rew_data thy); fun add_prf_rrule r = (Data.map o apfst o apfst) (cons (stamp (), r)); fun add_prf_rproc p = (Data.map o apfst o apsnd) (cons (stamp (), p)); fun set_preproc f = (Data.map o apsnd) (K (SOME f)); fun apply_preproc thy = (case #2 (Data.get thy) of NONE => I | SOME f => f thy); (** reconstruction of partial proof terms **) fun forall_intr_variables_term prop = fold_rev Logic.all (variables_of prop) prop; fun forall_intr_variables prop prf = fold_rev forall_intr_proof' (variables_of prop) prf; local fun app_types shift prop Ts prf = let val inst = type_variables_of prop ~~ Ts; fun subst_same A = (case AList.lookup (op =) inst A of SOME T => T | NONE => raise Same.SAME); val subst_type_same = Term_Subst.map_atypsT_same (fn TVar ((a, i), S) => subst_same (TVar ((a, i - shift), S)) | A => subst_same A); in Same.commit (map_proof_types_same subst_type_same) prf end; fun guess_name (PThm ({name, ...}, _)) = name | guess_name (prf %% Hyp _) = guess_name prf | guess_name (prf %% PClass _) = guess_name prf | guess_name (prf % NONE) = guess_name prf | guess_name (prf % SOME (Var _)) = guess_name prf | guess_name _ = ""; (* generate constraints for proof term *) fun mk_var env Ts T = let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T) in (list_comb (v, map Bound (length Ts - 1 downto 0)), env') end; fun mk_tvar S (Envir.Envir {maxidx, tenv, tyenv}) = (TVar (("'t", maxidx + 1), S), Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}); val mk_abs = fold (fn T => fn u => Abs ("", T, u)); fun unifyT thy env T U = let val Envir.Envir {maxidx, tenv, tyenv} = env; val (tyenv', maxidx') = Sign.typ_unify thy (T, U) (tyenv, maxidx); in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end; fun chaseT env (T as TVar v) = (case Type.lookup (Envir.type_env env) v of NONE => T | SOME T' => chaseT env T') | chaseT _ T = T; fun infer_type thy (env as Envir.Envir {maxidx, tenv, tyenv}) _ vTs (t as Const (s, T)) = if T = dummyT then (case Sign.const_type thy s of NONE => error ("reconstruct_proof: No such constant: " ^ quote s) | SOME T => let val T' = Type.strip_sorts (Logic.incr_tvar (maxidx + 1) T) in (Const (s, T'), T', vTs, Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}) end) else (t, T, vTs, env) | infer_type _ env _ vTs (t as Free (s, T)) = if T = dummyT then (case Symtab.lookup vTs s of NONE => let val (T, env') = mk_tvar [] env in (Free (s, T), T, Symtab.update_new (s, T) vTs, env') end | SOME T => (Free (s, T), T, vTs, env)) else (t, T, vTs, env) | infer_type _ _ _ _ (Var _) = error "reconstruct_proof: internal error" | infer_type thy env Ts vTs (Abs (s, T, t)) = let val (T', env') = if T = dummyT then mk_tvar [] env else (T, env); val (t', U, vTs', env'') = infer_type thy env' (T' :: Ts) vTs t in (Abs (s, T', t'), T' --> U, vTs', env'') end | infer_type thy env Ts vTs (t $ u) = let val (t', T, vTs1, env1) = infer_type thy env Ts vTs t; val (u', U, vTs2, env2) = infer_type thy env1 Ts vTs1 u; in (case chaseT env2 T of Type ("fun", [U', V]) => (t' $ u', V, vTs2, unifyT thy env2 U U') | _ => let val (V, env3) = mk_tvar [] env2 in (t' $ u', V, vTs2, unifyT thy env3 T (U --> V)) end) end | infer_type _ env Ts vTs (t as Bound i) = ((t, nth Ts i, vTs, env) handle General.Subscript => error ("infer_type: bad variable index " ^ string_of_int i)); fun cantunify thy (t, u) = error ("Non-unifiable terms:\n" ^ Syntax.string_of_term_global thy t ^ "\n\n" ^ Syntax.string_of_term_global thy u); fun decompose thy Ts (p as (t, u)) env = let fun rigrig (a, T) (b, U) uT ts us = if a <> b then cantunify thy p else apfst flat (fold_map (decompose thy Ts) (ts ~~ us) (uT env T U)) in case apply2 (strip_comb o Envir.head_norm env) p of ((Const c, ts), (Const d, us)) => rigrig c d (unifyT thy) ts us | ((Free c, ts), (Free d, us)) => rigrig c d (unifyT thy) ts us | ((Bound i, ts), (Bound j, us)) => rigrig (i, dummyT) (j, dummyT) (K o K) ts us | ((Abs (_, T, t), []), (Abs (_, U, u), [])) => decompose thy (T::Ts) (t, u) (unifyT thy env T U) | ((Abs (_, T, t), []), _) => decompose thy (T::Ts) (t, incr_boundvars 1 u $ Bound 0) env | (_, (Abs (_, T, u), [])) => decompose thy (T::Ts) (incr_boundvars 1 t $ Bound 0, u) env | _ => ([(mk_abs Ts t, mk_abs Ts u)], env) end; fun make_constraints_cprf thy env cprf = let fun add_cnstrt Ts prop prf cs env vTs (t, u) = let val t' = mk_abs Ts t; val u' = mk_abs Ts u in (prop, prf, cs, Pattern.unify (Context.Theory thy) (t', u') env, vTs) handle Pattern.Pattern => let val (cs', env') = decompose thy [] (t', u') env in (prop, prf, cs @ cs', env', vTs) end | Pattern.Unif => cantunify thy (Envir.norm_term env t', Envir.norm_term env u') end; fun mk_cnstrts_atom env vTs prop opTs prf = let val prop_types = type_variables_of prop; val (Ts, env') = (case opTs of NONE => fold_map (mk_tvar o Type.sort_of_atyp) prop_types env | SOME Ts => (Ts, env)); val prop' = subst_atomic_types (prop_types ~~ Ts) (forall_intr_variables_term prop) handle ListPair.UnequalLengths => error ("Wrong number of type arguments for " ^ quote (guess_name prf)) in (prop', change_types (SOME Ts) prf, [], env', vTs) end; fun head_norm (prop, prf, cnstrts, env, vTs) = (Envir.head_norm env prop, prf, cnstrts, env, vTs); fun mk_cnstrts env _ Hs vTs (PBound i) = ((nth Hs i, PBound i, [], env, vTs) handle General.Subscript => error ("mk_cnstrts: bad variable index " ^ string_of_int i)) | mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) = let val (T, env') = (case opT of NONE => mk_tvar [] env | SOME T => (T, env)); val (t, prf, cnstrts, env'', vTs') = mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf; in (Const ("Pure.all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf), cnstrts, env'', vTs') end | mk_cnstrts env Ts Hs vTs (AbsP (s, SOME t, cprf)) = let val (t', _, vTs', env') = infer_type thy env Ts vTs t; val (u, prf, cnstrts, env'', vTs'') = mk_cnstrts env' Ts (t'::Hs) vTs' cprf; in (Logic.mk_implies (t', u), AbsP (s, SOME t', prf), cnstrts, env'', vTs'') end | mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) = let val (t, env') = mk_var env Ts propT; val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf; in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs') end | mk_cnstrts env Ts Hs vTs (cprf1 %% cprf2) = let val (u, prf2, cnstrts, env', vTs') = mk_cnstrts env Ts Hs vTs cprf2 in (case head_norm (mk_cnstrts env' Ts Hs vTs' cprf1) of (Const ("Pure.imp", _) $ u' $ t', prf1, cnstrts', env'', vTs'') => add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts) env'' vTs'' (u, u') | (t, prf1, cnstrts', env'', vTs'') => let val (v, env''') = mk_var env'' Ts propT in add_cnstrt Ts v (prf1 %% prf2) (cnstrts' @ cnstrts) env''' vTs'' (t, Logic.mk_implies (u, v)) end) end | mk_cnstrts env Ts Hs vTs (cprf % SOME t) = let val (t', U, vTs1, env1) = infer_type thy env Ts vTs t in (case head_norm (mk_cnstrts env1 Ts Hs vTs1 cprf) of (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f, prf, cnstrts, env2, vTs2) => let val env3 = unifyT thy env2 T U in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2) end | (u, prf, cnstrts, env2, vTs2) => let val (v, env3) = mk_var env2 Ts (U --> propT); in add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2 (u, Const ("Pure.all", (U --> propT) --> propT) $ v) end) end | mk_cnstrts env Ts Hs vTs (cprf % NONE) = (case head_norm (mk_cnstrts env Ts Hs vTs cprf) of (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f, prf, cnstrts, env', vTs') => let val (t, env'') = mk_var env' Ts T in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs') end | (u, prf, cnstrts, env', vTs') => let val (T, env1) = mk_tvar [] env'; val (v, env2) = mk_var env1 Ts (T --> propT); val (t, env3) = mk_var env2 Ts T in add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs' (u, Const ("Pure.all", (T --> propT) --> propT) $ v) end) | mk_cnstrts env _ _ vTs (prf as PThm ({prop, types = opTs, ...}, _)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PClass (T, c)) = mk_cnstrts_atom env vTs (Logic.mk_of_class (T, c)) NONE prf | mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs) | mk_cnstrts _ _ _ _ MinProof = raise MIN_PROOF () in mk_cnstrts env [] [] Symtab.empty cprf end; (* update list of free variables of constraints *) fun upd_constrs env cs = let val tenv = Envir.term_env env; val tyenv = Envir.type_env env; val dom = [] |> Vartab.fold (cons o #1) tenv |> Vartab.fold (cons o #1) tyenv; val vran = [] |> Vartab.fold (Term.add_var_names o #2 o #2) tenv |> Vartab.fold (Term.add_tvar_namesT o #2 o #2) tyenv; fun check_cs [] = [] | check_cs ((u, p, vs) :: ps) = let val vs' = subtract (op =) dom vs in if vs = vs' then (u, p, vs) :: check_cs ps else (true, p, fold (insert op =) vs' vran) :: check_cs ps end; in check_cs cs end; (* solution of constraints *) fun solve _ [] bigenv = bigenv | solve thy cs bigenv = let fun search _ [] = error ("Unsolvable constraints:\n" ^ Pretty.string_of (Pretty.chunks (map (fn (_, p, _) => Syntax.pretty_flexpair (Syntax.init_pretty_global thy) (apply2 (Envir.norm_term bigenv) p)) cs))) | search env ((u, p as (t1, t2), vs)::ps) = if u then let val tn1 = Envir.norm_term bigenv t1; val tn2 = Envir.norm_term bigenv t2 in if Pattern.pattern tn1 andalso Pattern.pattern tn2 then (Pattern.unify (Context.Theory thy) (tn1, tn2) env, ps) handle Pattern.Unif => cantunify thy (tn1, tn2) else let val (cs', env') = decompose thy [] (tn1, tn2) env in if cs' = [(tn1, tn2)] then apsnd (cons (false, (tn1, tn2), vs)) (search env ps) else search env' (map (fn q => (true, q, vs)) cs' @ ps) end end else apsnd (cons (false, p, vs)) (search env ps); val Envir.Envir {maxidx, ...} = bigenv; val (env, cs') = search (Envir.empty maxidx) cs; in solve thy (upd_constrs env cs') (Envir.merge (bigenv, env)) end; in (* reconstruction of proofs *) fun reconstruct_proof thy prop cprf = let val (cprf' % SOME prop', thawf) = freeze_thaw_prf (cprf % SOME prop); val (t, prf, cs, env, _) = make_constraints_cprf thy (Envir.empty (maxidx_proof cprf ~1)) cprf'; val cs' = map (apply2 (Envir.norm_term env)) ((t, prop') :: cs) |> map (fn p => (true, p, Term.add_var_names (#1 p) (Term.add_var_names (#2 p) []))); val env' = solve thy cs' env; in thawf (norm_proof env' prf) end handle MIN_PROOF () => MinProof; fun prop_of_atom prop Ts = subst_atomic_types (type_variables_of prop ~~ Ts) (forall_intr_variables_term prop); val head_norm = Envir.head_norm Envir.init; fun prop_of0 Hs (PBound i) = nth Hs i | prop_of0 Hs (Abst (s, SOME T, prf)) = Logic.all_const T $ (Abs (s, T, prop_of0 Hs prf)) | prop_of0 Hs (AbsP (_, SOME t, prf)) = Logic.mk_implies (t, prop_of0 (t :: Hs) prf) | prop_of0 Hs (prf % SOME t) = (case head_norm (prop_of0 Hs prf) of Const ("Pure.all", _) $ f => f $ t | _ => error "prop_of: all expected") | prop_of0 Hs (prf1 %% _) = (case head_norm (prop_of0 Hs prf1) of Const ("Pure.imp", _) $ _ $ Q => Q | _ => error "prop_of: ==> expected") | prop_of0 _ (Hyp t) = t | prop_of0 _ (PThm ({prop, types = SOME Ts, ...}, _)) = prop_of_atom prop Ts | prop_of0 _ (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ (PClass (T, c)) = Logic.mk_of_class (T, c) | prop_of0 _ (Oracle (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ _ = error "prop_of: partial proof object"; val prop_of' = Envir.beta_eta_contract oo prop_of0; val prop_of = prop_of' []; (* expand and reconstruct subproofs *) fun expand_name_empty (header: thm_header) = if #name header = "" then SOME "" else NONE; fun expand_proof thy expand_name prf = let fun expand seen maxidx (AbsP (s, t, prf)) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', AbsP (s, t, prf')) end | expand seen maxidx (Abst (s, T, prf)) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', Abst (s, T, prf')) end | expand seen maxidx (prf1 %% prf2) = let val (seen', maxidx', prf1') = expand seen maxidx prf1; val (seen'', maxidx'', prf2') = expand seen' maxidx' prf2; in (seen'', maxidx'', prf1' %% prf2') end | expand seen maxidx (prf % t) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', prf' % t) end | expand seen maxidx (prf as PThm (header, thm_body)) = let val {serial, pos, theory_name, name, prop, types} = header in (case expand_name header of SOME name' => if name' = "" andalso is_some types then let val (seen', maxidx', prf') = (case Inttab.lookup seen serial of NONE => let val prf1 = thm_body_proof_open thm_body |> reconstruct_proof thy prop |> forall_intr_variables prop; val (seen1, maxidx1, prf2) = expand_init seen prf1 val seen2 = seen1 |> Inttab.update (serial, (maxidx1, prf2)); in (seen2, maxidx1, prf2) end | SOME (maxidx1, prf1) => (seen, maxidx1, prf1)); val prf'' = prf' |> incr_indexes (maxidx + 1) |> app_types (maxidx + 1) prop (the types); in (seen', maxidx' + maxidx + 1, prf'') end else if name' <> name then (seen, maxidx, PThm (thm_header serial pos theory_name name' prop types, thm_body)) else (seen, maxidx, prf) | NONE => (seen, maxidx, prf)) end | expand seen maxidx prf = (seen, maxidx, prf) and expand_init seen prf = expand seen (maxidx_proof prf ~1) prf; in #3 (expand_init Inttab.empty prf) end; end; (** promises **) fun fulfill_norm_proof thy ps body0 = let val _ = consolidate_bodies (map #2 ps @ [body0]); val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0; val oracles = unions_oracles (fold (fn (_, PBody {oracles, ...}) => not (null oracles) ? cons oracles) ps [oracles0]); val thms = unions_thms (fold (fn (_, PBody {thms, ...}) => not (null thms) ? cons thms) ps [thms0]); val proof = rew_proof thy proof0; in PBody {oracles = oracles, thms = thms, proof = proof} end; fun fulfill_proof_future thy promises (postproc: proof_body -> proof_body) body = let fun fulfill () = postproc (fulfill_norm_proof thy (map (apsnd Future.join) promises) (Future.join body)); in if null promises then Future.map postproc body else if Future.is_finished body andalso length promises = 1 then Future.map (fn _ => fulfill ()) (snd (hd promises)) else (singleton o Future.forks) {name = "Proofterm.fulfill_proof_future", group = NONE, deps = Future.task_of body :: map (Future.task_of o snd) promises, pri = 1, interrupts = true} fulfill end; (** theorems **) (* standardization of variables for export: only frees and named bounds *) local val declare_names_term = Term.declare_term_frees; val declare_names_term' = fn SOME t => declare_names_term t | NONE => I; val declare_names_proof = fold_proof_terms declare_names_term; fun variant names bs x = #1 (Name.variant x (fold Name.declare bs names)); fun variant_term bs (Abs (x, T, t)) = let val x' = variant (declare_names_term t Name.context) bs x; val t' = variant_term (x' :: bs) t; in Abs (x', T, t') end | variant_term bs (t $ u) = variant_term bs t $ variant_term bs u | variant_term _ t = t; fun variant_proof bs (Abst (x, T, prf)) = let val x' = variant (declare_names_proof prf Name.context) bs x; val prf' = variant_proof (x' :: bs) prf; in Abst (x', T, prf') end | variant_proof bs (AbsP (x, t, prf)) = let val x' = variant (declare_names_term' t (declare_names_proof prf Name.context)) bs x; val t' = Option.map (variant_term bs) t; val prf' = variant_proof (x' :: bs) prf; in AbsP (x', t', prf') end | variant_proof bs (prf % t) = variant_proof bs prf % Option.map (variant_term bs) t | variant_proof bs (prf1 %% prf2) = variant_proof bs prf1 %% variant_proof bs prf2 | variant_proof bs (Hyp t) = Hyp (variant_term bs t) | variant_proof _ prf = prf; val used_frees_type = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I); fun used_frees_term t = fold_types used_frees_type t #> Term.declare_term_frees t; val used_frees_proof = fold_proof_terms_types used_frees_term used_frees_type; val unvarifyT = Term.map_atyps (fn TVar ((a, _), S) => TFree (a, S) | T => T); val unvarify = Term.map_aterms (fn Var ((x, _), T) => Free (x, T) | t => t) #> map_types unvarifyT; val unvarify_proof = map_proof_terms unvarify unvarifyT; fun hidden_types prop proof = let val visible = (fold_types o fold_atyps) (insert (op =)) prop []; val add_hiddenT = fold_atyps (fn T => not (member (op =) visible T) ? insert (op =) T); in rev (fold_proof_terms_types (fold_types add_hiddenT) add_hiddenT proof []) end; fun standard_hidden_types term proof = let val hidden = hidden_types term proof; val idx = Term.maxidx_term term (maxidx_proof proof ~1) + 1; fun smash T = if member (op =) hidden T then (case Type.sort_of_atyp T of [] => dummyT | S => TVar (("'", idx), S)) else T; val smashT = map_atyps smash; in map_proof_terms (map_types smashT) smashT proof end; fun standard_hidden_terms term proof = let fun add_unless excluded x = ((is_Free x orelse is_Var x) andalso not (member (op =) excluded x)) ? insert (op =) x; val visible = fold_aterms (add_unless []) term []; val hidden = fold_proof_terms (fold_aterms (add_unless visible)) proof []; val dummy_term = Term.map_aterms (fn x => if member (op =) hidden x then Term.dummy_pattern (Term.fastype_of x) else x); in proof |> not (null hidden) ? map_proof_terms dummy_term I end; in fun standard_vars used (term, opt_proof) = let val proofs = opt_proof |> Option.map (standard_hidden_types term #> standard_hidden_terms term) |> the_list; val proof_terms = rev (fold (fold_proof_terms_types cons (cons o Logic.mk_type)) proofs []); val used_frees = used |> used_frees_term term |> fold used_frees_proof proofs; val inst = Term_Subst.zero_var_indexes_inst used_frees (term :: proof_terms); val term' = term |> Term_Subst.instantiate inst |> unvarify |> variant_term []; val proofs' = proofs |> map (instantiate inst #> unvarify_proof #> variant_proof []); in (term', try hd proofs') end; fun standard_vars_term used t = #1 (standard_vars used (t, NONE)); val add_standard_vars_term = fold_aterms (fn Free (x, T) => (fn env => (case AList.lookup (op =) env x of NONE => (x, T) :: env | SOME T' => if T = T' then env else raise TYPE ("standard_vars_env: type conflict for variable " ^ quote x, [T, T'], []))) | _ => I); val add_standard_vars = fold_proof_terms add_standard_vars_term; end; (* PThm nodes *) fun prune_body body = if Options.default_bool "prune_proofs" then (Future.map o map_proof_of) (K MinProof) body else body; fun export_enabled () = Options.default_bool "export_proofs"; fun export_standard_enabled () = Options.default_bool "export_standard_proofs"; fun export_proof_boxes_required thy = Context.theory_name thy = Context.PureN orelse (export_enabled () andalso not (export_standard_enabled ())); fun export_proof_boxes bodies = let fun export_thm (i, thm_node) boxes = if Inttab.defined boxes i then boxes else boxes |> Inttab.update (i, thm_node_export thm_node) |> fold export_thm (thm_node_thms thm_node); fun export_body (PBody {thms, ...}) = fold export_thm thms; - val exports = (bodies, Inttab.empty) |-> fold export_body |> Inttab.dest; + val exports = Inttab.build (fold export_body bodies) |> Inttab.dest; in List.app (Lazy.force o #2) exports end; local fun unconstrainT_proof algebra classrel_proof arity_proof (ucontext: Logic.unconstrain_context) = let fun hyp_map hyp = (case AList.lookup (op =) (#constraints ucontext) hyp of SOME t => Hyp t | NONE => raise Fail "unconstrainT_proof: missing constraint"); val typ = Term_Subst.map_atypsT_same (Type.strip_sorts o #atyp_map ucontext); fun ofclass (ty, c) = let val ty' = Term.map_atyps (#atyp_map ucontext) ty; in the_single (of_sort_proof algebra classrel_proof arity_proof hyp_map (ty', [c])) end; in Same.commit (map_proof_same (Term_Subst.map_types_same typ) typ ofclass) #> fold_rev (implies_intr_proof o snd) (#constraints ucontext) end; fun export_proof thy i prop prf0 = let val prf = prf0 |> reconstruct_proof thy prop |> apply_preproc thy; val (prop', SOME prf') = (prop, SOME prf) |> standard_vars Name.context; val args = [] |> add_standard_vars_term prop' |> add_standard_vars prf' |> rev; val typargs = [] |> Term.add_tfrees prop' |> fold_proof_terms Term.add_tfrees prf' |> rev; val consts = Sign.consts_of thy; val xml = (typargs, (args, (prop', no_thm_names prf'))) |> let open XML.Encode Term_XML.Encode; val encode_vars = list (pair string typ); val encode_term = encode_standard_term consts; val encode_proof = encode_standard_proof consts; in pair (list (pair string sort)) (pair encode_vars (pair encode_term encode_proof)) end; in Export.export_params {theory = thy, binding = Path.binding0 (Path.make ["proofs", string_of_int i]), executable = false, compress = true, strict = false} xml end; fun prepare_thm_proof unconstrain thy classrel_proof arity_proof (name, pos) shyps hyps concl promises body = let val named = name <> ""; val prop = Logic.list_implies (hyps, concl); val args = prop_args prop; val (ucontext, prop1) = Logic.unconstrainT shyps prop; val PBody {oracles = oracles0, thms = thms0, proof = prf} = body; val body0 = Future.value (PBody {oracles = oracles0, thms = thms0, proof = if proofs_enabled () then fold_rev implies_intr_proof hyps prf else MinProof}); fun new_prf () = let val i = serial (); val unconstrainT = unconstrainT_proof (Sign.classes_of thy) classrel_proof arity_proof ucontext; val postproc = map_proof_of (unconstrainT #> named ? rew_proof thy); in (i, fulfill_proof_future thy promises postproc body0) end; val (i, body') = (*somewhat non-deterministic proof boxes!*) if export_enabled () then new_prf () else (case strip_combt (fst (strip_combP prf)) of (PThm ({serial = ser, name = a, prop = prop', types = NONE, ...}, thm_body'), args') => if (a = "" orelse a = name) andalso prop' = prop1 andalso args' = args then let val Thm_Body {body = body', ...} = thm_body'; val i = if a = "" andalso named then serial () else ser; in (i, body' |> ser <> i ? Future.map (map_proof_of (rew_proof thy))) end else new_prf () | _ => new_prf ()); val open_proof = not named ? rew_proof thy; val export = if export_enabled () then Lazy.lazy (fn () => join_proof body' |> open_proof |> export_proof thy i prop1 handle exn => if Exn.is_interrupt exn then raise Fail ("Interrupt: potential resource problems while exporting proof " ^ string_of_int i) else Exn.reraise exn) else no_export; val thm_body = prune_body body'; val theory_name = Context.theory_long_name thy; val thm = (i, make_thm_node theory_name name prop1 thm_body export); val header = thm_header i ([pos, Position.thread_data ()]) theory_name name prop1 NONE; val head = PThm (header, Thm_Body {open_proof = open_proof, body = thm_body}); val proof = if unconstrain then proof_combt' (head, (map o Option.map o Term.map_types) (#map_atyps ucontext) args) else proof_combP (proof_combt' (head, args), map PClass (#outer_constraints ucontext) @ map Hyp hyps); in (thm, proof) end; in fun thm_proof thy = prepare_thm_proof false thy; fun unconstrain_thm_proof thy classrel_proof arity_proof shyps concl promises body = prepare_thm_proof true thy classrel_proof arity_proof ("", Position.none) shyps [] concl promises body; end; (* PThm identity *) fun get_identity shyps hyps prop prf = let val (_, prop) = Logic.unconstrainT shyps (Logic.list_implies (hyps, prop)) in (case fst (strip_combt (fst (strip_combP prf))) of PThm ({serial, theory_name, name, prop = prop', ...}, _) => if prop = prop' then SOME {serial = serial, theory_name = theory_name, name = name} else NONE | _ => NONE) end; fun get_approximative_name shyps hyps prop prf = Option.map #name (get_identity shyps hyps prop prf) |> the_default ""; (* thm_id *) type thm_id = {serial: serial, theory_name: string}; fun make_thm_id (serial, theory_name) : thm_id = {serial = serial, theory_name = theory_name}; fun thm_header_id ({serial, theory_name, ...}: thm_header) = make_thm_id (serial, theory_name); fun thm_id (serial, thm_node) : thm_id = make_thm_id (serial, thm_node_theory_name thm_node); fun get_id shyps hyps prop prf : thm_id option = (case get_identity shyps hyps prop prf of NONE => NONE | SOME {name = "", ...} => NONE | SOME {serial, theory_name, ...} => SOME (make_thm_id (serial, theory_name))); fun this_id NONE _ = false | this_id (SOME (thm_id: thm_id)) (thm_id': thm_id) = #serial thm_id = #serial thm_id'; (* proof boxes: intermediate PThm nodes *) fun proof_boxes {included, excluded} proofs = let fun boxes_of (Abst (_, _, prf)) = boxes_of prf | boxes_of (AbsP (_, _, prf)) = boxes_of prf | boxes_of (prf % _) = boxes_of prf | boxes_of (prf1 %% prf2) = boxes_of prf1 #> boxes_of prf2 | boxes_of (PThm (header as {serial = i, ...}, thm_body)) = (fn boxes => let val thm_id = thm_header_id header in if Inttab.defined boxes i orelse (excluded thm_id andalso not (included thm_id)) then boxes else let val prf' = thm_body_proof_open thm_body; val boxes' = Inttab.update (i, (header, prf')) boxes; in boxes_of prf' boxes' end end) | boxes_of MinProof = raise MIN_PROOF () | boxes_of _ = I; - in Inttab.fold_rev (cons o #2) (fold boxes_of proofs Inttab.empty) [] end; + in Inttab.fold_rev (cons o #2) (Inttab.build (fold boxes_of proofs)) [] end; end; structure Basic_Proofterm = struct datatype proof = datatype Proofterm.proof datatype proof_body = datatype Proofterm.proof_body val op %> = Proofterm.%> end; open Basic_Proofterm; diff --git a/src/Pure/raw_simplifier.ML b/src/Pure/raw_simplifier.ML --- a/src/Pure/raw_simplifier.ML +++ b/src/Pure/raw_simplifier.ML @@ -1,1462 +1,1463 @@ (* Title: Pure/raw_simplifier.ML Author: Tobias Nipkow and Stefan Berghofer, TU Muenchen Higher-order Simplification. *) infix 4 addsimps delsimps addsimprocs delsimprocs setloop addloop delloop setSSolver addSSolver setSolver addSolver; signature BASIC_RAW_SIMPLIFIER = sig val simp_depth_limit: int Config.T val simp_trace_depth_limit: int Config.T val simp_debug: bool Config.T val simp_trace: bool Config.T type cong_name = bool * string type rrule val mk_rrules: Proof.context -> thm list -> rrule list val eq_rrule: rrule * rrule -> bool type proc type solver val mk_solver: string -> (Proof.context -> int -> tactic) -> solver type simpset val empty_ss: simpset val merge_ss: simpset * simpset -> simpset val dest_ss: simpset -> {simps: (string * thm) list, procs: (string * term list) list, congs: (cong_name * thm) list, weak_congs: cong_name list, loopers: string list, unsafe_solvers: string list, safe_solvers: string list} type simproc val eq_simproc: simproc * simproc -> bool val cert_simproc: theory -> string -> {lhss: term list, proc: morphism -> Proof.context -> cterm -> thm option} -> simproc val transform_simproc: morphism -> simproc -> simproc val simpset_of: Proof.context -> simpset val put_simpset: simpset -> Proof.context -> Proof.context val simpset_map: Proof.context -> (Proof.context -> Proof.context) -> simpset -> simpset val map_theory_simpset: (Proof.context -> Proof.context) -> theory -> theory val empty_simpset: Proof.context -> Proof.context val clear_simpset: Proof.context -> Proof.context val addsimps: Proof.context * thm list -> Proof.context val delsimps: Proof.context * thm list -> Proof.context val addsimprocs: Proof.context * simproc list -> Proof.context val delsimprocs: Proof.context * simproc list -> Proof.context val setloop: Proof.context * (Proof.context -> int -> tactic) -> Proof.context val addloop: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context val delloop: Proof.context * string -> Proof.context val setSSolver: Proof.context * solver -> Proof.context val addSSolver: Proof.context * solver -> Proof.context val setSolver: Proof.context * solver -> Proof.context val addSolver: Proof.context * solver -> Proof.context val rewrite_rule: Proof.context -> thm list -> thm -> thm val rewrite_goals_rule: Proof.context -> thm list -> thm -> thm val rewrite_goals_tac: Proof.context -> thm list -> tactic val rewrite_goal_tac: Proof.context -> thm list -> int -> tactic val prune_params_tac: Proof.context -> tactic val fold_rule: Proof.context -> thm list -> thm -> thm val fold_goals_tac: Proof.context -> thm list -> tactic val norm_hhf: Proof.context -> thm -> thm val norm_hhf_protect: Proof.context -> thm -> thm end; signature RAW_SIMPLIFIER = sig include BASIC_RAW_SIMPLIFIER exception SIMPLIFIER of string * thm list type trace_ops val set_trace_ops: trace_ops -> theory -> theory val subgoal_tac: Proof.context -> int -> tactic val loop_tac: Proof.context -> int -> tactic val solvers: Proof.context -> solver list * solver list val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic val prems_of: Proof.context -> thm list val add_simp: thm -> Proof.context -> Proof.context val del_simp: thm -> Proof.context -> Proof.context val flip_simp: thm -> Proof.context -> Proof.context val init_simpset: thm list -> Proof.context -> Proof.context val add_eqcong: thm -> Proof.context -> Proof.context val del_eqcong: thm -> Proof.context -> Proof.context val add_cong: thm -> Proof.context -> Proof.context val del_cong: thm -> Proof.context -> Proof.context val mksimps: Proof.context -> thm -> thm list val set_mksimps: (Proof.context -> thm -> thm list) -> Proof.context -> Proof.context val set_mkcong: (Proof.context -> thm -> thm) -> Proof.context -> Proof.context val set_mksym: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context val set_mkeqTrue: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context val set_term_ord: term ord -> Proof.context -> Proof.context val set_subgoaler: (Proof.context -> int -> tactic) -> Proof.context -> Proof.context val solver: Proof.context -> solver -> int -> tactic val default_mk_sym: Proof.context -> thm -> thm option val add_prems: thm list -> Proof.context -> Proof.context val set_reorient: (Proof.context -> term list -> term -> term -> bool) -> Proof.context -> Proof.context val set_solvers: solver list -> Proof.context -> Proof.context val rewrite_cterm: bool * bool * bool -> (Proof.context -> thm -> thm option) -> Proof.context -> conv val rewrite_term: theory -> thm list -> (term -> term option) list -> term -> term val rewrite_thm: bool * bool * bool -> (Proof.context -> thm -> thm option) -> Proof.context -> thm -> thm val generic_rewrite_goal_tac: bool * bool * bool -> (Proof.context -> tactic) -> Proof.context -> int -> tactic val rewrite: Proof.context -> bool -> thm list -> conv end; structure Raw_Simplifier: RAW_SIMPLIFIER = struct (** datatype simpset **) (* congruence rules *) type cong_name = bool * string; fun cong_name (Const (a, _)) = SOME (true, a) | cong_name (Free (a, _)) = SOME (false, a) | cong_name _ = NONE; structure Congtab = Table(type key = cong_name val ord = prod_ord bool_ord fast_string_ord); (* rewrite rules *) type rrule = {thm: thm, (*the rewrite rule*) name: string, (*name of theorem from which rewrite rule was extracted*) lhs: term, (*the left-hand side*) elhs: cterm, (*the eta-contracted lhs*) extra: bool, (*extra variables outside of elhs*) fo: bool, (*use first-order matching*) perm: bool}; (*the rewrite rule is permutative*) fun trim_context_rrule ({thm, name, lhs, elhs, extra, fo, perm}: rrule) = {thm = Thm.trim_context thm, name = name, lhs = lhs, elhs = Thm.trim_context_cterm elhs, extra = extra, fo = fo, perm = perm}; (* Remarks: - elhs is used for matching, lhs only for preservation of bound variable names; - fo is set iff either elhs is first-order (no Var is applied), in which case fo-matching is complete, or elhs is not a pattern, in which case there is nothing better to do; *) fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) = Thm.eq_thm_prop (thm1, thm2); (* FIXME: it seems that the conditions on extra variables are too liberal if prems are nonempty: does solving the prems really guarantee instantiation of all its Vars? Better: a dynamic check each time a rule is applied. *) fun rewrite_rule_extra_vars prems elhs erhs = let val elhss = elhs :: prems; - val tvars = fold Term.add_tvars elhss []; - val vars = fold Term.add_vars elhss []; + val tvars = Term_Subst.TVars.build (fold Term_Subst.add_tvars elhss); + val vars = Term_Subst.Vars.build (fold Term_Subst.add_vars elhss); in erhs |> Term.exists_type (Term.exists_subtype - (fn TVar v => not (member (op =) tvars v) | _ => false)) orelse + (fn TVar v => not (Term_Subst.TVars.defined tvars v) | _ => false)) orelse erhs |> Term.exists_subterm - (fn Var v => not (member (op =) vars v) | _ => false) + (fn Var v => not (Term_Subst.Vars.defined vars v) | _ => false) end; fun rrule_extra_vars elhs thm = rewrite_rule_extra_vars [] (Thm.term_of elhs) (Thm.full_prop_of thm); fun mk_rrule2 {thm, name, lhs, elhs, perm} = let val t = Thm.term_of elhs; val fo = Pattern.first_order t orelse not (Pattern.pattern t); val extra = rrule_extra_vars elhs thm; in {thm = thm, name = name, lhs = lhs, elhs = elhs, extra = extra, fo = fo, perm = perm} end; (*simple test for looping rewrite rules and stupid orientations*) fun default_reorient ctxt prems lhs rhs = rewrite_rule_extra_vars prems lhs rhs orelse is_Var (head_of lhs) orelse (* turns t = x around, which causes a headache if x is a local variable - usually it is very useful :-( is_Free rhs andalso not(is_Free lhs) andalso not(Logic.occs(rhs,lhs)) andalso not(exists_subterm is_Var lhs) orelse *) exists (fn t => Logic.occs (lhs, t)) (rhs :: prems) orelse null prems andalso Pattern.matches (Proof_Context.theory_of ctxt) (lhs, rhs) (*the condition "null prems" is necessary because conditional rewrites with extra variables in the conditions may terminate although the rhs is an instance of the lhs; example: ?m < ?n \ f ?n \ f ?m *) orelse is_Const lhs andalso not (is_Const rhs); (* simplification procedures *) datatype proc = Proc of {name: string, lhs: term, proc: Proof.context -> cterm -> thm option, stamp: stamp}; fun eq_proc (Proc {stamp = stamp1, ...}, Proc {stamp = stamp2, ...}) = stamp1 = stamp2; (* solvers *) datatype solver = Solver of {name: string, solver: Proof.context -> int -> tactic, id: stamp}; fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()}; fun solver_name (Solver {name, ...}) = name; fun solver ctxt (Solver {solver = tac, ...}) = tac ctxt; fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2); (* simplification sets *) (*A simpset contains data required during conversion: rules: discrimination net of rewrite rules; prems: current premises; depth: simp_depth and exceeded flag; congs: association list of congruence rules and a list of `weak' congruence constants. A congruence is `weak' if it avoids normalization of some argument. procs: discrimination net of simplification procedures (functions that prove rewrite rules on the fly); mk_rews: mk: turn simplification thms into rewrite rules; mk_cong: prepare congruence rules; mk_sym: turn \ around; mk_eq_True: turn P into P \ True; term_ord: for ordered rewriting;*) datatype simpset = Simpset of {rules: rrule Net.net, prems: thm list, depth: int * bool Unsynchronized.ref} * {congs: thm Congtab.table * cong_name list, procs: proc Net.net, mk_rews: {mk: Proof.context -> thm -> thm list, mk_cong: Proof.context -> thm -> thm, mk_sym: Proof.context -> thm -> thm option, mk_eq_True: Proof.context -> thm -> thm option, reorient: Proof.context -> term list -> term -> term -> bool}, term_ord: term ord, subgoal_tac: Proof.context -> int -> tactic, loop_tacs: (string * (Proof.context -> int -> tactic)) list, solvers: solver list * solver list}; fun internal_ss (Simpset (_, ss2)) = ss2; fun make_ss1 (rules, prems, depth) = {rules = rules, prems = prems, depth = depth}; fun map_ss1 f {rules, prems, depth} = make_ss1 (f (rules, prems, depth)); fun make_ss2 (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) = {congs = congs, procs = procs, mk_rews = mk_rews, term_ord = term_ord, subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers}; fun map_ss2 f {congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers} = make_ss2 (f (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)); fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2); fun dest_ss (Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...})) = {simps = Net.entries rules |> map (fn {name, thm, ...} => (name, thm)), procs = Net.entries procs |> map (fn Proc {name, lhs, stamp, ...} => ((name, lhs), stamp)) |> partition_eq (eq_snd op =) |> map (fn ps => (fst (fst (hd ps)), map (snd o fst) ps)), congs = congs |> fst |> Congtab.dest, weak_congs = congs |> snd, loopers = map fst loop_tacs, unsafe_solvers = map solver_name (#1 solvers), safe_solvers = map solver_name (#2 solvers)}; (* empty *) fun init_ss depth mk_rews term_ord subgoal_tac solvers = make_simpset ((Net.empty, [], depth), ((Congtab.empty, []), Net.empty, mk_rews, term_ord, subgoal_tac, [], solvers)); fun default_mk_sym _ th = SOME (th RS Drule.symmetric_thm); val empty_ss = init_ss (0, Unsynchronized.ref false) {mk = fn _ => fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [], mk_cong = K I, mk_sym = default_mk_sym, mk_eq_True = K (K NONE), reorient = default_reorient} Term_Ord.term_ord (K (K no_tac)) ([], []); (* merge *) (*NOTE: ignores some fields of 2nd simpset*) fun merge_ss (ss1, ss2) = if pointer_eq (ss1, ss2) then ss1 else let val Simpset ({rules = rules1, prems = prems1, depth = depth1}, {congs = (congs1, weak1), procs = procs1, mk_rews, term_ord, subgoal_tac, loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1; val Simpset ({rules = rules2, prems = prems2, depth = depth2}, {congs = (congs2, weak2), procs = procs2, mk_rews = _, term_ord = _, subgoal_tac = _, loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2; val rules' = Net.merge eq_rrule (rules1, rules2); val prems' = Thm.merge_thms (prems1, prems2); val depth' = if #1 depth1 < #1 depth2 then depth2 else depth1; val congs' = Congtab.merge (K true) (congs1, congs2); val weak' = merge (op =) (weak1, weak2); val procs' = Net.merge eq_proc (procs1, procs2); val loop_tacs' = AList.merge (op =) (K true) (loop_tacs1, loop_tacs2); val unsafe_solvers' = merge eq_solver (unsafe_solvers1, unsafe_solvers2); val solvers' = merge eq_solver (solvers1, solvers2); in make_simpset ((rules', prems', depth'), ((congs', weak'), procs', mk_rews, term_ord, subgoal_tac, loop_tacs', (unsafe_solvers', solvers'))) end; (** context data **) structure Simpset = Generic_Data ( type T = simpset; val empty = empty_ss; val extend = I; val merge = merge_ss; ); val simpset_of = Simpset.get o Context.Proof; fun map_simpset f = Context.proof_map (Simpset.map f); fun map_simpset1 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (map_ss1 f ss1, ss2)); fun map_simpset2 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (ss1, map_ss2 f ss2)); fun put_simpset ss = map_simpset (K ss); fun simpset_map ctxt f ss = ctxt |> put_simpset ss |> f |> simpset_of; val empty_simpset = put_simpset empty_ss; fun map_theory_simpset f thy = let val ctxt' = f (Proof_Context.init_global thy); val thy' = Proof_Context.theory_of ctxt'; in Context.theory_map (Simpset.map (K (simpset_of ctxt'))) thy' end; fun map_ss f = Context.mapping (map_theory_simpset (f o Context_Position.not_really)) f; val clear_simpset = map_simpset (fn Simpset ({depth, ...}, {mk_rews, term_ord, subgoal_tac, solvers, ...}) => init_ss depth mk_rews term_ord subgoal_tac solvers); (* accessors for tactis *) fun subgoal_tac ctxt = (#subgoal_tac o internal_ss o simpset_of) ctxt ctxt; fun loop_tac ctxt = FIRST' (map (fn (_, tac) => tac ctxt) (rev ((#loop_tacs o internal_ss o simpset_of) ctxt))); val solvers = #solvers o internal_ss o simpset_of (* simp depth *) (* The simp_depth_limit is meant to abort infinite recursion of the simplifier early but should not terminate "normal" executions. As of 2017, 25 would suffice; 40 builds in a safety margin. *) val simp_depth_limit = Config.declare_int ("simp_depth_limit", \<^here>) (K 40); val simp_trace_depth_limit = Config.declare_int ("simp_trace_depth_limit", \<^here>) (K 1); fun inc_simp_depth ctxt = ctxt |> map_simpset1 (fn (rules, prems, (depth, exceeded)) => (rules, prems, (depth + 1, if depth = Config.get ctxt simp_trace_depth_limit then Unsynchronized.ref false else exceeded))); fun simp_depth ctxt = let val Simpset ({depth = (depth, _), ...}, _) = simpset_of ctxt in depth end; (* diagnostics *) exception SIMPLIFIER of string * thm list; val simp_debug = Config.declare_bool ("simp_debug", \<^here>) (K false); val simp_trace = Config.declare_bool ("simp_trace", \<^here>) (K false); fun cond_warning ctxt msg = if Context_Position.is_really_visible ctxt then warning (msg ()) else (); fun cond_tracing' ctxt flag msg = if Config.get ctxt flag then let val Simpset ({depth = (depth, exceeded), ...}, _) = simpset_of ctxt; val depth_limit = Config.get ctxt simp_trace_depth_limit; in if depth > depth_limit then if ! exceeded then () else (tracing "simp_trace_depth_limit exceeded!"; exceeded := true) else (tracing (enclose "[" "]" (string_of_int depth) ^ msg ()); exceeded := false) end else (); fun cond_tracing ctxt = cond_tracing' ctxt simp_trace; fun print_term ctxt s t = s ^ "\n" ^ Syntax.string_of_term ctxt t; fun print_thm ctxt s (name, th) = print_term ctxt (if name = "" then s else s ^ " " ^ quote name ^ ":") (Thm.full_prop_of th); (** simpset operations **) (* prems *) fun prems_of ctxt = let val Simpset ({prems, ...}, _) = simpset_of ctxt in prems end; fun add_prems ths = map_simpset1 (fn (rules, prems, depth) => (rules, ths @ prems, depth)); (* maintain simp rules *) fun del_rrule loud (rrule as {thm, elhs, ...}) ctxt = ctxt |> map_simpset1 (fn (rules, prems, depth) => (Net.delete_term eq_rrule (Thm.term_of elhs, rrule) rules, prems, depth)) handle Net.DELETE => (if not loud then () else cond_warning ctxt (fn () => print_thm ctxt "Rewrite rule not in simpset:" ("", thm)); ctxt); fun insert_rrule (rrule as {thm, name, ...}) ctxt = (cond_tracing ctxt (fn () => print_thm ctxt "Adding rewrite rule" (name, thm)); ctxt |> map_simpset1 (fn (rules, prems, depth) => let val rrule2 as {elhs, ...} = mk_rrule2 rrule; val rules' = Net.insert_term eq_rrule (Thm.term_of elhs, trim_context_rrule rrule2) rules; in (rules', prems, depth) end) handle Net.INSERT => (cond_warning ctxt (fn () => print_thm ctxt "Ignoring duplicate rewrite rule:" ("", thm)); ctxt)); +val vars_set = Term_Subst.Vars.build o Term_Subst.add_vars; + local fun vperm (Var _, Var _) = true | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t) | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2) | vperm (t, u) = (t = u); -fun var_perm (t, u) = - vperm (t, u) andalso eq_set (op =) (Term.add_vars t [], Term.add_vars u []); +fun var_perm (t, u) = vperm (t, u) andalso Term_Subst.Vars.eq_set (apply2 vars_set (t, u)); in fun decomp_simp thm = let val prop = Thm.prop_of thm; val prems = Logic.strip_imp_prems prop; val concl = Drule.strip_imp_concl (Thm.cprop_of thm); val (lhs, rhs) = Thm.dest_equals concl handle TERM _ => raise SIMPLIFIER ("Rewrite rule not a meta-equality", [thm]); val elhs = Thm.dest_arg (Thm.cprop_of (Thm.eta_conversion lhs)); val erhs = Envir.eta_contract (Thm.term_of rhs); val perm = var_perm (Thm.term_of elhs, erhs) andalso not (Thm.term_of elhs aconv erhs) andalso not (is_Var (Thm.term_of elhs)); in (prems, Thm.term_of lhs, elhs, Thm.term_of rhs, perm) end; end; fun decomp_simp' thm = let val (_, lhs, _, rhs, _) = decomp_simp thm in if Thm.nprems_of thm > 0 then raise SIMPLIFIER ("Bad conditional rewrite rule", [thm]) else (lhs, rhs) end; fun mk_eq_True ctxt (thm, name) = let val Simpset (_, {mk_rews = {mk_eq_True, ...}, ...}) = simpset_of ctxt in (case mk_eq_True ctxt thm of NONE => [] | SOME eq_True => let val (_, lhs, elhs, _, _) = decomp_simp eq_True; in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end) end; (*create the rewrite rule and possibly also the eq_True variant, in case there are extra vars on the rhs*) fun rrule_eq_True ctxt thm name lhs elhs rhs thm2 = let val rrule = {thm = thm, name = name, lhs = lhs, elhs = elhs, perm = false} in if rewrite_rule_extra_vars [] lhs rhs then mk_eq_True ctxt (thm2, name) @ [rrule] else [rrule] end; fun mk_rrule ctxt (thm, name) = let val (prems, lhs, elhs, rhs, perm) = decomp_simp thm in if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}] else (*weak test for loops*) if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (Thm.term_of elhs) then mk_eq_True ctxt (thm, name) else rrule_eq_True ctxt thm name lhs elhs rhs thm end |> map (fn {thm, name, lhs, elhs, perm} => {thm = Thm.trim_context thm, name = name, lhs = lhs, elhs = Thm.trim_context_cterm elhs, perm = perm}); fun orient_rrule ctxt (thm, name) = let val (prems, lhs, elhs, rhs, perm) = decomp_simp thm; val Simpset (_, {mk_rews = {reorient, mk_sym, ...}, ...}) = simpset_of ctxt; in if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}] else if reorient ctxt prems lhs rhs then if reorient ctxt prems rhs lhs then mk_eq_True ctxt (thm, name) else (case mk_sym ctxt thm of NONE => [] | SOME thm' => let val (_, lhs', elhs', rhs', _) = decomp_simp thm' in rrule_eq_True ctxt thm' name lhs' elhs' rhs' thm end) else rrule_eq_True ctxt thm name lhs elhs rhs thm end; fun extract_rews ctxt sym thms = let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt; val mk = if sym then fn ctxt => fn th => (mk ctxt th) RL [Drule.symmetric_thm] else mk in maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk ctxt thm)) thms end; fun extract_safe_rrules ctxt thm = maps (orient_rrule ctxt) (extract_rews ctxt false [thm]); fun mk_rrules ctxt thms = let val rews = extract_rews ctxt false thms val raw_rrules = flat (map (mk_rrule ctxt) rews) in map mk_rrule2 raw_rrules end (* add/del rules explicitly *) local fun comb_simps ctxt comb mk_rrule sym thms = let val rews = extract_rews ctxt sym (map (Thm.transfer' ctxt) thms); in fold (fold comb o mk_rrule) rews ctxt end; (* This code checks if the symetric version of a rule is already in the simpset. However, the variable names in the two versions of the rule may differ. Thus the current test modulo eq_rrule is too weak to be useful and needs to be refined. fun present ctxt rules (rrule as {thm, elhs, ...}) = (Net.insert_term eq_rrule (Thm.term_of elhs, trim_context_rrule rrule) rules; false) handle Net.INSERT => (cond_warning ctxt (fn () => print_thm ctxt "Symmetric rewrite rule already in simpset:" ("", thm)); true); fun sym_present ctxt thms = let val rews = extract_rews ctxt true (map (Thm.transfer' ctxt) thms); val rrules = map mk_rrule2 (flat(map (mk_rrule ctxt) rews)) val Simpset({rules, ...},_) = simpset_of ctxt in exists (present ctxt rules) rrules end *) in fun ctxt addsimps thms = comb_simps ctxt insert_rrule (mk_rrule ctxt) false thms; fun addsymsimps ctxt thms = comb_simps ctxt insert_rrule (mk_rrule ctxt) true thms; fun ctxt delsimps thms = comb_simps ctxt (del_rrule true) (map mk_rrule2 o mk_rrule ctxt) false thms; fun delsimps_quiet ctxt thms = comb_simps ctxt (del_rrule false) (map mk_rrule2 o mk_rrule ctxt) false thms; fun add_simp thm ctxt = ctxt addsimps [thm]; (* with check for presence of symmetric version: if sym_present ctxt [thm] then (cond_warning ctxt (fn () => print_thm ctxt "Ignoring rewrite rule:" ("", thm)); ctxt) else ctxt addsimps [thm]; *) fun del_simp thm ctxt = ctxt delsimps [thm]; fun flip_simp thm ctxt = addsymsimps (delsimps_quiet ctxt [thm]) [thm]; end; fun init_simpset thms ctxt = ctxt |> Context_Position.set_visible false |> empty_simpset |> fold add_simp thms |> Context_Position.restore_visible ctxt; (* congs *) local fun is_full_cong_prems [] [] = true | is_full_cong_prems [] _ = false | is_full_cong_prems (p :: prems) varpairs = (case Logic.strip_assums_concl p of Const ("Pure.eq", _) $ lhs $ rhs => let val (x, xs) = strip_comb lhs and (y, ys) = strip_comb rhs in is_Var x andalso forall is_Bound xs andalso not (has_duplicates (op =) xs) andalso xs = ys andalso member (op =) varpairs (x, y) andalso is_full_cong_prems prems (remove (op =) (x, y) varpairs) end | _ => false); fun is_full_cong thm = let val prems = Thm.prems_of thm and concl = Thm.concl_of thm; val (lhs, rhs) = Logic.dest_equals concl; val (f, xs) = strip_comb lhs and (g, ys) = strip_comb rhs; in f = g andalso not (has_duplicates (op =) (xs @ ys)) andalso length xs = length ys andalso is_full_cong_prems prems (xs ~~ ys) end; fun mk_cong ctxt = let val Simpset (_, {mk_rews = {mk_cong = f, ...}, ...}) = simpset_of ctxt in f ctxt end; in fun add_eqcong thm ctxt = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => let val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", [thm]); (*val lhs = Envir.eta_contract lhs;*) val a = the (cong_name (head_of lhs)) handle Option.Option => raise SIMPLIFIER ("Congruence must start with a constant or free variable", [thm]); val (xs, weak) = congs; val xs' = Congtab.update (a, Thm.trim_context thm) xs; val weak' = if is_full_cong thm then weak else a :: weak; in ((xs', weak'), procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) end); fun del_eqcong thm ctxt = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => let val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", [thm]); (*val lhs = Envir.eta_contract lhs;*) val a = the (cong_name (head_of lhs)) handle Option.Option => raise SIMPLIFIER ("Congruence must start with a constant", [thm]); val (xs, _) = congs; val xs' = Congtab.delete_safe a xs; val weak' = Congtab.fold (fn (a, th) => if is_full_cong th then I else insert (op =) a) xs' []; in ((xs', weak'), procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) end); fun add_cong thm ctxt = add_eqcong (mk_cong ctxt thm) ctxt; fun del_cong thm ctxt = del_eqcong (mk_cong ctxt thm) ctxt; end; (* simprocs *) datatype simproc = Simproc of {name: string, lhss: term list, proc: morphism -> Proof.context -> cterm -> thm option, stamp: stamp}; fun eq_simproc (Simproc {stamp = stamp1, ...}, Simproc {stamp = stamp2, ...}) = stamp1 = stamp2; fun cert_simproc thy name {lhss, proc} = Simproc {name = name, lhss = map (Sign.cert_term thy) lhss, proc = proc, stamp = stamp ()}; fun transform_simproc phi (Simproc {name, lhss, proc, stamp}) = Simproc {name = name, lhss = map (Morphism.term phi) lhss, proc = Morphism.transform phi proc, stamp = stamp}; local fun add_proc (proc as Proc {name, lhs, ...}) ctxt = (cond_tracing ctxt (fn () => print_term ctxt ("Adding simplification procedure " ^ quote name ^ " for") lhs); ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => (congs, Net.insert_term eq_proc (lhs, proc) procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)) handle Net.INSERT => (cond_warning ctxt (fn () => "Ignoring duplicate simplification procedure " ^ quote name); ctxt)); fun del_proc (proc as Proc {name, lhs, ...}) ctxt = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => (congs, Net.delete_term eq_proc (lhs, proc) procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)) handle Net.DELETE => (cond_warning ctxt (fn () => "Simplification procedure " ^ quote name ^ " not in simpset"); ctxt); fun prep_procs (Simproc {name, lhss, proc, stamp}) = lhss |> map (fn lhs => Proc {name = name, lhs = lhs, proc = Morphism.form proc, stamp = stamp}); in fun ctxt addsimprocs ps = fold (fold add_proc o prep_procs) ps ctxt; fun ctxt delsimprocs ps = fold (fold del_proc o prep_procs) ps ctxt; end; (* mk_rews *) local fun map_mk_rews f = map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => let val {mk, mk_cong, mk_sym, mk_eq_True, reorient} = mk_rews; val (mk', mk_cong', mk_sym', mk_eq_True', reorient') = f (mk, mk_cong, mk_sym, mk_eq_True, reorient); val mk_rews' = {mk = mk', mk_cong = mk_cong', mk_sym = mk_sym', mk_eq_True = mk_eq_True', reorient = reorient'}; in (congs, procs, mk_rews', term_ord, subgoal_tac, loop_tacs, solvers) end); in fun mksimps ctxt = let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt in mk ctxt end; fun set_mksimps mk = map_mk_rews (fn (_, mk_cong, mk_sym, mk_eq_True, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); fun set_mkcong mk_cong = map_mk_rews (fn (mk, _, mk_sym, mk_eq_True, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); fun set_mksym mk_sym = map_mk_rews (fn (mk, mk_cong, _, mk_eq_True, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); fun set_mkeqTrue mk_eq_True = map_mk_rews (fn (mk, mk_cong, mk_sym, _, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); fun set_reorient reorient = map_mk_rews (fn (mk, mk_cong, mk_sym, mk_eq_True, _) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); end; (* term_ord *) fun set_term_ord term_ord = map_simpset2 (fn (congs, procs, mk_rews, _, subgoal_tac, loop_tacs, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)); (* tactics *) fun set_subgoaler subgoal_tac = map_simpset2 (fn (congs, procs, mk_rews, term_ord, _, loop_tacs, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)); fun ctxt setloop tac = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, _, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, [("", tac)], solvers)); fun ctxt addloop (name, tac) = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, AList.update (op =) (name, tac) loop_tacs, solvers)); fun ctxt delloop name = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, (if AList.defined (op =) loop_tacs name then () else cond_warning ctxt (fn () => "No such looper in simpset: " ^ quote name); AList.delete (op =) name loop_tacs), solvers)); fun ctxt setSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, _)) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, [solver]))); fun ctxt addSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, insert eq_solver solver solvers))); fun ctxt setSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, ([solver], solvers))); fun ctxt addSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (insert eq_solver solver unsafe_solvers, solvers))); fun set_solvers solvers = map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, _) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (solvers, solvers))); (* trace operations *) type trace_ops = {trace_invoke: {depth: int, term: term} -> Proof.context -> Proof.context, trace_apply: {unconditional: bool, term: term, thm: thm, rrule: rrule} -> Proof.context -> (Proof.context -> (thm * term) option) -> (thm * term) option}; structure Trace_Ops = Theory_Data ( type T = trace_ops; val empty: T = {trace_invoke = fn _ => fn ctxt => ctxt, trace_apply = fn _ => fn ctxt => fn cont => cont ctxt}; val extend = I; fun merge (trace_ops, _) = trace_ops; ); val set_trace_ops = Trace_Ops.put; val trace_ops = Trace_Ops.get o Proof_Context.theory_of; fun trace_invoke args ctxt = #trace_invoke (trace_ops ctxt) args ctxt; fun trace_apply args ctxt = #trace_apply (trace_ops ctxt) args ctxt; (** rewriting **) (* Uses conversions, see: L C Paulson, A higher-order implementation of rewriting, Science of Computer Programming 3 (1983), pages 119-149. *) fun check_conv ctxt msg thm thm' = let val thm'' = Thm.transitive thm thm' handle THM _ => let val nthm' = Thm.transitive (Thm.symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm' in Thm.transitive thm nthm' handle THM _ => let val nthm = Thm.transitive thm (Drule.beta_eta_conversion (Thm.rhs_of thm)) in Thm.transitive nthm nthm' end end val _ = if msg then cond_tracing ctxt (fn () => print_thm ctxt "SUCCEEDED" ("", thm')) else (); in SOME thm'' end handle THM _ => let val _ $ _ $ prop0 = Thm.prop_of thm; val _ = cond_tracing ctxt (fn () => print_thm ctxt "Proved wrong theorem (bad subgoaler?)" ("", thm') ^ "\n" ^ print_term ctxt "Should have proved:" prop0); in NONE end; (* mk_procrule *) fun mk_procrule ctxt thm = let val (prems, lhs, elhs, rhs, _) = decomp_simp thm val thm' = Thm.close_derivation \<^here> thm; in if rewrite_rule_extra_vars prems lhs rhs then (cond_warning ctxt (fn () => print_thm ctxt "Extra vars on rhs:" ("", thm)); []) else [mk_rrule2 {thm = thm', name = "", lhs = lhs, elhs = elhs, perm = false}] end; (* rewritec: conversion to apply the meta simpset to a term *) (*Since the rewriting strategy is bottom-up, we avoid re-normalizing already normalized terms by carrying around the rhs of the rewrite rule just applied. This is called the `skeleton'. It is decomposed in parallel with the term. Once a Var is encountered, the corresponding term is already in normal form. skel0 is a dummy skeleton that is to enforce complete normalization.*) val skel0 = Bound 0; (*Use rhs as skeleton only if the lhs does not contain unnormalized bits. The latter may happen iff there are weak congruence rules for constants in the lhs.*) fun uncond_skel ((_, weak), (lhs, rhs)) = if null weak then rhs (*optimization*) else if exists_subterm (fn Const (a, _) => member (op =) weak (true, a) | Free (a, _) => member (op =) weak (false, a) | _ => false) lhs then skel0 else rhs; (*Behaves like unconditional rule if rhs does not contain vars not in the lhs. Otherwise those vars may become instantiated with unnormalized terms while the premises are solved.*) fun cond_skel (args as (_, (lhs, rhs))) = - if subset (op =) (Term.add_vars rhs [], Term.add_vars lhs []) then uncond_skel args + if Term_Subst.Vars.subset (vars_set rhs, vars_set lhs) then uncond_skel args else skel0; (* Rewriting -- we try in order: (1) beta reduction (2) unconditional rewrite rules (3) conditional rewrite rules (4) simplification procedures IMPORTANT: rewrite rules must not introduce new Vars or TVars! *) fun rewritec (prover, maxt) ctxt t = let val thy = Proof_Context.theory_of ctxt; val Simpset ({rules, ...}, {congs, procs, term_ord, ...}) = simpset_of ctxt; val eta_thm = Thm.eta_conversion t; val eta_t' = Thm.rhs_of eta_thm; val eta_t = Thm.term_of eta_t'; fun rew rrule = let val {thm = thm0, name, lhs, elhs = elhs0, extra, fo, perm} = rrule; val thm = Thm.transfer thy thm0; val elhs = Thm.transfer_cterm thy elhs0; val prop = Thm.prop_of thm; val (rthm, elhs') = if maxt = ~1 orelse not extra then (thm, elhs) else (Thm.incr_indexes (maxt + 1) thm, Thm.incr_indexes_cterm (maxt + 1) elhs); val insts = if fo then Thm.first_order_match (elhs', eta_t') else Thm.match (elhs', eta_t'); val thm' = Thm.instantiate insts (Thm.rename_boundvars lhs eta_t rthm); val prop' = Thm.prop_of thm'; val unconditional = (Logic.count_prems prop' = 0); val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop'); val trace_args = {unconditional = unconditional, term = eta_t, thm = thm', rrule = rrule}; in if perm andalso is_greater_equal (term_ord (rhs', lhs')) then (cond_tracing ctxt (fn () => print_thm ctxt "Cannot apply permutative rewrite rule" (name, thm) ^ "\n" ^ print_thm ctxt "Term does not become smaller:" ("", thm')); NONE) else (cond_tracing ctxt (fn () => print_thm ctxt "Applying instance of rewrite rule" (name, thm)); if unconditional then (cond_tracing ctxt (fn () => print_thm ctxt "Rewriting:" ("", thm')); trace_apply trace_args ctxt (fn ctxt' => let val lr = Logic.dest_equals prop; val SOME thm'' = check_conv ctxt' false eta_thm thm'; in SOME (thm'', uncond_skel (congs, lr)) end)) else (cond_tracing ctxt (fn () => print_thm ctxt "Trying to rewrite:" ("", thm')); if simp_depth ctxt > Config.get ctxt simp_depth_limit then (cond_tracing ctxt (fn () => "simp_depth_limit exceeded - giving up"); NONE) else trace_apply trace_args ctxt (fn ctxt' => (case prover ctxt' thm' of NONE => (cond_tracing ctxt' (fn () => print_thm ctxt' "FAILED" ("", thm')); NONE) | SOME thm2 => (case check_conv ctxt' true eta_thm thm2 of NONE => NONE | SOME thm2' => let val concl = Logic.strip_imp_concl prop; val lr = Logic.dest_equals concl; in SOME (thm2', cond_skel (congs, lr)) end))))) end; fun rews [] = NONE | rews (rrule :: rrules) = let val opt = rew rrule handle Pattern.MATCH => NONE in (case opt of NONE => rews rrules | some => some) end; fun sort_rrules rrs = let fun is_simple ({thm, ...}: rrule) = (case Thm.prop_of thm of Const ("Pure.eq", _) $ _ $ _ => true | _ => false); fun sort [] (re1, re2) = re1 @ re2 | sort (rr :: rrs) (re1, re2) = if is_simple rr then sort rrs (rr :: re1, re2) else sort rrs (re1, rr :: re2); in sort rrs ([], []) end; fun proc_rews [] = NONE | proc_rews (Proc {name, proc, lhs, ...} :: ps) = if Pattern.matches (Proof_Context.theory_of ctxt) (lhs, Thm.term_of t) then (cond_tracing' ctxt simp_debug (fn () => print_term ctxt ("Trying procedure " ^ quote name ^ " on:") eta_t); (case proc ctxt eta_t' of NONE => (cond_tracing' ctxt simp_debug (fn () => "FAILED"); proc_rews ps) | SOME raw_thm => (cond_tracing ctxt (fn () => print_thm ctxt ("Procedure " ^ quote name ^ " produced rewrite rule:") ("", raw_thm)); (case rews (mk_procrule ctxt raw_thm) of NONE => (cond_tracing ctxt (fn () => print_term ctxt ("IGNORED result of simproc " ^ quote name ^ " -- does not match") (Thm.term_of t)); proc_rews ps) | some => some)))) else proc_rews ps; in (case eta_t of Abs _ $ _ => SOME (Thm.transitive eta_thm (Thm.beta_conversion false eta_t'), skel0) | _ => (case rews (sort_rrules (Net.match_term rules eta_t)) of NONE => proc_rews (Net.match_term procs eta_t) | some => some)) end; (* conversion to apply a congruence rule to a term *) fun congc prover ctxt maxt cong t = let val rthm = Thm.incr_indexes (maxt + 1) cong; val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (Thm.cprop_of rthm))); val insts = Thm.match (rlhs, t) (* Thm.match can raise Pattern.MATCH; is handled when congc is called *) val thm' = Thm.instantiate insts (Thm.rename_boundvars (Thm.term_of rlhs) (Thm.term_of t) rthm); val _ = cond_tracing ctxt (fn () => print_thm ctxt "Applying congruence rule:" ("", thm')); fun err (msg, thm) = (cond_tracing ctxt (fn () => print_thm ctxt msg ("", thm)); NONE); in (case prover thm' of NONE => err ("Congruence proof failed. Could not prove", thm') | SOME thm2 => (case check_conv ctxt true (Drule.beta_eta_conversion t) thm2 of NONE => err ("Congruence proof failed. Should not have proved", thm2) | SOME thm2' => if op aconv (apply2 Thm.term_of (Thm.dest_equals (Thm.cprop_of thm2'))) then NONE else SOME thm2')) end; val vA = (("A", 0), propT); val vB = (("B", 0), propT); val vC = (("C", 0), propT); fun transitive1 NONE NONE = NONE | transitive1 (SOME thm1) NONE = SOME thm1 | transitive1 NONE (SOME thm2) = SOME thm2 | transitive1 (SOME thm1) (SOME thm2) = SOME (Thm.transitive thm1 thm2); fun transitive2 thm = transitive1 (SOME thm); fun transitive3 thm = transitive1 thm o SOME; fun bottomc ((simprem, useprem, mutsimp), prover, maxidx) = let fun botc skel ctxt t = if is_Var skel then NONE else (case subc skel ctxt t of some as SOME thm1 => (case rewritec (prover, maxidx) ctxt (Thm.rhs_of thm1) of SOME (thm2, skel2) => transitive2 (Thm.transitive thm1 thm2) (botc skel2 ctxt (Thm.rhs_of thm2)) | NONE => some) | NONE => (case rewritec (prover, maxidx) ctxt t of SOME (thm2, skel2) => transitive2 thm2 (botc skel2 ctxt (Thm.rhs_of thm2)) | NONE => NONE)) and try_botc ctxt t = (case botc skel0 ctxt t of SOME trec1 => trec1 | NONE => Thm.reflexive t) and subc skel ctxt t0 = let val Simpset (_, {congs, ...}) = simpset_of ctxt in (case Thm.term_of t0 of Abs (a, T, _) => let val (v, ctxt') = Variable.next_bound (a, T) ctxt; val b = #1 (Term.dest_Free v); val (v', t') = Thm.dest_abs (SOME b) t0; val b' = #1 (Term.dest_Free (Thm.term_of v')); val _ = if b <> b' then warning ("Bad Simplifier context: renamed bound variable " ^ quote b ^ " to " ^ quote b' ^ Position.here (Position.thread_data ())) else (); val skel' = (case skel of Abs (_, _, sk) => sk | _ => skel0); in (case botc skel' ctxt' t' of SOME thm => SOME (Thm.abstract_rule a v' thm) | NONE => NONE) end | t $ _ => (case t of Const ("Pure.imp", _) $ _ => impc t0 ctxt | Abs _ => let val thm = Thm.beta_conversion false t0 in (case subc skel0 ctxt (Thm.rhs_of thm) of NONE => SOME thm | SOME thm' => SOME (Thm.transitive thm thm')) end | _ => let fun appc () = let val (tskel, uskel) = (case skel of tskel $ uskel => (tskel, uskel) | _ => (skel0, skel0)); val (ct, cu) = Thm.dest_comb t0; in (case botc tskel ctxt ct of SOME thm1 => (case botc uskel ctxt cu of SOME thm2 => SOME (Thm.combination thm1 thm2) | NONE => SOME (Thm.combination thm1 (Thm.reflexive cu))) | NONE => (case botc uskel ctxt cu of SOME thm1 => SOME (Thm.combination (Thm.reflexive ct) thm1) | NONE => NONE)) end; val (h, ts) = strip_comb t; in (case cong_name h of SOME a => (case Congtab.lookup (fst congs) a of NONE => appc () | SOME cong => (*post processing: some partial applications h t1 ... tj, j <= length ts, may be a redex. Example: map (\x. x) = (\xs. xs) wrt map_cong*) (let val thm = congc (prover ctxt) ctxt maxidx cong t0; val t = the_default t0 (Option.map Thm.rhs_of thm); val (cl, cr) = Thm.dest_comb t val dVar = Var(("", 0), dummyT) val skel = list_comb (h, replicate (length ts) dVar) in (case botc skel ctxt cl of NONE => thm | SOME thm' => transitive3 thm (Thm.combination thm' (Thm.reflexive cr))) end handle Pattern.MATCH => appc ())) | _ => appc ()) end) | _ => NONE) end and impc ct ctxt = if mutsimp then mut_impc0 [] ct [] [] ctxt else nonmut_impc ct ctxt and rules_of_prem prem ctxt = if maxidx_of_term (Thm.term_of prem) <> ~1 then (cond_tracing ctxt (fn () => print_term ctxt "Cannot add premise as rewrite rule because it contains (type) unknowns:" (Thm.term_of prem)); (([], NONE), ctxt)) else let val (asm, ctxt') = Thm.assume_hyps prem ctxt in ((extract_safe_rrules ctxt' asm, SOME asm), ctxt') end and add_rrules (rrss, asms) ctxt = (fold o fold) insert_rrule rrss ctxt |> add_prems (map_filter I asms) and disch r prem eq = let val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of eq); val eq' = Thm.implies_elim (Thm.instantiate ([], [(vA, prem), (vB, lhs), (vC, rhs)]) Drule.imp_cong) (Thm.implies_intr prem eq); in if not r then eq' else let val (prem', concl) = Thm.dest_implies lhs; val (prem'', _) = Thm.dest_implies rhs; in Thm.transitive (Thm.transitive (Thm.instantiate ([], [(vA, prem'), (vB, prem), (vC, concl)]) Drule.swap_prems_eq) eq') (Thm.instantiate ([], [(vA, prem), (vB, prem''), (vC, concl)]) Drule.swap_prems_eq) end end and rebuild [] _ _ _ _ eq = eq | rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ctxt eq = let val ctxt' = add_rrules (rev rrss, rev asms) ctxt; val concl' = Drule.mk_implies (prem, the_default concl (Option.map Thm.rhs_of eq)); val dprem = Option.map (disch false prem); in (case rewritec (prover, maxidx) ctxt' concl' of NONE => rebuild prems concl' rrss asms ctxt (dprem eq) | SOME (eq', _) => transitive2 (fold (disch false) prems (the (transitive3 (dprem eq) eq'))) (mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ctxt)) end and mut_impc0 prems concl rrss asms ctxt = let val prems' = strip_imp_prems concl; val ((rrss', asms'), ctxt') = fold_map rules_of_prem prems' ctxt |>> split_list; in mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss') (asms @ asms') [] [] [] [] ctxt' ~1 ~1 end and mut_impc [] concl [] [] prems' rrss' asms' eqns ctxt changed k = transitive1 (fold (fn (eq1, prem) => fn eq2 => transitive1 eq1 (Option.map (disch false prem) eq2)) (eqns ~~ prems') NONE) (if changed > 0 then mut_impc (rev prems') concl (rev rrss') (rev asms') [] [] [] [] ctxt ~1 changed else rebuild prems' concl rrss' asms' ctxt (botc skel0 (add_rrules (rev rrss', rev asms') ctxt) concl)) | mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms) prems' rrss' asms' eqns ctxt changed k = (case (if k = 0 then NONE else botc skel0 (add_rrules (rev rrss' @ rrss, rev asms' @ asms) ctxt) prem) of NONE => mut_impc prems concl rrss asms (prem :: prems') (rrs :: rrss') (asm :: asms') (NONE :: eqns) ctxt changed (if k = 0 then 0 else k - 1) | SOME eqn => let val prem' = Thm.rhs_of eqn; val tprems = map Thm.term_of prems; val i = 1 + fold Integer.max (map (fn p => find_index (fn q => q aconv p) tprems) (Thm.hyps_of eqn)) ~1; val ((rrs', asm'), ctxt') = rules_of_prem prem' ctxt; in mut_impc prems concl rrss asms (prem' :: prems') (rrs' :: rrss') (asm' :: asms') (SOME (fold_rev (disch true) (take i prems) (Drule.imp_cong_rule eqn (Thm.reflexive (Drule.list_implies (drop i prems, concl))))) :: eqns) ctxt' (length prems') ~1 end) (*legacy code -- only for backwards compatibility*) and nonmut_impc ct ctxt = let val (prem, conc) = Thm.dest_implies ct; val thm1 = if simprem then botc skel0 ctxt prem else NONE; val prem1 = the_default prem (Option.map Thm.rhs_of thm1); val ctxt1 = if not useprem then ctxt else let val ((rrs, asm), ctxt') = rules_of_prem prem1 ctxt in add_rrules ([rrs], [asm]) ctxt' end; in (case botc skel0 ctxt1 conc of NONE => (case thm1 of NONE => NONE | SOME thm1' => SOME (Drule.imp_cong_rule thm1' (Thm.reflexive conc))) | SOME thm2 => let val thm2' = disch false prem1 thm2 in (case thm1 of NONE => SOME thm2' | SOME thm1' => SOME (Thm.transitive (Drule.imp_cong_rule thm1' (Thm.reflexive conc)) thm2')) end) end; in try_botc end; (* Meta-rewriting: rewrites t to u and returns the theorem t \ u *) (* Parameters: mode = (simplify A, use A in simplifying B, use prems of B (if B is again a meta-impl.) to simplify A) when simplifying A \ B prover: how to solve premises in conditional rewrites and congruences *) fun rewrite_cterm mode prover raw_ctxt raw_ct = let val thy = Proof_Context.theory_of raw_ctxt; val ct = raw_ct |> Thm.transfer_cterm thy |> Thm.adjust_maxidx_cterm ~1; val maxidx = Thm.maxidx_of_cterm ct; val ctxt = raw_ctxt |> Variable.set_body true |> Context_Position.set_visible false |> inc_simp_depth |> (fn ctxt => trace_invoke {depth = simp_depth ctxt, term = Thm.term_of ct} ctxt); val _ = cond_tracing ctxt (fn () => print_term ctxt "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" (Thm.term_of ct)); in ct |> bottomc (mode, Option.map (Drule.flexflex_unique (SOME ctxt)) oo prover, maxidx) ctxt |> Thm.solve_constraints end; val simple_prover = SINGLE o (fn ctxt => ALLGOALS (resolve_tac ctxt (prems_of ctxt))); fun rewrite _ _ [] = Thm.reflexive | rewrite ctxt full thms = rewrite_cterm (full, false, false) simple_prover (init_simpset thms ctxt); fun rewrite_rule ctxt = Conv.fconv_rule o rewrite ctxt true; (*simple term rewriting -- no proof*) fun rewrite_term thy rules procs = Pattern.rewrite_term thy (map decomp_simp' rules) procs; fun rewrite_thm mode prover ctxt = Conv.fconv_rule (rewrite_cterm mode prover ctxt); (*Rewrite the subgoals of a proof state (represented by a theorem)*) fun rewrite_goals_rule ctxt thms th = Conv.fconv_rule (Conv.prems_conv ~1 (rewrite_cterm (true, true, true) simple_prover (init_simpset thms ctxt))) th; (** meta-rewriting tactics **) (*Rewrite all subgoals*) fun rewrite_goals_tac ctxt defs = PRIMITIVE (rewrite_goals_rule ctxt defs); (*Rewrite one subgoal*) fun generic_rewrite_goal_tac mode prover_tac ctxt i thm = if 0 < i andalso i <= Thm.nprems_of thm then Seq.single (Conv.gconv_rule (rewrite_cterm mode (SINGLE o prover_tac) ctxt) i thm) else Seq.empty; fun rewrite_goal_tac ctxt thms = generic_rewrite_goal_tac (true, false, false) (K no_tac) (init_simpset thms ctxt); (*Prunes all redundant parameters from the proof state by rewriting.*) fun prune_params_tac ctxt = rewrite_goals_tac ctxt [Drule.triv_forall_equality]; (* for folding definitions, handling critical pairs *) (*The depth of nesting in a term*) fun term_depth (Abs (_, _, t)) = 1 + term_depth t | term_depth (f $ t) = 1 + Int.max (term_depth f, term_depth t) | term_depth _ = 0; val lhs_of_thm = #1 o Logic.dest_equals o Thm.prop_of; (*folding should handle critical pairs! E.g. K \ Inl 0, S \ Inr (Inl 0) Returns longest lhs first to avoid folding its subexpressions.*) fun sort_lhs_depths defs = let val keylist = AList.make (term_depth o lhs_of_thm) defs val keys = sort_distinct (rev_order o int_ord) (map #2 keylist) in map (AList.find (op =) keylist) keys end; val rev_defs = sort_lhs_depths o map Thm.symmetric; fun fold_rule ctxt defs = fold (rewrite_rule ctxt) (rev_defs defs); fun fold_goals_tac ctxt defs = EVERY (map (rewrite_goals_tac ctxt) (rev_defs defs)); (* HHF normal form: \ before \, outermost \ generalized *) local fun gen_norm_hhf ss ctxt0 th0 = let val (ctxt, th) = Thm.join_transfer_context (ctxt0, th0); val th' = if Drule.is_norm_hhf (Thm.prop_of th) then th else Conv.fconv_rule (rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset ss ctxt)) th; in th' |> Thm.adjust_maxidx_thm ~1 |> Variable.gen_all ctxt end; val hhf_ss = Context.the_local_context () |> init_simpset Drule.norm_hhf_eqs |> simpset_of; val hhf_protect_ss = Context.the_local_context () |> init_simpset Drule.norm_hhf_eqs |> add_eqcong Drule.protect_cong |> simpset_of; in val norm_hhf = gen_norm_hhf hhf_ss; val norm_hhf_protect = gen_norm_hhf hhf_protect_ss; end; end; structure Basic_Meta_Simplifier: BASIC_RAW_SIMPLIFIER = Raw_Simplifier; open Basic_Meta_Simplifier; diff --git a/src/Pure/sorts.ML b/src/Pure/sorts.ML --- a/src/Pure/sorts.ML +++ b/src/Pure/sorts.ML @@ -1,494 +1,492 @@ (* Title: Pure/sorts.ML Author: Markus Wenzel and Stefan Berghofer, TU Muenchen The order-sorted algebra of type classes. Classes denote (possibly empty) collections of types that are partially ordered by class inclusion. They are represented symbolically by strings. Sorts are intersections of finitely many classes. They are represented by lists of classes. Normal forms of sorts are sorted lists of minimal classes (wrt. current class inclusion). *) signature SORTS = sig val make: sort list -> sort Ord_List.T val subset: sort Ord_List.T * sort Ord_List.T -> bool val union: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T val subtract: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T val remove_sort: sort -> sort Ord_List.T -> sort Ord_List.T val insert_sort: sort -> sort Ord_List.T -> sort Ord_List.T val insert_typ: typ -> sort Ord_List.T -> sort Ord_List.T val insert_typs: typ list -> sort Ord_List.T -> sort Ord_List.T val insert_term: term -> sort Ord_List.T -> sort Ord_List.T val insert_terms: term list -> sort Ord_List.T -> sort Ord_List.T type algebra val classes_of: algebra -> serial Graph.T val arities_of: algebra -> (class * sort list) list Symtab.table val all_classes: algebra -> class list val super_classes: algebra -> class -> class list val class_less: algebra -> class * class -> bool val class_le: algebra -> class * class -> bool val sort_eq: algebra -> sort * sort -> bool val sort_le: algebra -> sort * sort -> bool val sorts_le: algebra -> sort list * sort list -> bool val inter_sort: algebra -> sort * sort -> sort val minimize_sort: algebra -> sort -> sort val complete_sort: algebra -> sort -> sort val add_class: Context.generic -> class * class list -> algebra -> algebra val add_classrel: Context.generic -> class * class -> algebra -> algebra val add_arities: Context.generic -> string * (class * sort list) list -> algebra -> algebra val empty_algebra: algebra val merge_algebra: Context.generic -> algebra * algebra -> algebra val dest_algebra: algebra list -> algebra -> {classrel: (class * class list) list, arities: (string * sort list * class) list} val subalgebra: Context.generic -> (class -> bool) -> (class * string -> sort list option) -> algebra -> (sort -> sort) * algebra type class_error val class_error: Context.generic -> class_error -> string exception CLASS_ERROR of class_error val has_instance: algebra -> string -> sort -> bool val mg_domain: algebra -> string -> sort -> sort list (*exception CLASS_ERROR*) val meet_sort: algebra -> typ * sort -> sort Vartab.table -> sort Vartab.table (*exception CLASS_ERROR*) val meet_sort_typ: algebra -> typ * sort -> typ -> typ (*exception CLASS_ERROR*) val of_sort: algebra -> typ * sort -> bool val of_sort_derivation: algebra -> {class_relation: typ -> bool -> 'a * class -> class -> 'a, type_constructor: string * typ list -> ('a * class) list list -> class -> 'a, type_variable: typ -> ('a * class) list} -> typ * sort -> 'a list (*exception CLASS_ERROR*) val classrel_derivation: algebra -> ('a * class -> class -> 'a) -> 'a * class -> class -> 'a (*exception CLASS_ERROR*) val witness_sorts: algebra -> string list -> (typ * sort) list -> sort list -> (typ * sort) list end; structure Sorts: SORTS = struct (** ordered lists of sorts **) val make = Ord_List.make Term_Ord.sort_ord; val subset = Ord_List.subset Term_Ord.sort_ord; val union = Ord_List.union Term_Ord.sort_ord; val subtract = Ord_List.subtract Term_Ord.sort_ord; val remove_sort = Ord_List.remove Term_Ord.sort_ord; val insert_sort = Ord_List.insert Term_Ord.sort_ord; fun insert_typ (TFree (_, S)) Ss = insert_sort S Ss | insert_typ (TVar (_, S)) Ss = insert_sort S Ss | insert_typ (Type (_, Ts)) Ss = insert_typs Ts Ss and insert_typs [] Ss = Ss | insert_typs (T :: Ts) Ss = insert_typs Ts (insert_typ T Ss); fun insert_term (Const (_, T)) Ss = insert_typ T Ss | insert_term (Free (_, T)) Ss = insert_typ T Ss | insert_term (Var (_, T)) Ss = insert_typ T Ss | insert_term (Bound _) Ss = Ss | insert_term (Abs (_, T, t)) Ss = insert_term t (insert_typ T Ss) | insert_term (t $ u) Ss = insert_term t (insert_term u Ss); fun insert_terms [] Ss = Ss | insert_terms (t :: ts) Ss = insert_terms ts (insert_term t Ss); (** order-sorted algebra **) (* classes: graph representing class declarations together with proper subclass relation, which needs to be transitive and acyclic. arities: table of association lists of all type arities; (t, ars) means that type constructor t has the arities ars; an element (c, Ss) of ars represents the arity t::(Ss)c. "Coregularity" of the arities structure requires that for any two declarations t::(Ss1)c1 and t::(Ss2)c2 such that c1 <= c2 holds Ss1 <= Ss2. *) datatype algebra = Algebra of {classes: serial Graph.T, arities: (class * sort list) list Symtab.table}; fun classes_of (Algebra {classes, ...}) = classes; fun arities_of (Algebra {arities, ...}) = arities; fun make_algebra (classes, arities) = Algebra {classes = classes, arities = arities}; fun map_classes f (Algebra {classes, arities}) = make_algebra (f classes, arities); fun map_arities f (Algebra {classes, arities}) = make_algebra (classes, f arities); (* classes *) fun all_classes (Algebra {classes, ...}) = Graph.all_preds classes (Graph.maximals classes); val super_classes = Graph.immediate_succs o classes_of; (* class relations *) val class_less : algebra -> class * class -> bool = Graph.is_edge o classes_of; fun class_le algebra (c1, c2) = c1 = c2 orelse class_less algebra (c1, c2); (* sort relations *) fun sort_le algebra (S1: sort, S2: sort) = S1 = S2 orelse forall (fn c2 => exists (fn c1 => class_le algebra (c1, c2)) S1) S2; fun sorts_le algebra (Ss1, Ss2) = ListPair.all (sort_le algebra) (Ss1, Ss2); fun sort_eq algebra (S1, S2) = sort_le algebra (S1, S2) andalso sort_le algebra (S2, S1); (* intersection *) fun inter_class algebra c S = let fun intr [] = [c] | intr (S' as c' :: c's) = if class_le algebra (c', c) then S' else if class_le algebra (c, c') then intr c's else c' :: intr c's in intr S end; fun inter_sort algebra (S1, S2) = sort_strings (fold (inter_class algebra) S1 S2); (* normal forms *) fun minimize_sort _ [] = [] | minimize_sort _ (S as [_]) = S | minimize_sort algebra S = filter (fn c => not (exists (fn c' => class_less algebra (c', c)) S)) S |> sort_distinct string_ord; fun complete_sort algebra = Graph.all_succs (classes_of algebra) o minimize_sort algebra; (** build algebras **) (* classes *) fun err_dup_class c = error ("Duplicate declaration of class: " ^ quote c); fun err_cyclic_classes context css = error (cat_lines (map (fn cs => "Cycle in class relation: " ^ Syntax.string_of_classrel (Syntax.init_pretty context) cs) css)); fun add_class context (c, cs) = map_classes (fn classes => let val classes' = classes |> Graph.new_node (c, serial ()) handle Graph.DUP dup => err_dup_class dup; val classes'' = classes' |> fold Graph.add_edge_trans_acyclic (map (pair c) cs) handle Graph.CYCLES css => err_cyclic_classes context css; in classes'' end); (* arities *) local fun for_classes _ NONE = "" | for_classes ctxt (SOME (c1, c2)) = " for classes " ^ Syntax.string_of_classrel ctxt [c1, c2]; fun err_conflict context t cc (c, Ss) (c', Ss') = let val ctxt = Syntax.init_pretty context in error ("Conflict of type arities" ^ for_classes ctxt cc ^ ":\n " ^ Syntax.string_of_arity ctxt (t, Ss, [c]) ^ " and\n " ^ Syntax.string_of_arity ctxt (t, Ss', [c'])) end; fun coregular context algebra t (c, Ss) ars = let fun conflict (c', Ss') = if class_le algebra (c, c') andalso not (sorts_le algebra (Ss, Ss')) then SOME ((c, c'), (c', Ss')) else if class_le algebra (c', c) andalso not (sorts_le algebra (Ss', Ss)) then SOME ((c', c), (c', Ss')) else NONE; in (case get_first conflict ars of SOME ((c1, c2), (c', Ss')) => err_conflict context t (SOME (c1, c2)) (c, Ss) (c', Ss') | NONE => (c, Ss) :: ars) end; fun complete algebra (c, Ss) = map (rpair Ss) (c :: super_classes algebra c); fun insert context algebra t (c, Ss) ars = (case AList.lookup (op =) ars c of NONE => coregular context algebra t (c, Ss) ars | SOME Ss' => if sorts_le algebra (Ss, Ss') then ars else if sorts_le algebra (Ss', Ss) then coregular context algebra t (c, Ss) (remove (op =) (c, Ss') ars) else err_conflict context t NONE (c, Ss) (c, Ss')); in fun insert_ars context algebra t = fold_rev (insert context algebra t); fun insert_complete_ars context algebra (t, ars) arities = let val ars' = Symtab.lookup_list arities t |> fold_rev (insert_ars context algebra t) (map (complete algebra) ars); in Symtab.update (t, ars') arities end; fun add_arities context arg algebra = algebra |> map_arities (insert_complete_ars context algebra arg); fun add_arities_table context algebra = Symtab.fold (fn (t, ars) => insert_complete_ars context algebra (t, ars)); end; (* classrel *) -fun rebuild_arities context algebra = algebra |> map_arities (fn arities => - Symtab.empty - |> add_arities_table context algebra arities); +fun rebuild_arities context algebra = algebra + |> map_arities (Symtab.build o add_arities_table context algebra); fun add_classrel context rel = rebuild_arities context o map_classes (fn classes => classes |> Graph.add_edge_trans_acyclic rel handle Graph.CYCLES css => err_cyclic_classes context css); (* empty and merge *) val empty_algebra = make_algebra (Graph.empty, Symtab.empty); fun merge_algebra context (Algebra {classes = classes1, arities = arities1}, Algebra {classes = classes2, arities = arities2}) = let val classes' = Graph.merge_trans_acyclic (op =) (classes1, classes2) handle Graph.DUP c => err_dup_class c | Graph.CYCLES css => err_cyclic_classes context css; val algebra0 = make_algebra (classes', Symtab.empty); val arities' = (case (pointer_eq (classes1, classes2), pointer_eq (arities1, arities2)) of (true, true) => arities1 | (true, false) => (*no completion*) (arities1, arities2) |> Symtab.join (fn t => fn (ars1, ars2) => if pointer_eq (ars1, ars2) then raise Symtab.SAME else insert_ars context algebra0 t ars2 ars1) | (false, true) => (*unary completion*) - Symtab.empty - |> add_arities_table context algebra0 arities1 + Symtab.build (add_arities_table context algebra0 arities1) | (false, false) => (*binary completion*) - Symtab.empty - |> add_arities_table context algebra0 arities1 - |> add_arities_table context algebra0 arities2); + Symtab.build + (add_arities_table context algebra0 arities1 #> + add_arities_table context algebra0 arities2)); in make_algebra (classes', arities') end; (* destruct *) fun dest_algebra parents (Algebra {classes, arities}) = let fun new_classrel rel = not (exists (fn algebra => class_less algebra rel) parents); val classrel = - (classes, []) |-> Graph.fold (fn (c, (_, (_, ds))) => + build (classes |> Graph.fold (fn (c, (_, (_, ds))) => (case filter (fn d => new_classrel (c, d)) (Graph.Keys.dest ds) of [] => I - | ds' => cons (c, sort_strings ds'))) + | ds' => cons (c, sort_strings ds')))) |> sort_by #1; fun is_arity t ar algebra = member (op =) (Symtab.lookup_list (arities_of algebra) t) ar; fun add_arity t (c, Ss) = not (exists (is_arity t (c, Ss)) parents) ? cons (t, Ss, c); val arities = - (arities, []) |-> Symtab.fold (fn (t, ars) => fold_rev (add_arity t) (sort_by #1 ars)) + build (arities |> Symtab.fold (fn (t, ars) => fold_rev (add_arity t) (sort_by #1 ars))) |> sort_by #1; in {classrel = classrel, arities = arities} end; (* algebra projections *) (* FIXME potentially violates abstract type integrity *) fun subalgebra context P sargs (algebra as Algebra {classes, arities}) = let val restrict_sort = minimize_sort algebra o filter P o Graph.all_succs classes; fun restrict_arity t (c, Ss) = if P c then (case sargs (c, t) of SOME sorts => SOME (c, Ss |> map2 (curry (inter_sort algebra)) sorts |> map restrict_sort) | NONE => NONE) else NONE; val classes' = classes |> Graph.restrict P; val arities' = arities |> Symtab.map (map_filter o restrict_arity); in (restrict_sort, rebuild_arities context (make_algebra (classes', arities'))) end; (** sorts of types **) (* errors -- performance tuning via delayed message composition *) datatype class_error = No_Classrel of class * class | No_Arity of string * class | No_Subsort of sort * sort; fun class_error context = let val ctxt = Syntax.init_pretty context in fn No_Classrel (c1, c2) => "No class relation " ^ Syntax.string_of_classrel ctxt [c1, c2] | No_Arity (a, c) => "No type arity " ^ Syntax.string_of_arity ctxt (a, [], [c]) | No_Subsort (S1, S2) => "Cannot derive subsort relation " ^ Syntax.string_of_sort ctxt S1 ^ " < " ^ Syntax.string_of_sort ctxt S2 end; exception CLASS_ERROR of class_error; (* instances *) fun has_instance algebra a = forall (AList.defined (op =) (Symtab.lookup_list (arities_of algebra) a)); fun mg_domain algebra a S = let val ars = Symtab.lookup_list (arities_of algebra) a; fun dom c = (case AList.lookup (op =) ars c of NONE => raise CLASS_ERROR (No_Arity (a, c)) | SOME Ss => Ss); fun dom_inter c Ss = ListPair.map (inter_sort algebra) (dom c, Ss); in (case S of [] => raise Fail "Unknown domain of empty intersection" | c :: cs => fold dom_inter cs (dom c)) end; (* meet_sort *) fun meet_sort algebra = let fun inters S S' = inter_sort algebra (S, S'); fun meet _ [] = I | meet (TFree (_, S)) S' = if sort_le algebra (S, S') then I else raise CLASS_ERROR (No_Subsort (S, S')) | meet (TVar (v, S)) S' = if sort_le algebra (S, S') then I else Vartab.map_default (v, S) (inters S') | meet (Type (a, Ts)) S = fold2 meet Ts (mg_domain algebra a S); in uncurry meet end; fun meet_sort_typ algebra (T, S) = - let val tab = meet_sort algebra (T, S) Vartab.empty; + let val tab = Vartab.build (meet_sort algebra (T, S)); in Term.map_type_tvar (fn (v, _) => TVar (v, (the o Vartab.lookup tab) v)) end; (* of_sort *) fun of_sort algebra = let fun ofS (_, []) = true | ofS (TFree (_, S), S') = sort_le algebra (S, S') | ofS (TVar (_, S), S') = sort_le algebra (S, S') | ofS (Type (a, Ts), S) = let val Ss = mg_domain algebra a S in ListPair.all ofS (Ts, Ss) end handle CLASS_ERROR _ => false; in ofS end; (* animating derivations *) fun of_sort_derivation algebra {class_relation, type_constructor, type_variable} = let val arities = arities_of algebra; fun weaken T D1 S2 = let val S1 = map snd D1 in if S1 = S2 then map fst D1 else S2 |> map (fn c2 => (case D1 |> filter (fn (_, c1) => class_le algebra (c1, c2)) of [d1] => class_relation T true d1 c2 | (d1 :: _ :: _) => class_relation T false d1 c2 | [] => raise CLASS_ERROR (No_Subsort (S1, S2)))) end; fun derive (_, []) = [] | derive (Type (a, Us), S) = let val Ss = mg_domain algebra a S; val dom = map2 (fn U => fn S => derive (U, S) ~~ S) Us Ss; in S |> map (fn c => let val Ss' = the (AList.lookup (op =) (Symtab.lookup_list arities a) c); val dom' = map (fn ((U, d), S') => weaken U d S' ~~ S') ((Us ~~ dom) ~~ Ss'); in type_constructor (a, Us) dom' c end) end | derive (T, S) = weaken T (type_variable T) S; in derive end; fun classrel_derivation algebra class_relation = let fun path (x, c1 :: c2 :: cs) = path (class_relation (x, c1) c2, c2 :: cs) | path (x, _) = x; in fn (x, c1) => fn c2 => (case Graph.irreducible_paths (classes_of algebra) (c1, c2) of [] => raise CLASS_ERROR (No_Classrel (c1, c2)) | cs :: _ => path (x, cs)) end; (* witness_sorts *) fun witness_sorts algebra ground_types hyps sorts = let fun le S1 S2 = sort_le algebra (S1, S2); fun get S2 (T, S1) = if le S1 S2 then SOME (T, S2) else NONE; fun mg_dom t S = SOME (mg_domain algebra t S) handle CLASS_ERROR _ => NONE; fun witn_sort _ [] solved_failed = (SOME (propT, []), solved_failed) | witn_sort path S (solved, failed) = if exists (le S) failed then (NONE, (solved, failed)) else (case get_first (get S) solved of SOME w => (SOME w, (solved, failed)) | NONE => (case get_first (get S) hyps of SOME w => (SOME w, (w :: solved, failed)) | NONE => witn_types path ground_types S (solved, failed))) and witn_sorts path x = fold_map (witn_sort path) x and witn_types _ [] S (solved, failed) = (NONE, (solved, S :: failed)) | witn_types path (t :: ts) S solved_failed = (case mg_dom t S of SOME SS => (*do not descend into stronger args (achieving termination)*) if exists (fn D => le D S orelse exists (le D) path) SS then witn_types path ts S solved_failed else let val (ws, (solved', failed')) = witn_sorts (S :: path) SS solved_failed in if forall is_some ws then let val w = (Type (t, map (#1 o the) ws), S) in (SOME w, (w :: solved', failed')) end else witn_types path ts S (solved', failed') end | NONE => witn_types path ts S solved_failed); in map_filter I (#1 (witn_sorts [] sorts ([], []))) end; end; diff --git a/src/Pure/term_subst.ML b/src/Pure/term_subst.ML --- a/src/Pure/term_subst.ML +++ b/src/Pure/term_subst.ML @@ -1,270 +1,289 @@ (* Title: Pure/term_subst.ML Author: Makarius Efficient type/term substitution. *) signature INST_TABLE = sig include TABLE val add: key * 'a -> 'a table -> 'a table val table: (key * 'a) list -> 'a table end; functor Inst_Table(Key: KEY): INST_TABLE = struct structure Tab = Table(Key); fun add entry = Tab.insert (K true) entry; -fun table entries = fold add entries Tab.empty; +fun table entries = Tab.build (fold add entries); open Tab; end; signature TERM_SUBST = sig structure TFrees: INST_TABLE structure TVars: INST_TABLE structure Frees: INST_TABLE structure Vars: INST_TABLE + val add_tfreesT: typ -> TFrees.set -> TFrees.set + val add_tfrees: term -> TFrees.set -> TFrees.set + val add_tvarsT: typ -> TVars.set -> TVars.set + val add_tvars: term -> TVars.set -> TVars.set + val add_frees: term -> Frees.set -> Frees.set + val add_vars: term -> Vars.set -> Vars.set + val add_tfree_namesT: typ -> Symtab.set -> Symtab.set + val add_tfree_names: term -> Symtab.set -> Symtab.set + val add_free_names: term -> Symtab.set -> Symtab.set val map_atypsT_same: typ Same.operation -> typ Same.operation val map_types_same: typ Same.operation -> term Same.operation val map_aterms_same: term Same.operation -> term Same.operation val generalizeT_same: Symtab.set -> int -> typ Same.operation val generalize_same: Symtab.set * Symtab.set -> int -> term Same.operation val generalizeT: Symtab.set -> int -> typ -> typ val generalize: Symtab.set * Symtab.set -> int -> term -> term val instantiateT_maxidx: (typ * int) TVars.table -> typ -> int -> typ * int val instantiate_maxidx: (typ * int) TVars.table * (term * int) Vars.table -> term -> int -> term * int val instantiateT_frees_same: typ TFrees.table -> typ Same.operation val instantiate_frees_same: typ TFrees.table * term Frees.table -> term Same.operation val instantiateT_frees: typ TFrees.table -> typ -> typ val instantiate_frees: typ TFrees.table * term Frees.table -> term -> term val instantiateT_same: typ TVars.table -> typ Same.operation val instantiate_same: typ TVars.table * term Vars.table -> term Same.operation val instantiateT: typ TVars.table -> typ -> typ val instantiate: typ TVars.table * term Vars.table -> term -> term val zero_var_indexes_inst: Name.context -> term list -> typ TVars.table * term Vars.table val zero_var_indexes: term -> term val zero_var_indexes_list: term list -> term list end; structure Term_Subst: TERM_SUBST = struct (* instantiation as table *) structure TFrees = Inst_Table( type key = string * sort val ord = pointer_eq_ord (prod_ord fast_string_ord Term_Ord.sort_ord) ); structure TVars = Inst_Table( type key = indexname * sort val ord = pointer_eq_ord (prod_ord Term_Ord.fast_indexname_ord Term_Ord.sort_ord) ); structure Frees = Inst_Table( type key = string * typ val ord = pointer_eq_ord (prod_ord fast_string_ord Term_Ord.typ_ord) ); structure Vars = Inst_Table( type key = indexname * typ val ord = pointer_eq_ord (prod_ord Term_Ord.fast_indexname_ord Term_Ord.typ_ord) ); +val add_tfreesT = fold_atyps (fn TFree v => TFrees.add (v, ()) | _ => I); +val add_tfrees = fold_types add_tfreesT; +val add_tvarsT = fold_atyps (fn TVar v => TVars.add (v, ()) | _ => I); +val add_tvars = fold_types add_tvarsT; +val add_frees = fold_aterms (fn Free v => Frees.add (v, ()) | _ => I); +val add_vars = fold_aterms (fn Var v => Vars.add (v, ()) | _ => I); +val add_tfree_namesT = fold_atyps (fn TFree (a, _) => Symtab.insert_set a | _ => I); +val add_tfree_names = fold_types add_tfree_namesT; +val add_free_names = fold_aterms (fn Free (x, _) => Symtab.insert_set x | _ => I); + (* generic mapping *) fun map_atypsT_same f = let fun typ (Type (a, Ts)) = Type (a, Same.map typ Ts) | typ T = f T; in typ end; fun map_types_same f = let fun term (Const (a, T)) = Const (a, f T) | term (Free (a, T)) = Free (a, f T) | term (Var (v, T)) = Var (v, f T) | term (Bound _) = raise Same.SAME | term (Abs (x, T, t)) = (Abs (x, f T, Same.commit term t) handle Same.SAME => Abs (x, T, term t)) | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u); in term end; fun map_aterms_same f = let fun term (Abs (x, T, t)) = Abs (x, T, term t) | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u) | term a = f a; in term end; (* generalization of fixed variables *) fun generalizeT_same tfrees idx ty = if Symtab.is_empty tfrees then raise Same.SAME else let fun gen (Type (a, Ts)) = Type (a, Same.map gen Ts) | gen (TFree (a, S)) = if Symtab.defined tfrees a then TVar ((a, idx), S) else raise Same.SAME | gen _ = raise Same.SAME; in gen ty end; fun generalize_same (tfrees, frees) idx tm = if Symtab.is_empty tfrees andalso Symtab.is_empty frees then raise Same.SAME else let val genT = generalizeT_same tfrees idx; fun gen (Free (x, T)) = if Symtab.defined frees x then Var (Name.clean_index (x, idx), Same.commit genT T) else Free (x, genT T) | gen (Var (xi, T)) = Var (xi, genT T) | gen (Const (c, T)) = Const (c, genT T) | gen (Bound _) = raise Same.SAME | gen (Abs (x, T, t)) = (Abs (x, genT T, Same.commit gen t) handle Same.SAME => Abs (x, T, gen t)) | gen (t $ u) = (gen t $ Same.commit gen u handle Same.SAME => t $ gen u); in gen tm end; fun generalizeT names i ty = Same.commit (generalizeT_same names i) ty; fun generalize names i tm = Same.commit (generalize_same names i) tm; (* instantiation of free variables (types before terms) *) fun instantiateT_frees_same instT ty = if TFrees.is_empty instT then raise Same.SAME else let fun subst (Type (a, Ts)) = Type (a, Same.map subst Ts) | subst (TFree v) = (case TFrees.lookup instT v of SOME T => T | NONE => raise Same.SAME) | subst _ = raise Same.SAME; in subst ty end; fun instantiate_frees_same (instT, inst) tm = if TFrees.is_empty instT andalso Frees.is_empty inst then raise Same.SAME else let val substT = instantiateT_frees_same instT; fun subst (Const (c, T)) = Const (c, substT T) | subst (Free (x, T)) = let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in (case Frees.lookup inst (x, T') of SOME t => t | NONE => if same then raise Same.SAME else Free (x, T')) end | subst (Var (xi, T)) = Var (xi, substT T) | subst (Bound _) = raise Same.SAME | subst (Abs (x, T, t)) = (Abs (x, substT T, Same.commit subst t) handle Same.SAME => Abs (x, T, subst t)) | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); in subst tm end; fun instantiateT_frees instT ty = Same.commit (instantiateT_frees_same instT) ty; fun instantiate_frees inst tm = Same.commit (instantiate_frees_same inst) tm; (* instantiation of schematic variables (types before terms) -- recomputes maxidx *) local fun no_indexesT instT = TVars.map (fn _ => rpair ~1) instT; fun no_indexes inst = Vars.map (fn _ => rpair ~1) inst; fun instT_same maxidx instT ty = let fun maxify i = if i > ! maxidx then maxidx := i else (); fun subst_typ (Type (a, Ts)) = Type (a, subst_typs Ts) | subst_typ (TVar ((a, i), S)) = (case TVars.lookup instT ((a, i), S) of SOME (T, j) => (maxify j; T) | NONE => (maxify i; raise Same.SAME)) | subst_typ _ = raise Same.SAME and subst_typs (T :: Ts) = (subst_typ T :: Same.commit subst_typs Ts handle Same.SAME => T :: subst_typs Ts) | subst_typs [] = raise Same.SAME; in subst_typ ty end; fun inst_same maxidx (instT, inst) tm = let fun maxify i = if i > ! maxidx then maxidx := i else (); val substT = instT_same maxidx instT; fun subst (Const (c, T)) = Const (c, substT T) | subst (Free (x, T)) = Free (x, substT T) | subst (Var ((x, i), T)) = let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in (case Vars.lookup inst ((x, i), T') of SOME (t, j) => (maxify j; t) | NONE => (maxify i; if same then raise Same.SAME else Var ((x, i), T'))) end | subst (Bound _) = raise Same.SAME | subst (Abs (x, T, t)) = (Abs (x, substT T, Same.commit subst t) handle Same.SAME => Abs (x, T, subst t)) | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); in subst tm end; in fun instantiateT_maxidx instT ty i = let val maxidx = Unsynchronized.ref i in (Same.commit (instT_same maxidx instT) ty, ! maxidx) end; fun instantiate_maxidx insts tm i = let val maxidx = Unsynchronized.ref i in (Same.commit (inst_same maxidx insts) tm, ! maxidx) end; fun instantiateT_same instT ty = if TVars.is_empty instT then raise Same.SAME else instT_same (Unsynchronized.ref ~1) (no_indexesT instT) ty; fun instantiate_same (instT, inst) tm = if TVars.is_empty instT andalso Vars.is_empty inst then raise Same.SAME else inst_same (Unsynchronized.ref ~1) (no_indexesT instT, no_indexes inst) tm; fun instantiateT instT ty = Same.commit (instantiateT_same instT) ty; fun instantiate inst tm = Same.commit (instantiate_same inst) tm; end; (* zero var indexes *) fun zero_var_inst ins mk (v as ((x, i), X)) (inst, used) = let val (x', used') = Name.variant (if Name.is_bound x then "u" else x) used in if x = x' andalso i = 0 then (inst, used') else (ins (v, mk ((x', 0), X)) inst, used') end; fun zero_var_indexes_inst used ts = let val (instT, _) = (TVars.empty, used) |> TVars.fold (zero_var_inst TVars.add TVar o #1) - ((fold o fold_types o fold_atyps) (fn TVar v => - TVars.add (v, ()) | _ => I) ts TVars.empty); + (TVars.build (ts |> (fold o fold_types o fold_atyps) + (fn TVar v => TVars.add (v, ()) | _ => I))); val (inst, _) = (Vars.empty, used) |> Vars.fold (zero_var_inst Vars.add Var o #1) - ((fold o fold_aterms) (fn Var (xi, T) => - Vars.add ((xi, instantiateT instT T), ()) | _ => I) ts Vars.empty); + (Vars.build (ts |> (fold o fold_aterms) + (fn Var (xi, T) => Vars.add ((xi, instantiateT instT T), ()) | _ => I))); in (instT, inst) end; fun zero_var_indexes_list ts = map (instantiate (zero_var_indexes_inst Name.context ts)) ts; val zero_var_indexes = singleton zero_var_indexes_list; end; diff --git a/src/Pure/theory.ML b/src/Pure/theory.ML --- a/src/Pure/theory.ML +++ b/src/Pure/theory.ML @@ -1,342 +1,342 @@ (* Title: Pure/theory.ML Author: Lawrence C Paulson and Markus Wenzel Logical theory content: axioms, definitions, and begin/end wrappers. *) signature THEORY = sig val parents_of: theory -> theory list val ancestors_of: theory -> theory list val nodes_of: theory -> theory list val setup: (theory -> theory) -> unit val local_setup: (Proof.context -> Proof.context) -> unit val install_pure: theory -> unit val get_pure: unit -> theory val get_pure_bootstrap: unit -> theory val get_markup: theory -> Markup.T val check: {long: bool} -> Proof.context -> string * Position.T -> theory val axiom_table: theory -> term Name_Space.table val axiom_space: theory -> Name_Space.T val all_axioms_of: theory -> (string * term) list val defs_of: theory -> Defs.T val at_begin: (theory -> theory option) -> theory -> theory val at_end: (theory -> theory option) -> theory -> theory val join_theory: theory list -> theory val begin_theory: string * Position.T -> theory list -> theory val end_theory: theory -> theory val add_axiom: Proof.context -> binding * term -> theory -> theory val const_dep: theory -> string * typ -> Defs.entry val type_dep: string * typ list -> Defs.entry val add_deps: Defs.context -> string -> Defs.entry -> Defs.entry list -> theory -> theory val add_deps_global: string -> Defs.entry -> Defs.entry list -> theory -> theory val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> theory val specify_const: (binding * typ) * mixfix -> theory -> term * theory val check_overloading: Proof.context -> bool -> string * typ -> unit end structure Theory: THEORY = struct (** theory context operations **) val parents_of = Context.parents_of; val ancestors_of = Context.ancestors_of; fun nodes_of thy = thy :: ancestors_of thy; fun setup f = Context.>> (Context.map_theory f); fun local_setup f = Context.>> (Context.map_proof f); (* implicit theory Pure *) val pure: theory Single_Assignment.var = Single_Assignment.var "pure"; fun install_pure thy = Single_Assignment.assign pure thy; fun get_pure () = (case Single_Assignment.peek pure of SOME thy => thy | NONE => raise Fail "Theory Pure not present"); fun get_pure_bootstrap () = (case Single_Assignment.peek pure of SOME thy => thy | NONE => Context.the_global_context ()); (** datatype thy **) type wrapper = (theory -> theory option) * stamp; fun apply_wrappers (wrappers: wrapper list) = perhaps (perhaps_loop (perhaps_apply (map fst wrappers))); datatype thy = Thy of {pos: Position.T, id: serial, axioms: term Name_Space.table, defs: Defs.T, wrappers: wrapper list * wrapper list}; fun make_thy (pos, id, axioms, defs, wrappers) = Thy {pos = pos, id = id, axioms = axioms, defs = defs, wrappers = wrappers}; structure Thy = Theory_Data' ( type T = thy; val extend = I; val empty = make_thy (Position.none, 0, Name_Space.empty_table Markup.axiomN, Defs.empty, ([], [])); fun merge old_thys (thy1, thy2) = let val Thy {pos, id, axioms = axioms1, defs = defs1, wrappers = (bgs1, ens1)} = thy1; val Thy {pos = _, id = _, axioms = axioms2, defs = defs2, wrappers = (bgs2, ens2)} = thy2; val axioms' = Name_Space.merge_tables (axioms1, axioms2); val defs' = Defs.merge (Defs.global_context (fst old_thys)) (defs1, defs2); val bgs' = Library.merge (eq_snd op =) (bgs1, bgs2); val ens' = Library.merge (eq_snd op =) (ens1, ens2); in make_thy (pos, id, axioms', defs', (bgs', ens')) end; ); fun rep_theory thy = Thy.get thy |> (fn Thy args => args); fun map_thy f = Thy.map (fn (Thy {pos, id, axioms, defs, wrappers}) => make_thy (f (pos, id, axioms, defs, wrappers))); fun map_axioms f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, f axioms, defs, wrappers)); fun map_defs f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, axioms, f defs, wrappers)); fun map_wrappers f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, axioms, defs, f wrappers)); (* entity markup *) fun theory_markup def name id pos = if id = 0 then Markup.empty else Position.make_entity_markup def id Markup.theoryN (name, pos); fun init_markup (name, pos) thy = let val id = serial (); val _ = Context_Position.reports_global thy [(pos, theory_markup true name id pos)]; in map_thy (fn (_, _, axioms, defs, wrappers) => (pos, id, axioms, defs, wrappers)) thy end; fun get_markup thy = let val {pos, id, ...} = rep_theory thy in theory_markup false (Context.theory_long_name thy) id pos end; fun check long ctxt (name, pos) = let val thy = Proof_Context.theory_of ctxt; val thy' = Context.get_theory long thy name handle ERROR msg => let val completion_report = Completion.make_report (name, pos) (fn completed => map (Context.theory_name' long) (ancestors_of thy) |> filter (completed o Long_Name.base_name) |> sort_strings |> map (fn a => (a, (Markup.theoryN, a)))); in error (msg ^ Position.here pos ^ completion_report) end; val _ = Context_Position.report ctxt pos (get_markup thy'); in thy' end; (* basic operations *) val axiom_table = #axioms o rep_theory; val axiom_space = Name_Space.space_of_table o axiom_table; val all_axioms_of = Name_Space.dest_table o axiom_table; val defs_of = #defs o rep_theory; (* join theory *) fun join_theory [] = raise List.Empty | join_theory [thy] = thy | join_theory thys = foldl1 Context.join_thys thys; (* begin/end theory *) val begin_wrappers = rev o #1 o #wrappers o rep_theory; val end_wrappers = rev o #2 o #wrappers o rep_theory; fun at_begin f = map_wrappers (apfst (cons (f, stamp ()))); fun at_end f = map_wrappers (apsnd (cons (f, stamp ()))); fun begin_theory (name, pos) imports = if name = Context.PureN then (case imports of [thy] => init_markup (name, pos) thy | _ => error "Bad bootstrapping of theory Pure") else let val thy = Context.begin_thy name imports; val wrappers = begin_wrappers thy; in thy |> init_markup (name, pos) |> Sign.init_naming |> Sign.local_path |> apply_wrappers wrappers |> tap (Syntax.force_syntax o Sign.syn_of) end; fun end_theory thy = thy |> apply_wrappers (end_wrappers thy) |> Sign.change_check |> Context.finish_thy; (** primitive specifications **) (* raw axioms *) fun cert_axm ctxt (b, raw_tm) = let val thy = Proof_Context.theory_of ctxt; val t = Sign.cert_prop thy raw_tm handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg; val _ = Term.no_dummy_patterns t handle TERM (msg, _) => error msg; val bad_sorts = rev ((fold_types o fold_atyps_sorts) (fn (_, []) => I | (T, _) => insert (op =) T) t []); val _ = null bad_sorts orelse error ("Illegal sort constraints in primitive specification: " ^ commas (map (Syntax.string_of_typ (Config.put show_sorts true ctxt)) bad_sorts)); in (b, Sign.no_vars ctxt t) end handle ERROR msg => cat_error msg ("The error(s) above occurred in axiom " ^ Binding.print b); fun add_axiom ctxt raw_axm thy = thy |> map_axioms (fn axioms => let val axm = apsnd Logic.varify_global (cert_axm ctxt raw_axm); val context = ctxt |> Sign.inherit_naming thy |> Context_Position.set_visible_generic false; val (_, axioms') = Name_Space.define context true axm axioms; in axioms' end); (* dependencies *) fun const_dep thy (c, T) = ((Defs.Const, c), Sign.const_typargs thy (c, T)); fun type_dep (c, args) = ((Defs.Type, c), args); fun dependencies (context as (ctxt, _)) unchecked def description lhs rhs = let fun prep (item, args) = (case fold Term.add_tvarsT args [] of [] => (item, map Logic.varifyT_global args) | vs => raise TYPE ("Illegal schematic type variable(s)", map TVar vs, [])); - val lhs_vars = fold Term.add_tfreesT (snd lhs) []; + val lhs_vars = Term_Subst.TFrees.build (snd lhs |> fold Term_Subst.add_tfreesT); val rhs_extras = - fold (fn (_, args) => args |> (fold o Term.fold_atyps) (fn TFree v => - if member (op =) lhs_vars v then I else insert (op =) v)) rhs []; + build (rhs |> fold (#2 #> (fold o Term.fold_atyps) + (fn TFree v => not (Term_Subst.TFrees.defined lhs_vars v) ? insert (op =) v | _ => I))); val _ = if null rhs_extras then () else error ("Specification depends on extra type variables: " ^ commas_quote (map (Syntax.string_of_typ ctxt o TFree) rhs_extras) ^ "\nThe error(s) above occurred in " ^ quote description); in Defs.define context unchecked def description (prep lhs) (map prep rhs) end; fun cert_entry thy ((Defs.Const, c), args) = Sign.cert_term thy (Const (c, Sign.const_instance thy (c, args))) |> dest_Const |> const_dep thy | cert_entry thy ((Defs.Type, c), args) = Sign.certify_typ thy (Type (c, args)) |> dest_Type |> type_dep; fun add_deps context a raw_lhs raw_rhs thy = let val (lhs as ((_, lhs_name), _)) :: rhs = map (cert_entry thy) (raw_lhs :: raw_rhs); val description = if a = "" then lhs_name ^ " axiom" else a; in thy |> map_defs (dependencies context false NONE description lhs rhs) end; fun add_deps_global a x y thy = add_deps (Defs.global_context thy) a x y thy; fun specify_const decl thy = let val (t, thy') = Sign.declare_const_global decl thy; in (t, add_deps_global "" (const_dep thy' (dest_Const t)) [] thy') end; (* overloading *) fun check_overloading ctxt overloaded (c, T) = let val thy = Proof_Context.theory_of ctxt; val declT = Sign.the_const_constraint thy c handle TYPE (msg, _, _) => error msg; val T' = Logic.varifyT_global T; fun message sorts txt = [Pretty.block [Pretty.str "Specification of constant ", Pretty.str c, Pretty.str " ::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ (Config.put show_sorts sorts ctxt) T)], Pretty.str txt] |> Pretty.chunks |> Pretty.string_of; in if Sign.typ_instance thy (declT, T') then () else if Type.raw_instance (declT, T') then error (message true "imposes additional sort constraints on the constant declaration") else if overloaded then () else error (message false "is strictly less general than the declared type (overloading required)") end; (* definitional axioms *) local fun check_def (context as (ctxt, _)) thy unchecked overloaded (b, tm) defs = let val name = Sign.full_name thy b; val ((lhs, rhs), _, _) = Primitive_Defs.dest_def ctxt {check_head = Term.is_Const, check_free_lhs = K true, check_free_rhs = K false, check_tfree = K false} tm handle TERM (msg, _) => error msg; val lhs_const = Term.dest_Const (Term.head_of lhs); val rhs_consts = fold_aterms (fn Const const => insert (op =) (const_dep thy const) | _ => I) rhs []; val rhs_types = (fold_types o fold_subtypes) (fn Type t => insert (op =) (type_dep t) | _ => I) rhs []; val rhs_deps = rhs_consts @ rhs_types; val _ = check_overloading ctxt overloaded lhs_const; in defs |> dependencies context unchecked (SOME name) name (const_dep thy lhs_const) rhs_deps end handle ERROR msg => cat_error msg (Pretty.string_of (Pretty.block [Pretty.str ("The error(s) above occurred in definition " ^ Binding.print b ^ ":"), Pretty.fbrk, Pretty.quote (Syntax.pretty_term ctxt tm)])); in fun add_def (context as (ctxt, _)) unchecked overloaded raw_axm thy = let val axm = cert_axm ctxt raw_axm in thy |> map_defs (check_def context thy unchecked overloaded axm) |> add_axiom ctxt axm end; end; end; diff --git a/src/Pure/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2373 +1,2377 @@ (* 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 fold_terms: {hyps: bool} -> (term -> 'a -> 'a) -> thm -> 'a -> 'a + val fold_atomic_ctyps: {hyps: bool} -> (typ -> bool) -> (ctyp -> 'a -> 'a) -> thm -> 'a -> 'a + val fold_atomic_cterms: {hyps: bool} -> (term -> bool) -> (cterm -> 'a -> 'a) -> thm -> 'a -> 'a val terms_of_tpairs: (term * term) list -> term list val full_prop_of: thm -> term val theory_id: thm -> Context.theory_id val theory_name: thm -> string val maxidx_of: thm -> int val maxidx_thm: thm -> int -> int val shyps_of: thm -> sort Ord_List.T val hyps_of: thm -> term list val prop_of: thm -> term val tpairs_of: thm -> (term * term) list val concl_of: thm -> term val prems_of: thm -> term list val nprems_of: thm -> int val no_prems: thm -> bool val major_prem_of: thm -> term val cprop_of: thm -> cterm val cprem_of: thm -> int -> cterm val cconcl_of: thm -> cterm val cprems_of: thm -> cterm list val chyps_of: thm -> cterm list exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option val theory_of_cterm: cterm -> theory val theory_of_thm: thm -> theory val trim_context_ctyp: ctyp -> ctyp val trim_context_cterm: cterm -> cterm val transfer_ctyp: theory -> ctyp -> ctyp val transfer_cterm: theory -> cterm -> cterm val transfer: theory -> thm -> thm val transfer': Proof.context -> thm -> thm val transfer'': Context.generic -> thm -> thm val join_transfer: theory -> thm -> thm val join_transfer_context: Proof.context * thm -> Proof.context * thm val renamed_prop: term -> thm -> thm val weaken: cterm -> thm -> thm val weaken_sorts: sort list -> cterm -> cterm val proof_bodies_of: thm list -> proof_body list val proof_body_of: thm -> proof_body val proof_of: thm -> proof val reconstruct_proof_of: thm -> Proofterm.proof val consolidate: thm list -> unit val expose_proofs: theory -> thm list -> unit val expose_proof: theory -> thm -> unit val future: thm future -> cterm -> thm val thm_deps: thm -> Proofterm.thm Ord_List.T val extra_shyps: thm -> sort list val strip_shyps: thm -> thm val derivation_closed: thm -> bool val derivation_name: thm -> string val derivation_id: thm -> Proofterm.thm_id option val raw_derivation_name: thm -> string val expand_name: thm -> Proofterm.thm_header -> string option val name_derivation: string * Position.T -> thm -> thm val close_derivation: Position.T -> thm -> thm val trim_context: thm -> thm val axiom: theory -> string -> thm val all_axioms_of: theory -> (string * thm) list val get_tags: thm -> Properties.T val map_tags: (Properties.T -> Properties.T) -> thm -> thm val norm_proof: thm -> thm val adjust_maxidx_thm: int -> thm -> thm (*type classes*) val the_classrel: theory -> class * class -> thm val the_arity: theory -> string * sort list * class -> thm val classrel_proof: theory -> class * class -> proof val arity_proof: theory -> string * sort list * class -> proof (*oracles*) val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory val oracle_space: theory -> Name_Space.T val pretty_oracle: Proof.context -> string -> Pretty.T val extern_oracles: bool -> Proof.context -> (Markup.T * xstring) list val check_oracle: Proof.context -> xstring * Position.T -> string (*inference rules*) val assume: cterm -> thm val implies_intr: cterm -> thm -> thm val implies_elim: thm -> thm -> thm val forall_intr: cterm -> thm -> thm val forall_elim: cterm -> thm -> thm val reflexive: cterm -> thm val symmetric: thm -> thm val transitive: thm -> thm -> thm val beta_conversion: bool -> conv val eta_conversion: conv val eta_long_conversion: conv val abstract_rule: string -> cterm -> thm -> thm val combination: thm -> thm -> thm val equal_intr: thm -> thm -> thm val equal_elim: thm -> thm -> thm val solve_constraints: thm -> thm val flexflex_rule: Proof.context option -> thm -> thm Seq.seq val generalize: Symtab.set * Symtab.set -> int -> thm -> thm val generalize_cterm: Symtab.set * Symtab.set -> int -> cterm -> cterm val generalize_ctyp: Symtab.set -> int -> ctyp -> ctyp val instantiate: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> thm -> thm val instantiate_cterm: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> cterm -> cterm val trivial: cterm -> thm val of_class: ctyp * class -> thm val unconstrainT: thm -> thm - val varifyT_global': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm + val varifyT_global': Term_Subst.TFrees.set -> thm -> ((string * sort) * indexname) list * thm val varifyT_global: thm -> thm val legacy_freezeT: thm -> thm val plain_prop_of: thm -> term val dest_state: thm * int -> (term * term) list * term list * term * term val lift_rule: cterm -> thm -> thm val incr_indexes: int -> thm -> thm val assumption: Proof.context option -> int -> thm -> thm Seq.seq val eq_assumption: int -> thm -> thm val rotate_rule: int -> int -> thm -> thm val permute_prems: int -> int -> thm -> thm val bicompose: Proof.context option -> {flatten: bool, match: bool, incremented: bool} -> bool * thm * int -> int -> thm -> thm Seq.seq val biresolution: Proof.context option -> bool -> (bool * thm) list -> int -> thm -> thm Seq.seq val thynames_of_arity: theory -> string * class -> string list val add_classrel: thm -> theory -> theory val add_arity: thm -> theory -> theory end; structure Thm: THM = struct (*** Certified terms and types ***) (** certified types **) datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: sort Ord_List.T}; fun typ_of (Ctyp {T, ...}) = T; fun global_ctyp_of thy raw_T = let val T = Sign.certify_typ thy raw_T; val maxidx = Term.maxidx_of_typ T; val sorts = Sorts.insert_typ T []; in Ctyp {cert = Context.Certificate thy, T = T, maxidx = maxidx, sorts = sorts} end; val ctyp_of = global_ctyp_of o Proof_Context.theory_of; fun dest_ctyp (Ctyp {cert, T = Type (_, Ts), maxidx, sorts}) = map (fn T => Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}) Ts | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []); fun dest_ctypN n (Ctyp {cert, T, maxidx, sorts}) = let fun err () = raise TYPE ("dest_ctypN", [T], []) in (case T of Type (_, Ts) => Ctyp {cert = cert, T = nth Ts n handle General.Subscript => err (), maxidx = maxidx, sorts = sorts} | _ => err ()) end; val dest_ctyp0 = dest_ctypN 0; val dest_ctyp1 = dest_ctypN 1; fun join_certificate_ctyp (Ctyp {cert, ...}) cert0 = Context.join_certificate (cert0, cert); fun union_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sorts.union sorts0 sorts; fun maxidx_ctyp (Ctyp {maxidx, ...}) maxidx0 = Int.max (maxidx0, maxidx); fun make_ctyp (Ctyp {cert, T, maxidx = _, sorts = _}) cargs = let val As = map typ_of cargs; fun err () = raise TYPE ("make_ctyp", T :: As, []); in (case T of Type (a, args) => Ctyp { cert = fold join_certificate_ctyp cargs cert, maxidx = fold maxidx_ctyp cargs ~1, sorts = fold union_sorts_ctyp cargs [], T = if length args = length cargs then Type (a, As) else err ()} | _ => err ()) end; (** certified terms **) (*certified terms with checked typ, maxidx, and sorts*) datatype cterm = Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: sort Ord_List.T}; exception CTERM of string * cterm list; fun term_of (Cterm {t, ...}) = t; fun typ_of_cterm (Cterm {T, ...}) = T; fun ctyp_of_cterm (Cterm {cert, T, maxidx, sorts, ...}) = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}; fun maxidx_of_cterm (Cterm {maxidx, ...}) = maxidx; fun global_cterm_of thy tm = let val (t, T, maxidx) = Sign.certify_term thy tm; val sorts = Sorts.insert_term t []; in Cterm {cert = Context.Certificate thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end; val cterm_of = global_cterm_of o Proof_Context.theory_of; fun join_certificate0 (Cterm {cert = cert1, ...}, Cterm {cert = cert2, ...}) = Context.join_certificate (cert1, cert2); fun renamed_term t' (Cterm {cert, t, T, maxidx, sorts}) = if t aconv t' then Cterm {cert = cert, t = t', T = T, maxidx = maxidx, sorts = sorts} else raise TERM ("renamed_term: terms disagree", [t, t']); (* destructors *) fun dest_comb (Cterm {t = c $ a, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in (Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_comb ct = raise CTERM ("dest_comb", [ct]); fun dest_fun (Cterm {t = c $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun ct = raise CTERM ("dest_fun", [ct]); fun dest_arg (Cterm {t = c $ a, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg ct = raise CTERM ("dest_arg", [ct]); fun dest_fun2 (Cterm {t = c $ _ $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0; val B = Term.argument_type_of c 1; in Cterm {t = c, T = A --> B --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun2 ct = raise CTERM ("dest_fun2", [ct]); fun dest_arg1 (Cterm {t = c $ a $ _, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]); fun dest_abs a (Cterm {t = Abs (x, T, t), T = Type ("fun", [_, U]), cert, maxidx, sorts}) = let val (y', t') = Term.dest_abs (the_default x a, T, t) in (Cterm {t = Free (y', T), T = T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = t', T = U, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_abs _ ct = raise CTERM ("dest_abs", [ct]); (* constructors *) fun rename_tvar (a, i) (Ctyp {cert, T, maxidx, sorts}) = let val S = (case T of TFree (_, S) => S | TVar (_, S) => S | _ => raise TYPE ("rename_tvar: no variable", [T], [])); val _ = if i < 0 then raise TYPE ("rename_tvar: bad index", [TVar ((a, i), S)], []) else (); in Ctyp {cert = cert, T = TVar ((a, i), S), maxidx = Int.max (i, maxidx), sorts = sorts} end; fun var ((x, i), Ctyp {cert, T, maxidx, sorts}) = if i < 0 then raise TERM ("var: bad index", [Var ((x, i), T)]) else Cterm {cert = cert, t = Var ((x, i), T), T = T, maxidx = Int.max (i, maxidx), sorts = sorts}; fun apply (cf as Cterm {t = f, T = Type ("fun", [dty, rty]), maxidx = maxidx1, sorts = sorts1, ...}) (cx as Cterm {t = x, T, maxidx = maxidx2, sorts = sorts2, ...}) = if T = dty then Cterm {cert = join_certificate0 (cf, cx), t = f $ x, T = rty, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} else raise CTERM ("apply: types don't agree", [cf, cx]) | apply cf cx = raise CTERM ("apply: first arg is not a function", [cf, cx]); fun lambda_name (x, ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...}) (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) = let val t = Term.lambda_name (x, t1) t2 in Cterm {cert = join_certificate0 (ct1, ct2), t = t, T = T1 --> T2, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} end; fun lambda t u = lambda_name ("", t) u; (* indexes *) fun adjust_maxidx_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if maxidx = i then ct else if maxidx < i then Cterm {maxidx = i, cert = cert, t = t, T = T, sorts = sorts} else Cterm {maxidx = Int.max (maxidx_of_term t, i), cert = cert, t = t, T = T, sorts = sorts}; fun incr_indexes_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if i < 0 then raise CTERM ("negative increment", [ct]) else if i = 0 then ct else Cterm {cert = cert, t = Logic.incr_indexes ([], [], i) t, T = Logic.incr_tvar i T, maxidx = maxidx + i, sorts = sorts}; (*** Derivations and Theorems ***) (* sort constraints *) type constraint = {theory: theory, typ: typ, sort: sort}; local val constraint_ord : constraint ord = Context.theory_id_ord o apply2 (Context.theory_id o #theory) ||| Term_Ord.typ_ord o apply2 #typ ||| Term_Ord.sort_ord o apply2 #sort; val smash_atyps = map_atyps (fn TVar (_, S) => Term.aT S | TFree (_, S) => Term.aT S | T => T); in val union_constraints = Ord_List.union constraint_ord; fun insert_constraints thy (T, S) = let val ignored = S = [] orelse (case T of TFree (_, S') => S = S' | TVar (_, S') => S = S' | _ => false); in if ignored then I else Ord_List.insert constraint_ord {theory = thy, typ = smash_atyps T, sort = S} end; fun insert_constraints_env thy env = let val tyenv = Envir.type_env env; fun insert ([], _) = I | insert (S, T) = insert_constraints thy (Envir.norm_type tyenv T, S); in tyenv |> Vartab.fold (insert o #2) end; end; (* datatype thm *) datatype thm = Thm of deriv * (*derivation*) {cert: Context.certificate, (*background theory certificate*) tags: Properties.T, (*additional annotations/comments*) maxidx: int, (*maximum index of any Var or TVar*) constraints: constraint Ord_List.T, (*implicit proof obligations for sort constraints*) shyps: sort Ord_List.T, (*sort hypotheses*) hyps: term Ord_List.T, (*hypotheses*) tpairs: (term * term) list, (*flex-flex pairs*) prop: term} (*conclusion*) and deriv = Deriv of {promises: (serial * thm future) Ord_List.T, body: Proofterm.proof_body}; type conv = cterm -> thm; (*errors involving theorems*) exception THM of string * int * thm list; fun rep_thm (Thm (_, args)) = args; -fun fold_terms f (Thm (_, {tpairs, prop, hyps, ...})) = - fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps; +fun fold_terms h f (Thm (_, {tpairs, prop, hyps, ...})) = + fold (fn (t, u) => f t #> f u) tpairs #> f prop #> #hyps h ? fold f hyps; -fun fold_atomic_ctyps f (th as Thm (_, {cert, maxidx, shyps, ...})) = +fun fold_atomic_ctyps h g 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; + in (fold_terms h o fold_types o fold_atyps) (fn T => if g T then f (ctyp T) else I) 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) +fun fold_atomic_cterms h g f (th as Thm (_, {cert, maxidx, shyps, ...})) = + let + fun cterm t T = Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = shyps}; + fun apply t T = if g t then f (cterm t T) else I; + in + (fold_terms h o fold_aterms) + (fn t as Const (_, T) => apply t T + | t as Free (_, T) => apply t T + | t as Var (_, T) => apply t T | _ => I) th end; fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs []; fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u'; fun union_tpairs ts us = Library.merge eq_tpairs (ts, us); val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u); fun attach_tpairs tpairs prop = Logic.list_implies (map Logic.mk_equals tpairs, prop); fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop; val union_hyps = Ord_List.union Term_Ord.fast_term_ord; val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord; val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord; fun join_certificate1 (Cterm {cert = cert1, ...}, Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); fun join_certificate2 (Thm (_, {cert = cert1, ...}), Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); (* basic components *) val cert_of = #cert o rep_thm; val theory_id = Context.certificate_theory_id o cert_of; val theory_name = Context.theory_id_name o theory_id; val maxidx_of = #maxidx o rep_thm; fun maxidx_thm th i = Int.max (maxidx_of th, i); val shyps_of = #shyps o rep_thm; val hyps_of = #hyps o rep_thm; val prop_of = #prop o rep_thm; val tpairs_of = #tpairs o rep_thm; val concl_of = Logic.strip_imp_concl o prop_of; val prems_of = Logic.strip_imp_prems o prop_of; val nprems_of = Logic.count_prems o prop_of; fun no_prems th = nprems_of th = 0; fun major_prem_of th = (case prems_of th of prem :: _ => Logic.strip_assums_concl prem | [] => raise THM ("major_prem_of: rule with no premises", 0, [th])); fun cprop_of (Thm (_, {cert, maxidx, shyps, prop, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, t = prop, sorts = shyps}; fun cprem_of (th as Thm (_, {cert, maxidx, shyps, prop, ...})) i = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])}; fun cconcl_of (th as Thm (_, {cert, maxidx, shyps, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = concl_of th}; fun cprems_of (th as Thm (_, {cert, maxidx, shyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = t}) (prems_of th); fun chyps_of (Thm (_, {cert, shyps, hyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t}) hyps; (* implicit theory context *) exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option; fun theory_of_cterm (ct as Cterm {cert, ...}) = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct], [], NONE); fun theory_of_thm th = Context.certificate_theory (cert_of th) handle ERROR msg => raise CONTEXT (msg, [], [], [th], NONE); fun trim_context_ctyp cT = (case cT of Ctyp {cert = Context.Certificate_Id _, ...} => cT | Ctyp {cert = Context.Certificate thy, T, maxidx, sorts} => Ctyp {cert = Context.Certificate_Id (Context.theory_id thy), T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_cterm ct = (case ct of Cterm {cert = Context.Certificate_Id _, ...} => ct | Cterm {cert = Context.Certificate thy, t, T, maxidx, sorts} => Cterm {cert = Context.Certificate_Id (Context.theory_id thy), t = t, T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_thm th = (case th of Thm (_, {constraints = _ :: _, ...}) => raise THM ("trim_context: pending sort constraints", 0, [th]) | Thm (_, {cert = Context.Certificate_Id _, ...}) => th | Thm (der, {cert = Context.Certificate thy, tags, maxidx, constraints = [], shyps, hyps, tpairs, prop}) => Thm (der, {cert = Context.Certificate_Id (Context.theory_id thy), tags = tags, maxidx = maxidx, constraints = [], shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})); fun transfer_ctyp thy' cT = let val Ctyp {cert, T, maxidx, sorts} = cT; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [cT], [], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then cT else Ctyp {cert = cert', T = T, maxidx = maxidx, sorts = sorts} end; fun transfer_cterm thy' ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [ct], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then ct else Cterm {cert = cert', t = t, T = T, maxidx = maxidx, sorts = sorts} end; fun transfer thy' th = let val Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) = th; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [], [th], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then th else Thm (der, {cert = cert', tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; val transfer' = transfer o Proof_Context.theory_of; val transfer'' = transfer o Context.theory_of; fun join_transfer thy th = (Context.subthy_id (theory_id th, Context.theory_id thy) ? transfer thy) th; fun join_transfer_context (ctxt, th) = if Context.subthy_id (theory_id th, Context.theory_id (Proof_Context.theory_of ctxt)) then (ctxt, transfer' ctxt th) else (Context.raw_transfer (theory_of_thm th) ctxt, th); (* matching *) local fun gen_match match (ct1 as Cterm {t = t1, sorts = sorts1, ...}, ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) = let val cert = join_certificate0 (ct1, ct2); val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct1, ct2], [], NONE); val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty); val sorts = Sorts.union sorts1 sorts2; fun mk_cTinst ((a, i), (S, T)) = (((a, i), S), Ctyp {T = T, cert = cert, maxidx = maxidx2, sorts = sorts}); fun mk_ctinst ((x, i), (U, t)) = let val T = Envir.subst_type Tinsts U in (((x, i), T), Cterm {t = t, T = T, cert = cert, maxidx = maxidx2, sorts = sorts}) end; in (Vartab.fold (cons o mk_cTinst) Tinsts [], Vartab.fold (cons o mk_ctinst) tinsts []) end; in val match = gen_match Pattern.match; val first_order_match = gen_match Pattern.first_order_match; end; (*implicit alpha-conversion*) fun renamed_prop prop' (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if prop aconv prop' then Thm (der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) else raise TERM ("renamed_prop: props disagree", [prop, prop']); fun make_context ths NONE cert = (Context.Theory (Context.certificate_theory cert) handle ERROR msg => raise CONTEXT (msg, [], [], ths, NONE)) | make_context ths (SOME ctxt) cert = let val thy_id = Context.certificate_theory_id cert; val thy_id' = Context.theory_id (Proof_Context.theory_of ctxt); in if Context.subthy_id (thy_id, thy_id') then Context.Proof ctxt else raise CONTEXT ("Bad context", [], [], ths, SOME (Context.Proof ctxt)) end; fun make_context_certificate ths opt_ctxt cert = let val context = make_context ths opt_ctxt cert; val cert' = Context.Certificate (Context.theory_of context); in (context, cert') end; (*explicit weakening: maps |- B to A |- B*) fun weaken raw_ct th = let val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct; val Thm (der, {tags, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; in if T <> propT then raise THM ("weaken: assumptions must have type prop", 0, []) else if maxidxA <> ~1 then raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, []) else Thm (der, {cert = join_certificate1 (ct, th), tags = tags, maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = insert_hyps A hyps, tpairs = tpairs, prop = prop}) end; fun weaken_sorts raw_sorts ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val thy = theory_of_cterm ct; val more_sorts = Sorts.make (map (Sign.certify_sort thy) raw_sorts); val sorts' = Sorts.union sorts more_sorts; in Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = sorts'} end; (** derivations and promised proofs **) fun make_deriv promises oracles thms proof = Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}}; val empty_deriv = make_deriv [] [] [] MinProof; (* inference rules *) val promise_ord: (serial * thm future) ord = fn ((i, _), (j, _)) => int_ord (j, i); fun bad_proofs i = error ("Illegal level of detail for proof objects: " ^ string_of_int i); fun deriv_rule2 f (Deriv {promises = ps1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}}) (Deriv {promises = ps2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) = let val ps = Ord_List.union promise_ord ps1 ps2; val oracles = Proofterm.unions_oracles [oracles1, oracles2]; val thms = Proofterm.unions_thms [thms1, thms2]; val prf = (case ! Proofterm.proofs of 2 => f prf1 prf2 | 1 => MinProof | 0 => MinProof | i => bad_proofs i); in make_deriv ps oracles thms prf end; fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv; fun deriv_rule0 make_prf = if ! Proofterm.proofs <= 1 then empty_deriv else deriv_rule1 I (make_deriv [] [] [] (make_prf ())); fun deriv_rule_unconditional f (Deriv {promises, body = PBody {oracles, thms, proof}}) = make_deriv promises oracles thms (f proof); (* fulfilled proofs *) fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises; fun join_promises [] = () | join_promises promises = join_promises_of (Future.joins (map snd promises)) and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms)); fun fulfill_body (th as Thm (Deriv {promises, body}, _)) = let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises)) in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end; fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms); val proof_body_of = singleton proof_bodies_of; val proof_of = Proofterm.proof_of o proof_body_of; fun reconstruct_proof_of thm = Proofterm.reconstruct_proof (theory_of_thm thm) (prop_of thm) (proof_of thm); val consolidate = ignore o proof_bodies_of; fun expose_proofs thy thms = if Proofterm.export_proof_boxes_required thy then Proofterm.export_proof_boxes (proof_bodies_of (map (transfer thy) thms)) else (); fun expose_proof thy = expose_proofs thy o single; (* future rule *) fun future_result i orig_cert orig_shyps orig_prop thm = let fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]); val Thm (Deriv {promises, ...}, {cert, constraints, shyps, hyps, tpairs, prop, ...}) = thm; val _ = Context.eq_certificate (cert, orig_cert) orelse err "bad theory"; val _ = prop aconv orig_prop orelse err "bad prop"; val _ = null constraints orelse err "bad sort constraints"; val _ = null tpairs orelse err "bad flex-flex constraints"; val _ = null hyps orelse err "bad hyps"; val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps"; val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies"; val _ = join_promises promises; in thm end; fun future future_thm ct = let val Cterm {cert = cert, t = prop, T, maxidx, sorts} = ct; val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]); val _ = if Proofterm.proofs_enabled () then raise CTERM ("future: proof terms enabled", [ct]) else (); val i = serial (); val future = future_thm |> Future.map (future_result i cert sorts prop); in Thm (make_deriv [(i, future)] [] [] MinProof, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end; (** Axioms **) fun axiom thy name = (case Name_Space.lookup (Theory.axiom_table thy) name of SOME prop => let val der = deriv_rule0 (fn () => Proofterm.axm_proof name prop); val cert = Context.Certificate thy; val maxidx = maxidx_of_term prop; val shyps = Sorts.insert_term prop []; in Thm (der, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = shyps, hyps = [], tpairs = [], prop = prop}) end | NONE => raise THEORY ("No axiom " ^ quote name, [thy])); fun all_axioms_of thy = map (fn (name, _) => (name, axiom thy name)) (Theory.all_axioms_of thy); (* tags *) val get_tags = #tags o rep_thm; fun map_tags f (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = Thm (der, {cert = cert, tags = f tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (* technical adjustments *) fun norm_proof (th as Thm (der, args)) = Thm (deriv_rule1 (Proofterm.rew_proof (theory_of_thm th)) der, args); fun adjust_maxidx_thm i (th as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if maxidx = i then th else if maxidx < i then Thm (der, {maxidx = i, cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) else Thm (der, {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i), cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (*** Theory data ***) (* type classes *) structure Aritytab = Table( type key = string * sort list * class; val ord = fast_string_ord o apply2 #1 ||| fast_string_ord o apply2 #3 ||| list_ord Term_Ord.sort_ord o apply2 #2; ); datatype classes = Classes of {classrels: thm Symreltab.table, arities: (thm * string * serial) Aritytab.table}; fun make_classes (classrels, arities) = Classes {classrels = classrels, arities = arities}; val empty_classes = make_classes (Symreltab.empty, Aritytab.empty); (*see Theory.at_begin hook for transitive closure of classrels and arity completion*) fun merge_classes (Classes {classrels = classrels1, arities = arities1}, Classes {classrels = classrels2, arities = arities2}) = let val classrels' = Symreltab.merge (K true) (classrels1, classrels2); val arities' = Aritytab.merge (K true) (arities1, arities2); in make_classes (classrels', arities') end; (* data *) structure Data = Theory_Data ( type T = unit Name_Space.table * (*oracles: authentic derivation names*) classes; (*type classes within the logic*) val empty : T = (Name_Space.empty_table Markup.oracleN, empty_classes); val extend = I; fun merge ((oracles1, sorts1), (oracles2, sorts2)) : T = (Name_Space.merge_tables (oracles1, oracles2), merge_classes (sorts1, sorts2)); ); val get_oracles = #1 o Data.get; val map_oracles = Data.map o apfst; val get_classes = (fn (_, Classes args) => args) o Data.get; val get_classrels = #classrels o get_classes; val get_arities = #arities o get_classes; fun map_classes f = (Data.map o apsnd) (fn Classes {classrels, arities} => make_classes (f (classrels, arities))); fun map_classrels f = map_classes (fn (classrels, arities) => (f classrels, arities)); fun map_arities f = map_classes (fn (classrels, arities) => (classrels, f arities)); (* type classes *) fun the_classrel thy (c1, c2) = (case Symreltab.lookup (get_classrels thy) (c1, c2) of SOME thm => transfer thy thm | NONE => error ("Unproven class relation " ^ Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2])); fun the_arity thy (a, Ss, c) = (case Aritytab.lookup (get_arities thy) (a, Ss, c) of SOME (thm, _, _) => transfer thy thm | NONE => error ("Unproven type arity " ^ Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c]))); val classrel_proof = proof_of oo the_classrel; val arity_proof = proof_of oo the_arity; (* solve sort constraints by pro-forma proof *) local fun union_digest (oracles1, thms1) (oracles2, thms2) = (Proofterm.unions_oracles [oracles1, oracles2], Proofterm.unions_thms [thms1, thms2]); fun thm_digest (Thm (Deriv {body = PBody {oracles, thms, ...}, ...}, _)) = (oracles, thms); fun constraint_digest ({theory = thy, typ, sort, ...}: constraint) = Sorts.of_sort_derivation (Sign.classes_of thy) {class_relation = fn _ => fn _ => fn (digest, c1) => fn c2 => if c1 = c2 then ([], []) else union_digest digest (thm_digest (the_classrel thy (c1, c2))), type_constructor = fn (a, _) => fn dom => fn c => let val arity_digest = thm_digest (the_arity thy (a, (map o map) #2 dom, c)) in (fold o fold) (union_digest o #1) dom arity_digest end, type_variable = fn T => map (pair ([], [])) (Type.sort_of_atyp T)} (typ, sort); in fun solve_constraints (thm as Thm (_, {constraints = [], ...})) = thm | solve_constraints (thm as Thm (der, args)) = let val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args; val thy = Context.certificate_theory cert; val bad_thys = constraints |> map_filter (fn {theory = thy', ...} => if Context.eq_thy (thy, thy') then NONE else SOME thy'); val () = if null bad_thys then () else raise THEORY ("solve_constraints: bad theories for theorem\n" ^ Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys); val Deriv {promises, body = PBody {oracles, thms, proof}} = der; val (oracles', thms') = (oracles, thms) |> fold (fold union_digest o constraint_digest) constraints; val body' = PBody {oracles = oracles', thms = thms', proof = proof}; in Thm (Deriv {promises = promises, body = body'}, {constraints = [], cert = cert, tags = tags, maxidx = maxidx, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; end; (*Dangling sort constraints of a thm*) fun extra_shyps (th as Thm (_, {shyps, ...})) = - Sorts.subtract (fold_terms Sorts.insert_term th []) shyps; + Sorts.subtract (fold_terms {hyps = true} Sorts.insert_term th []) shyps; (*Remove extra sorts that are witnessed by type signature information*) fun strip_shyps thm = (case thm of Thm (_, {shyps = [], ...}) => thm | Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) => let val thy = theory_of_thm thm; val algebra = Sign.classes_of thy; val minimize = Sorts.minimize_sort algebra; val le = Sorts.sort_le algebra; fun lt (S1, S2) = le (S1, S2) andalso not (le (S2, S1)); fun rel (S1, S2) = if S1 = S2 then [] else [(Term.aT S1, S2)]; - val present = (fold_terms o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm []; + val present = + (fold_terms {hyps = true} o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm []; val extra = fold (Sorts.remove_sort o #2) present shyps; val witnessed = Sign.witness_sorts thy present extra; val non_witnessed = fold (Sorts.remove_sort o #2) witnessed extra |> map (`minimize); val extra' = non_witnessed |> map_filter (fn (S, _) => if non_witnessed |> exists (fn (S', _) => lt (S', S)) then NONE else SOME S) |> Sorts.make; val constrs' = non_witnessed |> maps (fn (S1, S2) => let val S0 = the (find_first (fn S => le (S, S1)) extra') in rel (S0, S1) @ rel (S1, S2) end); val constraints' = fold (insert_constraints thy) (witnessed @ constrs') constraints; val shyps' = fold (Sorts.insert_sort o #2) present extra'; in Thm (deriv_rule_unconditional (Proofterm.strip_shyps_proof algebra present witnessed extra') der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop}) end) |> solve_constraints; (*** Closed theorems with official name ***) (*non-deterministic, depends on unknown promises*) fun derivation_closed (Thm (Deriv {body, ...}, _)) = Proofterm.compact_proof (Proofterm.proof_of body); (*non-deterministic, depends on unknown promises*) fun raw_derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (Proofterm.proof_of body); fun expand_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = let val self_id = (case Proofterm.get_identity shyps hyps prop (Proofterm.proof_of body) of NONE => K false | SOME {serial, ...} => fn (header: Proofterm.thm_header) => serial = #serial header); fun expand header = if self_id header orelse #name header = "" then SOME "" else NONE; in expand end; (*deterministic name of finished proof*) fun derivation_name (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (proof_of thm); (*identified PThm node*) fun derivation_id (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_id shyps hyps prop (proof_of thm); (*dependencies of PThm node*) fun thm_deps (thm as Thm (Deriv {promises = [], body = PBody {thms, ...}, ...}, _)) = (case (derivation_id thm, thms) of (SOME {serial = i, ...}, [(j, thm_node)]) => if i = j then Proofterm.thm_node_thms thm_node else thms | _ => thms) | thm_deps thm = raise THM ("thm_deps: bad promises", 0, [thm]); fun name_derivation name_pos = strip_shyps #> (fn thm as Thm (der, args) => let val thy = theory_of_thm thm; val Deriv {promises, body} = der; val {shyps, hyps, prop, tpairs, ...} = args; val _ = null tpairs orelse raise THM ("name_derivation: bad flex-flex constraints", 0, [thm]); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.thm_proof thy (classrel_proof thy) (arity_proof thy) name_pos shyps hyps prop ps body; val der' = make_deriv [] [] [pthm] proof; in Thm (der', args) end); fun close_derivation pos = solve_constraints #> (fn thm => if not (null (tpairs_of thm)) orelse derivation_closed thm then thm else name_derivation ("", pos) thm); val trim_context = solve_constraints #> trim_context_thm; (*** Oracles ***) fun add_oracle (b, oracle_fn) thy = let val (name, oracles') = Name_Space.define (Context.Theory thy) true (b, ()) (get_oracles thy); val thy' = map_oracles (K oracles') thy; fun invoke_oracle arg = let val Cterm {cert = cert2, t = prop, T, maxidx, sorts} = oracle_fn arg in if T <> propT then raise THM ("Oracle's result must have type prop: " ^ name, 0, []) else let val (oracle, prf) = (case ! Proofterm.proofs of 2 => (((name, Position.thread_data ()), SOME prop), Proofterm.oracle_proof name prop) | 1 => (((name, Position.thread_data ()), SOME prop), MinProof) | 0 => (((name, Position.none), NONE), MinProof) | i => bad_proofs i); in Thm (make_deriv [] [oracle] [] prf, {cert = Context.join_certificate (Context.Certificate thy', cert2), tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end end; in ((name, invoke_oracle), thy') end; val oracle_space = Name_Space.space_of_table o get_oracles; fun pretty_oracle ctxt = Name_Space.pretty ctxt (oracle_space (Proof_Context.theory_of ctxt)); fun extern_oracles verbose ctxt = map #1 (Name_Space.markup_table verbose ctxt (get_oracles (Proof_Context.theory_of ctxt))); fun check_oracle ctxt = Name_Space.check (Context.Proof ctxt) (get_oracles (Proof_Context.theory_of ctxt)) #> #1; (*** Meta rules ***) (** primitive rules **) (*The assumption rule A |- A*) fun assume raw_ct = let val Cterm {cert, t = prop, T, maxidx, sorts} = adjust_maxidx_cterm ~1 raw_ct in if T <> propT then raise THM ("assume: prop", 0, []) else if maxidx <> ~1 then raise THM ("assume: variables", maxidx, []) else Thm (deriv_rule0 (fn () => Proofterm.Hyp prop), {cert = cert, tags = [], maxidx = ~1, constraints = [], shyps = sorts, hyps = [prop], tpairs = [], prop = prop}) end; (*Implication introduction [A] : B ------- A \ B *) fun implies_intr (ct as Cterm {t = A, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...})) = if T <> propT then raise THM ("implies_intr: assumptions must have type prop", 0, [th]) else Thm (deriv_rule1 (Proofterm.implies_intr_proof A) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = remove_hyps A hyps, tpairs = tpairs, prop = Logic.mk_implies (A, prop)}); (*Implication elimination A \ B A ------------ B *) fun implies_elim thAB thA = let val Thm (derA, {maxidx = maxidx1, hyps = hypsA, constraints = constraintsA, shyps = shypsA, tpairs = tpairsA, prop = propA, ...}) = thA and Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...}) = thAB; fun err () = raise THM ("implies_elim: major premise", 0, [thAB, thA]); in (case prop of Const ("Pure.imp", _) $ A $ B => if A aconv propA then Thm (deriv_rule2 (curry Proofterm.%%) der derA, {cert = join_certificate2 (thAB, thA), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraintsA constraints, shyps = Sorts.union shypsA shyps, hyps = union_hyps hypsA hyps, tpairs = union_tpairs tpairsA tpairs, prop = B}) else err () | _ => err ()) end; (*Forall introduction. The Free or Var x must not be free in the hypotheses. [x] : A ------ \x. A *) fun forall_intr (ct as Cterm {maxidx = maxidx1, t = x, T, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = let fun result a = Thm (deriv_rule1 (Proofterm.forall_intr_proof (a, x) NONE) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.all_const T $ Abs (a, T, abstract_over (x, prop))}); fun check_occs a x ts = if exists (fn t => Logic.occs (x, t)) ts then raise THM ("forall_intr: variable " ^ quote a ^ " free in assumptions", 0, [th]) else (); in (case x of Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result a) | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result a) | _ => raise THM ("forall_intr: not a variable", 0, [th])) end; (*Forall elimination \x. A ------ A[t/x] *) fun forall_elim (ct as Cterm {t, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of Const ("Pure.all", Type ("fun", [Type ("fun", [qary, _]), _])) $ A => if T <> qary then raise THM ("forall_elim: type mismatch", 0, [th]) else Thm (deriv_rule1 (Proofterm.% o rpair (SOME t)) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Term.betapply (A, t)}) | _ => raise THM ("forall_elim: not quantified", 0, [th])); (* Equality *) (*Reflexivity t \ t *) fun reflexive (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t)}); (*Symmetry t \ u ------ u \ t *) fun symmetric (th as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of (eq as Const ("Pure.eq", _)) $ t $ u => Thm (deriv_rule1 Proofterm.symmetric_proof der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = eq $ u $ t}) | _ => raise THM ("symmetric", 0, [th])); (*Transitivity t1 \ u u \ t2 ------------------ t1 \ t2 *) fun transitive th1 th2 = let val Thm (der1, {maxidx = maxidx1, hyps = hyps1, constraints = constraints1, shyps = shyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, hyps = hyps2, constraints = constraints2, shyps = shyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("transitive: " ^ msg, 0, [th1, th2]); in case (prop1, prop2) of ((eq as Const ("Pure.eq", Type (_, [U, _]))) $ t1 $ u, Const ("Pure.eq", _) $ u' $ t2) => if not (u aconv u') then err "middle term" else Thm (deriv_rule2 (Proofterm.transitive_proof U u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = eq $ t1 $ t2}) | _ => err "premises" end; (*Beta-conversion (\x. t) u \ t[u/x] fully beta-reduces the term if full = true *) fun beta_conversion full (Cterm {cert, t, T = _, maxidx, sorts}) = let val t' = if full then Envir.beta_norm t else (case t of Abs (_, _, bodt) $ u => subst_bound (u, bodt) | _ => raise THM ("beta_conversion: not a redex", 0, [])); in Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t')}) end; fun eta_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_contract t)}); fun eta_long_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_long [] t)}); (*The abstraction rule. The Free or Var x must not be free in the hypotheses. The bound variable will be named "a" (since x will be something like x320) t \ u -------------- \x. t \ \x. u *) fun abstract_rule a (Cterm {t = x, T, sorts, ...}) (th as Thm (der, {cert, maxidx, hyps, constraints, shyps, tpairs, prop, ...})) = let val (t, u) = Logic.dest_equals prop handle TERM _ => raise THM ("abstract_rule: premise not an equality", 0, [th]); val result = Thm (deriv_rule1 (Proofterm.abstract_rule_proof (a, x)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.mk_equals (Abs (a, T, abstract_over (x, t)), Abs (a, T, abstract_over (x, u)))}); fun check_occs a x ts = if exists (fn t => Logic.occs (x, t)) ts then raise THM ("abstract_rule: variable " ^ quote a ^ " free in assumptions", 0, [th]) else (); in (case x of Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result) | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result) | _ => raise THM ("abstract_rule: not a variable", 0, [th])) end; (*The combination rule f \ g t \ u ------------- f t \ g u *) fun combination th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun chktypes fT tT = (case fT of Type ("fun", [T1, _]) => if T1 <> tT then raise THM ("combination: types", 0, [th1, th2]) else () | _ => raise THM ("combination: not function type", 0, [th1, th2])); in (case (prop1, prop2) of (Const ("Pure.eq", Type ("fun", [fT, _])) $ f $ g, Const ("Pure.eq", Type ("fun", [tT, _])) $ t $ u) => (chktypes fT tT; Thm (deriv_rule2 (Proofterm.combination_proof f g t u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (f $ t, g $ u)})) | _ => raise THM ("combination: premises", 0, [th1, th2])) end; (*Equality introduction A \ B B \ A ---------------- A \ B *) fun equal_intr th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_intr: " ^ msg, 0, [th1, th2]); in (case (prop1, prop2) of (Const("Pure.imp", _) $ A $ B, Const("Pure.imp", _) $ B' $ A') => if A aconv A' andalso B aconv B' then Thm (deriv_rule2 (Proofterm.equal_intr_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (A, B)}) else err "not equal" | _ => err "premises") end; (*The equal propositions rule A \ B A --------- B *) fun equal_elim th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_elim: " ^ msg, 0, [th1, th2]); in (case prop1 of Const ("Pure.eq", _) $ A $ B => if prop2 aconv A then Thm (deriv_rule2 (Proofterm.equal_elim_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = B}) else err "not equal" | _ => err "major premise") end; (**** Derived rules ****) (*Smash unifies the list of term pairs leaving no flex-flex pairs. Instantiates the theorem and deletes trivial tpairs. Resulting sequence may contain multiple elements if the tpairs are not all flex-flex.*) fun flexflex_rule opt_ctxt = solve_constraints #> (fn th => let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val (context, cert') = make_context_certificate [th] opt_ctxt cert; in Unify.smash_unifiers context tpairs (Envir.empty maxidx) |> Seq.map (fn env => if Envir.is_empty env then th else let val tpairs' = tpairs |> map (apply2 (Envir.norm_term env)) (*remove trivial tpairs, of the form t \ t*) |> filter_out (op aconv); val der' = deriv_rule1 (Proofterm.norm_proof' env) der; val constraints' = insert_constraints_env (Context.certificate_theory cert') env constraints; val prop' = Envir.norm_term env prop; val maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop'); val shyps = Envir.insert_sorts env shyps; in Thm (der', {cert = cert', tags = [], maxidx = maxidx, constraints = constraints', shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end) end); (*Generalization of fixed variables A -------------------- A[?'a/'a, ?x/x, ...] *) fun generalize (tfrees, frees) idx th = if Symtab.is_empty tfrees andalso Symtab.is_empty frees then th else let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]); val bad_type = if Symtab.is_empty tfrees then K false else Term.exists_subtype (fn TFree (a, _) => Symtab.defined tfrees a | _ => false); fun bad_term (Free (x, T)) = bad_type T orelse Symtab.defined frees x | bad_term (Var (_, T)) = bad_type T | bad_term (Const (_, T)) = bad_type T | bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t | bad_term (t $ u) = bad_term t orelse bad_term u | bad_term (Bound _) = false; val _ = exists bad_term hyps andalso raise THM ("generalize: variable free in assumptions", 0, [th]); val generalize = Term_Subst.generalize (tfrees, frees) idx; val prop' = generalize prop; val tpairs' = map (apply2 generalize) tpairs; val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop'); in Thm (deriv_rule1 (Proofterm.generalize_proof (tfrees, frees) idx prop) der, {cert = cert, tags = [], maxidx = maxidx', constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end; fun generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) = if Symtab.is_empty tfrees andalso Symtab.is_empty frees then ct else if idx <= maxidx then raise CTERM ("generalize_cterm: bad index", [ct]) else Cterm {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, t = Term_Subst.generalize (tfrees, frees) idx t, maxidx = Int.max (maxidx, idx)}; fun generalize_ctyp tfrees idx (cT as Ctyp {cert, T, maxidx, sorts}) = if Symtab.is_empty tfrees then cT else if idx <= maxidx then raise CTERM ("generalize_ctyp: bad index", []) else Ctyp {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, maxidx = Int.max (maxidx, idx)}; (*Instantiation of schematic variables A -------------------- A[t1/v1, ..., tn/vn] *) local fun add_cert cert_of (_, c) cert = Context.join_certificate (cert, cert_of c); val add_instT_cert = add_cert (fn Ctyp {cert, ...} => cert); val add_inst_cert = add_cert (fn Cterm {cert, ...} => cert); fun add_sorts sorts_of (_, c) sorts = Sorts.union (sorts_of c) sorts; val add_instT_sorts = add_sorts (fn Ctyp {sorts, ...} => sorts); val add_inst_sorts = add_sorts (fn Cterm {sorts, ...} => sorts); fun add_instT thy (v as (_, S), Ctyp {T = U, maxidx, ...}) = if Sign.of_sort thy (U, S) then Term_Subst.TVars.add (v, (U, maxidx)) else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy S, [U], []); fun add_inst thy (v as (_, T), Cterm {t = u, T = U, maxidx, ...}) = if T = U then Term_Subst.Vars.add (v, (u, maxidx)) else let fun pretty_typing t ty = Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy ty]; val msg = Pretty.string_of (Pretty.block [Pretty.str "instantiate: type conflict", Pretty.fbrk, pretty_typing (Var v) T, Pretty.fbrk, pretty_typing u U]) in raise TYPE (msg, [T, U], [Var v, u]) end; fun prep_insts (instT, inst) (cert, sorts) = let val cert' = cert |> fold add_instT_cert instT |> fold add_inst_cert inst; val thy' = Context.certificate_theory cert' handle ERROR msg => raise CONTEXT (msg, map #2 instT, map #2 inst, [], NONE); val sorts' = sorts |> fold add_instT_sorts instT |> fold add_inst_sorts inst; - val instT' = fold (add_instT thy') instT Term_Subst.TVars.empty; - val inst' = fold (add_inst thy') inst Term_Subst.Vars.empty; + val instT' = Term_Subst.TVars.build (fold (add_instT thy') instT); + val inst' = Term_Subst.Vars.build (fold (add_inst thy') inst); in ((instT', inst'), (cert', sorts')) end; in (*Left-to-right replacements: ctpairs = [..., (vi, ti), ...]. Instantiates distinct Vars by terms of same type. Does NOT normalize the resulting theorem!*) fun instantiate ([], []) th = th | instantiate (instT, inst) th = let val Thm (der, {cert, hyps, constraints, shyps, tpairs, prop, ...}) = th; val ((instT', inst'), (cert', shyps')) = prep_insts (instT, inst) (cert, shyps) handle CONTEXT (msg, cTs, cts, ths, context) => raise CONTEXT (msg, cTs, cts, th :: ths, context); val subst = Term_Subst.instantiate_maxidx (instT', inst'); val (prop', maxidx1) = subst prop ~1; val (tpairs', maxidx') = fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1; val thy' = Context.certificate_theory cert'; val constraints' = Term_Subst.TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) instT' constraints; in Thm (deriv_rule1 (fn d => Proofterm.instantiate (Term_Subst.TVars.map (K #1) instT', Term_Subst.Vars.map (K #1) inst') d) der, {cert = cert', tags = [], maxidx = maxidx', constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs', prop = prop'}) |> solve_constraints end handle TYPE (msg, _, _) => raise THM (msg, 0, [th]); fun instantiate_cterm ([], []) ct = ct | instantiate_cterm (instT, inst) ct = let val Cterm {cert, t, T, sorts, ...} = ct; val ((instT', inst'), (cert', sorts')) = prep_insts (instT, inst) (cert, sorts); val subst = Term_Subst.instantiate_maxidx (instT', inst'); val substT = Term_Subst.instantiateT_maxidx instT'; val (t', maxidx1) = subst t ~1; val (T', maxidx') = substT T maxidx1; in Cterm {cert = cert', t = t', T = T', sorts = sorts', maxidx = maxidx'} end handle TYPE (msg, _, _) => raise CTERM (msg, [ct]); end; (*The trivial implication A \ A, justified by assume and forall rules. A can contain Vars, not so for assume!*) fun trivial (Cterm {cert, t = A, T, maxidx, sorts}) = if T <> propT then raise THM ("trivial: the term must have type prop", 0, []) else Thm (deriv_rule0 (fn () => Proofterm.trivial_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_implies (A, A)}); (*Axiom-scheme reflecting signature contents T :: c ------------------- OFCLASS(T, c_class) *) fun of_class (cT, raw_c) = let val Ctyp {cert, T, ...} = cT; val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [cT], [], [], NONE); val c = Sign.certify_class thy raw_c; val Cterm {t = prop, maxidx, sorts, ...} = global_cterm_of thy (Logic.mk_of_class (T, c)); in if Sign.of_sort thy (T, [c]) then Thm (deriv_rule0 (fn () => Proofterm.PClass (T, c)), {cert = cert, tags = [], maxidx = maxidx, constraints = insert_constraints thy (T, [c]) [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, []) end |> solve_constraints; (*Internalize sort constraints of type variables*) val unconstrainT = strip_shyps #> (fn thm as Thm (der, args) => let val Deriv {promises, body} = der; val {cert, shyps, hyps, tpairs, prop, ...} = args; val thy = theory_of_thm thm; fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]); val _ = null hyps orelse err "bad hyps"; val _ = null tpairs orelse err "bad flex-flex constraints"; - val tfrees = rev (Term.add_tfree_names prop []); + val tfrees = build_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 tfrees = fold Term_Subst.add_tfrees hyps fixed; val prop1 = attach_tpairs tpairs prop; val (al, prop2) = Type.varify_global tfrees prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in (al, Thm (deriv_rule1 (Proofterm.varify_proof prop tfrees) der, {cert = cert, tags = [], maxidx = Int.max (0, maxidx), constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3})) end; -val varifyT_global = #2 o varifyT_global' []; +val varifyT_global = #2 o varifyT_global' Term_Subst.TFrees.empty; (*Replace all TVars by TFrees that are often new*) fun legacy_freezeT (Thm (der, {cert, constraints, shyps, hyps, tpairs, prop, ...})) = let val prop1 = attach_tpairs tpairs prop; val prop2 = Type.legacy_freeze prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in Thm (deriv_rule1 (Proofterm.legacy_freezeT prop1) der, {cert = cert, tags = [], maxidx = maxidx_of_term prop2, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3}) end; fun plain_prop_of raw_thm = let val thm = strip_shyps raw_thm; fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]); in if not (null (hyps_of thm)) then err "theorem may not contain hypotheses" else if not (null (extra_shyps thm)) then err "theorem may not contain sort hypotheses" else if not (null (tpairs_of thm)) then err "theorem may not contain flex-flex pairs" else prop_of thm end; (*** Inference rules for tactics ***) (*Destruct proof state into constraints, other goals, goal(i), rest *) fun dest_state (state as Thm (_, {prop,tpairs,...}), i) = (case Logic.strip_prems(i, [], prop) of (B::rBs, C) => (tpairs, rev rBs, B, C) | _ => raise THM("dest_state", i, [state])) handle TERM _ => raise THM("dest_state", i, [state]); (*Prepare orule for resolution by lifting it over the parameters and assumptions of goal.*) fun lift_rule goal orule = let val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal; val inc = gmax + 1; val lift_abs = Logic.lift_abs inc gprop; val lift_all = Logic.lift_all inc gprop; val Thm (der, {maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = orule; val (As, B) = Logic.strip_horn prop; in if T <> propT then raise THM ("lift_rule: the term must have type prop", 0, []) else Thm (deriv_rule1 (Proofterm.lift_proof gprop inc prop) der, {cert = join_certificate1 (goal, orule), tags = [], maxidx = maxidx + inc, constraints = constraints, shyps = Sorts.union shyps sorts, (*sic!*) hyps = hyps, tpairs = map (apply2 lift_abs) tpairs, prop = Logic.list_implies (map lift_all As, lift_all B)}) end; fun incr_indexes i (thm as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = if i < 0 then raise THM ("negative increment", 0, [thm]) else if i = 0 then thm else Thm (deriv_rule1 (Proofterm.incr_indexes i) der, {cert = cert, tags = [], maxidx = maxidx + i, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = map (apply2 (Logic.incr_indexes ([], [], i))) tpairs, prop = Logic.incr_indexes ([], [], i) prop}); (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) fun assumption opt_ctxt i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (context, cert') = make_context_certificate [state] opt_ctxt cert; val (tpairs, Bs, Bi, C) = dest_state (state, i); fun newth n (env, tpairs) = let val normt = Envir.norm_term env; fun assumption_proof prf = Proofterm.assumption_proof (map normt Bs) (normt Bi) n prf; in Thm (deriv_rule1 (assumption_proof #> not (Envir.is_empty env) ? Proofterm.norm_proof' env) der, {tags = [], maxidx = Envir.maxidx_of env, constraints = insert_constraints_env (Context.certificate_theory cert') env constraints, shyps = Envir.insert_sorts env shyps, hyps = hyps, tpairs = if Envir.is_empty env then tpairs else map (apply2 normt) tpairs, prop = if Envir.is_empty env then Logic.list_implies (Bs, C) (*avoid wasted normalizations*) else normt (Logic.list_implies (Bs, C)) (*normalize the new rule fully*), cert = cert'}) end; val (close, asms, concl) = Logic.assum_problems (~1, Bi); val concl' = close concl; fun addprfs [] _ = Seq.empty | addprfs (asm :: rest) n = Seq.make (fn () => Seq.pull (Seq.mapp (newth n) (if Term.could_unify (asm, concl) then (Unify.unifiers (context, Envir.empty maxidx, (close asm, concl') :: tpairs)) else Seq.empty) (addprfs rest (n + 1)))) in addprfs asms 1 end; (*Solve subgoal Bi of proof state B1...Bn/C by assumption. Checks if Bi's conclusion is alpha/eta-convertible to one of its assumptions*) fun eq_assumption i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val (_, asms, concl) = Logic.assum_problems (~1, Bi); in (case find_index (fn asm => Envir.aeconv (asm, concl)) asms of ~1 => raise THM ("eq_assumption", 0, [state]) | n => Thm (deriv_rule1 (Proofterm.assumption_proof Bs Bi (n + 1)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs, C)})) end; (*For rotate_tac: fast rotation of assumptions of subgoal i*) fun rotate_rule k i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val params = Term.strip_all_vars Bi; val rest = Term.strip_all_body Bi; val asms = Logic.strip_imp_prems rest val concl = Logic.strip_imp_concl rest; val n = length asms; val m = if k < 0 then n + k else k; val Bi' = if 0 = m orelse m = n then Bi else if 0 < m andalso m < n then let val (ps, qs) = chop m asms in Logic.list_all (params, Logic.list_implies (qs @ ps, concl)) end else raise THM ("rotate_rule", k, [state]); in Thm (deriv_rule1 (Proofterm.rotate_proof Bs Bi' params asms m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs @ [Bi'], C)}) end; (*Rotates a rule's premises to the left by k, leaving the first j premises unchanged. Does nothing if k=0 or if k equals n-j, where n is the number of premises. Useful with eresolve_tac and underlies defer_tac*) fun permute_prems j k rl = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = rl; val prems = Logic.strip_imp_prems prop and concl = Logic.strip_imp_concl prop; val moved_prems = List.drop (prems, j) and fixed_prems = List.take (prems, j) handle General.Subscript => raise THM ("permute_prems: j", j, [rl]); val n_j = length moved_prems; val m = if k < 0 then n_j + k else k; val (prems', prop') = if 0 = m orelse m = n_j then (prems, prop) else if 0 < m andalso m < n_j then let val (ps, qs) = chop m moved_prems; val prems' = fixed_prems @ qs @ ps; in (prems', Logic.list_implies (prems', concl)) end else raise THM ("permute_prems: k", k, [rl]); in Thm (deriv_rule1 (Proofterm.permute_prems_proof prems' j m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) end; (* strip_apply f B A strips off all assumptions/parameters from A introduced by lifting over B, and applies f to remaining part of A*) fun strip_apply f = let fun strip (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ A2 $ B2) = Logic.mk_implies (A2, strip B1 B2) | strip ((c as Const ("Pure.all", _)) $ Abs (_, _, t1)) ( Const ("Pure.all", _) $ Abs (a, T, t2)) = c $ Abs (a, T, strip t1 t2) | strip _ A = f A in strip end; fun strip_lifted (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ _ $ B2) = strip_lifted B1 B2 | strip_lifted (Const ("Pure.all", _) $ Abs (_, _, t1)) (Const ("Pure.all", _) $ Abs (_, _, t2)) = strip_lifted t1 t2 | strip_lifted _ A = A; (*Use the alist to rename all bound variables and some unknowns in a term dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); Preserves unknowns in tpairs and on lhs of dpairs. *) fun rename_bvs [] _ _ _ _ = K I | rename_bvs al dpairs tpairs B As = let val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I); val vids = [] |> fold (add_var o fst) dpairs |> fold (add_var o fst) tpairs |> fold (add_var o snd) tpairs; val vids' = fold (add_var o strip_lifted B) As []; (*unknowns appearing elsewhere be preserved!*) val al' = distinct ((op =) o apply2 fst) (filter_out (fn (x, y) => not (member (op =) vids' x) orelse member (op =) vids x orelse member (op =) vids y) al); val unchanged = filter_out (AList.defined (op =) al') vids'; fun del_clashing clash xs _ [] qs = if clash then del_clashing false xs xs qs [] else qs | del_clashing clash xs ys ((p as (x, y)) :: ps) qs = if member (op =) ys y then del_clashing true (x :: xs) (x :: ys) ps qs else del_clashing clash xs (y :: ys) ps (p :: qs); val al'' = del_clashing false unchanged unchanged al' []; fun rename (t as Var ((x, i), T)) = (case AList.lookup (op =) al'' x of SOME y => Var ((y, i), T) | NONE => t) | rename (Abs (x, T, t)) = Abs (the_default x (AList.lookup (op =) al x), T, rename t) | rename (f $ t) = rename f $ rename t | rename t = t; fun strip_ren f Ai = f rename B Ai in strip_ren end; (*Function to rename bounds/unknowns in the argument, lifted over B*) fun rename_bvars dpairs = rename_bvs (fold_rev Term.match_bvars dpairs []) dpairs; (*** RESOLUTION ***) (** Lifting optimizations **) (*strip off pairs of assumptions/parameters in parallel -- they are identical because of lifting*) fun strip_assums2 (Const("Pure.imp", _) $ _ $ B1, Const("Pure.imp", _) $ _ $ B2) = strip_assums2 (B1,B2) | strip_assums2 (Const("Pure.all",_)$Abs(a,T,t1), Const("Pure.all",_)$Abs(_,_,t2)) = let val (B1,B2) = strip_assums2 (t1,t2) in (Abs(a,T,B1), Abs(a,T,B2)) end | strip_assums2 BB = BB; (*Faster normalization: skip assumptions that were lifted over*) fun norm_term_skip env 0 t = Envir.norm_term env t | norm_term_skip env n (Const ("Pure.all", _) $ Abs (a, T, t)) = let val T' = Envir.norm_type (Envir.type_env env) T (*Must instantiate types of parameters because they are flattened; this could be a NEW parameter*) in Logic.all_const T' $ Abs (a, T', norm_term_skip env n t) end | norm_term_skip env n (Const ("Pure.imp", _) $ A $ B) = Logic.mk_implies (A, norm_term_skip env (n - 1) B) | norm_term_skip _ _ _ = error "norm_term_skip: too few assumptions??"; (*unify types of schematic variables (non-lifted case)*) fun unify_var_types context (th1, th2) env = let fun unify_vars (T :: Us) = fold (fn U => Pattern.unify_types context (T, U)) Us | unify_vars _ = I; val add_vars = full_prop_of #> fold_aterms (fn Var v => Vartab.insert_list (op =) v | _ => I); - val vars = Vartab.empty |> add_vars th1 |> add_vars th2; + val vars = Vartab.build (add_vars th1 #> add_vars th2); in SOME (Vartab.fold (unify_vars o #2) vars env) end handle Pattern.Unif => NONE; (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) Unifies B with Bi, replacing subgoal i (1 <= i <= n) If match then forbid instantiations in proof state If lifted then shorten the dpair using strip_assums2. If eres_flg then simultaneously proves A1 by assumption. nsubgoal is the number of new subgoals (written m above). Curried so that resolution calls dest_state only once. *) local exception COMPOSE in fun bicompose_aux opt_ctxt {flatten, match, incremented} (state, (stpairs, Bs, Bi, C), lifted) (eres_flg, orule, nsubgoal) = let val Thm (sder, {maxidx=smax, constraints = constraints2, shyps = shyps2, hyps = hyps2, ...}) = state and Thm (rder, {maxidx=rmax, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs=rtpairs, prop=rprop,...}) = orule (*How many hyps to skip over during normalization*) and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0) val (context, cert) = make_context_certificate [state, orule] opt_ctxt (join_certificate2 (state, orule)); (*Add new theorem with prop = "\Bs; As\ \ C" to thq*) fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) = let val normt = Envir.norm_term env; (*perform minimal copying here by examining env*) val (ntpairs, normp) = if Envir.is_empty env then (tpairs, (Bs @ As, C)) else let val ntps = map (apply2 normt) tpairs in if Envir.above env smax then (*no assignments in state; normalize the rule only*) if lifted then (ntps, (Bs @ map (norm_term_skip env nlift) As, C)) else (ntps, (Bs @ map normt As, C)) else if match then raise COMPOSE else (*normalize the new rule fully*) (ntps, (map normt (Bs @ As), normt C)) end val constraints' = union_constraints constraints1 constraints2 |> insert_constraints_env (Context.certificate_theory cert) env; fun bicompose_proof prf1 prf2 = Proofterm.bicompose_proof flatten (map normt Bs) (map normt As) A oldAs n (nlift+1) prf1 prf2 val th = Thm (deriv_rule2 (if Envir.is_empty env then bicompose_proof else if Envir.above env smax then bicompose_proof o Proofterm.norm_proof' env else Proofterm.norm_proof' env oo bicompose_proof) rder' sder, {tags = [], maxidx = Envir.maxidx_of env, constraints = constraints', shyps = Envir.insert_sorts env (Sorts.union shyps1 shyps2), hyps = union_hyps hyps1 hyps2, tpairs = ntpairs, prop = Logic.list_implies normp, cert = cert}) in Seq.cons th thq end handle COMPOSE => thq; val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop) handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]); (*Modify assumptions, deleting n-th if n>0 for e-resolution*) fun newAs(As0, n, dpairs, tpairs) = let val (As1, rder') = if not lifted then (As0, rder) else let val rename = rename_bvars dpairs tpairs B As0 in (map (rename strip_apply) As0, deriv_rule1 (Proofterm.map_proof_terms (rename K) I) rder) end; in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n) handle TERM _ => raise THM("bicompose: 1st premise", 0, [orule]) end; val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); val dpairs = BBi :: (rtpairs@stpairs); (*elim-resolution: try each assumption in turn*) fun eres _ [] = raise THM ("bicompose: no premises", 0, [orule, state]) | eres env (A1 :: As) = let val A = SOME A1; val (close, asms, concl) = Logic.assum_problems (nlift + 1, A1); val concl' = close concl; fun tryasms [] _ = Seq.empty | tryasms (asm :: rest) n = if Term.could_unify (asm, concl) then let val asm' = close asm in (case Seq.pull (Unify.unifiers (context, env, (asm', concl') :: dpairs)) of NONE => tryasms rest (n + 1) | cell as SOME ((_, tpairs), _) => Seq.it_right (addth A (newAs (As, n, [BBi, (concl', asm')], tpairs))) (Seq.make (fn () => cell), Seq.make (fn () => Seq.pull (tryasms rest (n + 1))))) end else tryasms rest (n + 1); in tryasms asms 1 end; (*ordinary resolution*) fun res env = (case Seq.pull (Unify.unifiers (context, env, dpairs)) of NONE => Seq.empty | cell as SOME ((_, tpairs), _) => Seq.it_right (addth NONE (newAs (rev rAs, 0, [BBi], tpairs))) (Seq.make (fn () => cell), Seq.empty)); val env0 = Envir.empty (Int.max (rmax, smax)); in (case if incremented then SOME env0 else unify_var_types context (state, orule) env0 of NONE => Seq.empty | SOME env => if eres_flg then eres env (rev rAs) else res env) end; end; fun bicompose opt_ctxt flags arg i state = bicompose_aux opt_ctxt flags (state, dest_state (state,i), false) arg; (*Quick test whether rule is resolvable with the subgoal with hyps Hs and conclusion B. If eres_flg then checks 1st premise of rule also*) fun could_bires (Hs, B, eres_flg, rule) = let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs | could_reshyp [] = false; (*no premise -- illegal*) in Term.could_unify(concl_of rule, B) andalso (not eres_flg orelse could_reshyp (prems_of rule)) end; (*Bi-resolution of a state with a list of (flag,rule) pairs. Puts the rule above: rule/state. Renames vars in the rules. *) fun biresolution opt_ctxt match brules i state = let val (stpairs, Bs, Bi, C) = dest_state(state,i); val lift = lift_rule (cprem_of state i); val B = Logic.strip_assums_concl Bi; val Hs = Logic.strip_assums_hyp Bi; val compose = bicompose_aux opt_ctxt {flatten = true, match = match, incremented = true} (state, (stpairs, Bs, Bi, C), true); fun res [] = Seq.empty | res ((eres_flg, rule)::brules) = if Config.get_generic (make_context [state] opt_ctxt (cert_of state)) Pattern.unify_trace_failure orelse could_bires (Hs, B, eres_flg, rule) then Seq.make (*delay processing remainder till needed*) (fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule), res brules)) else res brules in Seq.flat (res brules) end; (*Resolution: exactly one resolvent must be produced*) fun tha RSN (i, thb) = (case Seq.chop 2 (biresolution NONE false [(false, tha)] i thb) of ([th], _) => solve_constraints th | ([], _) => raise THM ("RSN: no unifiers", i, [tha, thb]) | _ => raise THM ("RSN: multiple unifiers", i, [tha, thb])); (*Resolution: P \ Q, Q \ R gives P \ R*) fun tha RS thb = tha RSN (1,thb); (**** Type classes ****) fun standard_tvars thm = let val thy = theory_of_thm thm; - val tvars = rev (Term.add_tvars (prop_of thm) []); + val tvars = build_rev (Term.add_tvars (prop_of thm)); val names = Name.invent Name.context Name.aT (length tvars); val tinst = map2 (fn (ai, S) => fn b => ((ai, S), global_ctyp_of thy (TVar ((b, 0), S)))) tvars names; in instantiate (tinst, []) thm end (* class relations *) val is_classrel = Symreltab.defined o get_classrels; fun complete_classrels thy = let fun complete (c, (_, (all_preds, all_succs))) (finished1, thy1) = let fun compl c1 c2 (finished2, thy2) = if is_classrel thy2 (c1, c2) then (finished2, thy2) else (false, thy2 |> (map_classrels o Symreltab.update) ((c1, c2), (the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy2) |> trim_context)); val proven = is_classrel thy1; val preds = Graph.Keys.fold (fn c1 => proven (c1, c) ? cons c1) all_preds []; val succs = Graph.Keys.fold (fn c2 => proven (c, c2) ? cons c2) all_succs []; in fold_product compl preds succs (finished1, thy1) end; in (case Graph.fold complete (Sorts.classes_of (Sign.classes_of thy)) (true, thy) of (true, _) => NONE | (_, thy') => SOME thy') end; (* type arities *) fun thynames_of_arity thy (a, c) = - (get_arities thy, []) |-> Aritytab.fold - (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser)) + build (get_arities thy |> Aritytab.fold + (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser))) |> sort (int_ord o apply2 #2) |> map #1; fun insert_arity_completions thy ((t, Ss, c), (th, thy_name, ser)) (finished, arities) = let val completions = Sign.super_classes thy c |> map_filter (fn c1 => if Aritytab.defined arities (t, Ss, c1) then NONE else let val th1 = (th RS the_classrel thy (c, c1)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy) |> trim_context; in SOME ((t, Ss, c1), (th1, thy_name, ser)) end); val finished' = finished andalso null completions; val arities' = fold Aritytab.update completions arities; in (finished', arities') end; fun complete_arities thy = let val arities = get_arities thy; val (finished, arities') = Aritytab.fold (insert_arity_completions thy) arities (true, get_arities thy); in if finished then NONE else SOME (map_arities (K arities') thy) end; val _ = Theory.setup (Theory.at_begin complete_classrels #> Theory.at_begin complete_arities); (* primitive rules *) fun add_classrel raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (c1, c2) = Logic.dest_classrel prop; in thy |> Sign.primitive_classrel (c1, c2) |> map_classrels (Symreltab.update ((c1, c2), th')) |> perhaps complete_classrels |> perhaps complete_arities end; fun add_arity raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (t, Ss, c) = Logic.dest_arity prop; val ar = ((t, Ss, c), (th', Context.theory_name thy, serial ())); in thy |> Sign.primitive_arity (t, Ss, [c]) |> map_arities (Aritytab.update ar #> curry (insert_arity_completions thy ar) true #> #2) end; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Pure/thm_deps.ML b/src/Pure/thm_deps.ML --- a/src/Pure/thm_deps.ML +++ b/src/Pure/thm_deps.ML @@ -1,143 +1,144 @@ (* Title: Pure/thm_deps.ML Author: Stefan Berghofer, TU Muenchen Author: Makarius Dependencies of theorems wrt. internal derivation. *) signature THM_DEPS = sig val all_oracles: thm list -> Proofterm.oracle list val has_skip_proof: thm list -> bool val pretty_thm_oracles: Proof.context -> thm list -> Pretty.T val thm_deps: theory -> thm list -> (Proofterm.thm_id * Thm_Name.T) list val pretty_thm_deps: theory -> thm list -> Pretty.T val unused_thms_cmd: theory list * theory list -> (string * thm) list end; structure Thm_Deps: THM_DEPS = struct (* oracles *) fun all_oracles thms = let fun collect (PBody {oracles, thms, ...}) = (if null oracles then I else apfst (cons oracles)) #> (tap Proofterm.join_thms thms |> fold (fn (i, thm_node) => fn (res, seen) => if Inttab.defined seen i then (res, seen) else let val body = Future.join (Proofterm.thm_node_body thm_node) in collect body (res, Inttab.update (i, ()) seen) end)); val bodies = map Thm.proof_body_of thms; in fold collect bodies ([], Inttab.empty) |> #1 |> Proofterm.unions_oracles end; fun has_skip_proof thms = all_oracles thms |> exists (fn ((name, _), _) => name = \<^oracle_name>\skip_proof\); fun pretty_thm_oracles ctxt thms = let val thy = Proof_Context.theory_of ctxt; fun prt_ora (name, pos) = Thm.pretty_oracle ctxt name :: Pretty.here pos; fun prt_oracle (ora, NONE) = prt_ora ora | prt_oracle (ora, SOME prop) = prt_ora ora @ [Pretty.str ":", Pretty.brk 1, Syntax.pretty_term_global thy prop]; val oracles = (case try all_oracles thms of SOME oracles => oracles | NONE => error "Malformed proofs") in Pretty.big_list "oracles:" (map (Pretty.item o prt_oracle) oracles) end; (* thm_deps *) fun thm_deps thy = let val lookup = Global_Theory.lookup_thm_id thy; fun deps (i, thm_node) res = if Inttab.defined res i then res else let val thm_id = Proofterm.thm_id (i, thm_node) in (case lookup thm_id of SOME thm_name => Inttab.update (i, SOME (thm_id, thm_name)) res | NONE => Inttab.update (i, NONE) res |> fold deps (Proofterm.thm_node_thms thm_node)) end; in fn thms => - (fold (fold deps o Thm.thm_deps o Thm.transfer thy) thms Inttab.empty, []) + (Inttab.build (fold (fold deps o Thm.thm_deps o Thm.transfer thy) thms), []) |-> Inttab.fold_rev (fn (_, SOME entry) => cons entry | _ => I) end; fun pretty_thm_deps thy thms = let val ctxt = Proof_Context.init_global thy; val deps = (case try (thm_deps thy) thms of SOME deps => deps | NONE => error "Malformed proofs"); val items = map #2 deps |> map (fn (name, i) => (Proof_Context.markup_extern_fact ctxt name, i)) |> sort_by (#2 o #1) |> map (fn ((marks, xname), i) => Pretty.item [Pretty.marks_str (marks, Thm_Name.print (xname, i))]); in Pretty.big_list ("dependencies: " ^ string_of_int (length items)) items end; (* unused_thms_cmd *) fun unused_thms_cmd (base_thys, thys) = let fun add_fact transfer space (name, ths) = if exists (fn thy => Global_Theory.defined_fact thy name) base_thys then I else let val {concealed, group, ...} = Name_Space.the_entry space name in fold_rev (transfer #> (fn th => (case Thm.derivation_name th of "" => I | a => cons (a, (th, concealed, group))))) ths end; fun add_facts thy = let val transfer = Global_Theory.transfer_theories thy; val facts = Global_Theory.facts_of thy; in Facts.fold_static (add_fact transfer (Facts.space_of facts)) facts end; val new_thms = fold add_facts thys [] |> sort_distinct (string_ord o apply2 #1); val used = Proofterm.fold_body_thms (fn {name = a, ...} => a <> "" ? Symtab.update (a, ())) (map Proofterm.strip_thm_body (Thm.proof_bodies_of (map (#1 o #2) new_thms))) Symtab.empty; fun is_unused a = not (Symtab.defined used a); (*groups containing at least one used theorem*) - val used_groups = fold (fn (a, (_, _, group)) => - if is_unused a then I - else - (case group of - NONE => I - | SOME grp => Inttab.update (grp, ()))) new_thms Inttab.empty; + val used_groups = + Inttab.build (new_thms |> fold (fn (a, (_, _, group)) => + if is_unused a then I + else + (case group of + NONE => I + | SOME grp => Inttab.update (grp, ())))); val (thms', _) = fold (fn (a, (th, concealed, group)) => fn q as (thms, seen_groups) => if not concealed andalso (* FIXME replace by robust treatment of thm groups *) Thm.legacy_get_kind th = Thm.theoremK andalso is_unused a then (case group of NONE => ((a, th) :: thms, seen_groups) | SOME grp => if Inttab.defined used_groups grp orelse Inttab.defined seen_groups grp then q else ((a, th) :: thms, Inttab.update (grp, ()) seen_groups)) else q) new_thms ([], Inttab.empty); in rev thms' end; end; diff --git a/src/Pure/type.ML b/src/Pure/type.ML --- a/src/Pure/type.ML +++ b/src/Pure/type.ML @@ -1,713 +1,714 @@ (* Title: Pure/type.ML Author: Tobias Nipkow, Lawrence C Paulson, and Markus Wenzel Type signatures and certified types, special treatment of type vars, matching and unification of types, extend and merge type signatures. *) signature TYPE = sig (*constraints*) val mark_polymorphic: typ -> typ val constraint: typ -> term -> term val constraint_type: Proof.context -> typ -> typ val strip_constraints: term -> term val appl_error: Proof.context -> term -> typ -> term -> typ -> string (*type signatures and certified types*) datatype decl = LogicalType of int | Abbreviation of string list * typ * bool | Nonterminal type tsig val eq_tsig: tsig * tsig -> bool val rep_tsig: tsig -> {classes: Name_Space.T * Sorts.algebra, default: sort, types: decl Name_Space.table, log_types: string list} val change_base: bool -> tsig -> tsig val change_ignore: tsig -> tsig val empty_tsig: tsig val class_space: tsig -> Name_Space.T val defaultS: tsig -> sort val logical_types: tsig -> string list val eq_sort: tsig -> sort * sort -> bool val subsort: tsig -> sort * sort -> bool val of_sort: tsig -> typ * sort -> bool val inter_sort: tsig -> sort * sort -> sort val cert_class: tsig -> class -> class val cert_sort: tsig -> sort -> sort val minimize_sort: tsig -> sort -> sort val witness_sorts: tsig -> (typ * sort) list -> sort list -> (typ * sort) list type mode val mode_default: mode val mode_syntax: mode val mode_abbrev: mode val get_mode: Proof.context -> mode val set_mode: mode -> Proof.context -> Proof.context val restore_mode: Proof.context -> Proof.context -> Proof.context val type_space: tsig -> Name_Space.T val type_alias: Name_Space.naming -> binding -> string -> tsig -> tsig val check_decl: Context.generic -> tsig -> xstring * Position.T -> (string * Position.report list) * decl val the_decl: tsig -> string * Position.T -> decl val cert_typ_mode: mode -> tsig -> typ -> typ val cert_typ: tsig -> typ -> typ val arity_number: tsig -> string -> int val arity_sorts: Context.generic -> tsig -> string -> sort -> sort list (*special treatment of type vars*) val sort_of_atyp: typ -> sort val strip_sorts: typ -> typ val strip_sorts_dummy: typ -> typ val no_tvars: typ -> typ - val varify_global: (string * sort) list -> term -> ((string * sort) * indexname) list * term + val varify_global: Term_Subst.TFrees.set -> term -> ((string * sort) * indexname) list * term val legacy_freeze_thaw_type: typ -> typ * (typ -> typ) val legacy_freeze_type: typ -> typ val legacy_freeze_thaw: term -> term * (term -> term) val legacy_freeze: term -> term (*matching and unification*) exception TYPE_MATCH type tyenv = (sort * typ) Vartab.table val lookup: tyenv -> indexname * sort -> typ option val devar: tyenv -> typ -> typ val typ_match: tsig -> typ * typ -> tyenv -> tyenv val typ_instance: tsig -> typ * typ -> bool val raw_match: typ * typ -> tyenv -> tyenv val raw_matches: typ list * typ list -> tyenv -> tyenv val could_match: typ * typ -> bool val could_matches: typ list * typ list -> bool val raw_instance: typ * typ -> bool exception TUNIFY val unify: tsig -> typ * typ -> tyenv * int -> tyenv * int val raw_unify: typ * typ -> tyenv -> tyenv val raw_unifys: typ list * typ list -> tyenv -> tyenv val could_unify: typ * typ -> bool val could_unifys: typ list * typ list -> bool val unified: tyenv -> typ * typ -> bool (*extend and merge type signatures*) val add_class: Context.generic -> binding * class list -> tsig -> tsig val hide_class: bool -> string -> tsig -> tsig val set_defsort: sort -> tsig -> tsig val add_type: Context.generic -> binding * int -> tsig -> tsig val add_abbrev: Context.generic -> binding * string list * typ -> tsig -> tsig val add_nonterminal: Context.generic -> binding -> tsig -> tsig val hide_type: bool -> string -> tsig -> tsig val add_arity: Context.generic -> arity -> tsig -> tsig val add_classrel: Context.generic -> class * class -> tsig -> tsig val merge_tsig: Context.generic -> tsig * tsig -> tsig end; structure Type: TYPE = struct (** constraints **) (*indicate polymorphic Vars*) fun mark_polymorphic T = Type ("_polymorphic_", [T]); fun constraint T t = if T = dummyT then t else Const ("_type_constraint_", T --> T) $ t; fun constraint_type ctxt T = let fun err () = error ("Malformed internal type constraint: " ^ Syntax.string_of_typ ctxt T); in (case T of Type ("fun", [A, B]) => if A = B then A else err () | _ => err ()) end; fun strip_constraints (Const ("_type_constraint_", _) $ t) = strip_constraints t | strip_constraints (t $ u) = strip_constraints t $ strip_constraints u | strip_constraints (Abs (x, T, t)) = Abs (x, T, strip_constraints t) | strip_constraints a = a; fun appl_error ctxt (Const ("_type_constraint_", Type ("fun", [T, _]))) _ u U = cat_lines ["Failed to meet type constraint:", "", Pretty.string_of (Pretty.block [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt u, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt U]), Pretty.string_of (Pretty.block [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt T])] | appl_error ctxt t T u U = cat_lines ["Type error in application: " ^ (case T of Type ("fun", _) => "incompatible operand type" | _ => "operator not of function type"), "", Pretty.string_of (Pretty.block [Pretty.str "Operator:", Pretty.brk 2, Syntax.pretty_term ctxt t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt T]), Pretty.string_of (Pretty.block [Pretty.str "Operand:", Pretty.brk 3, Syntax.pretty_term ctxt u, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt U])]; (** type signatures and certified types **) (* type declarations *) datatype decl = LogicalType of int | Abbreviation of string list * typ * bool | Nonterminal; (* type tsig *) datatype tsig = TSig of { classes: Name_Space.T * Sorts.algebra, (*order-sorted algebra of type classes*) default: sort, (*default sort on input*) types: decl Name_Space.table, (*declared types*) log_types: string list}; (*logical types sorted by number of arguments*) fun eq_tsig (TSig {classes = classes1, default = default1, types = types1, log_types = _}, TSig {classes = classes2, default = default2, types = types2, log_types = _}) = pointer_eq (classes1, classes2) andalso default1 = default2 andalso pointer_eq (types1, types2); fun rep_tsig (TSig comps) = comps; fun make_tsig (classes, default, types, log_types) = TSig {classes = classes, default = default, types = types, log_types = log_types}; fun change_base begin (TSig {classes, default, types, log_types}) = make_tsig (classes, default, Name_Space.change_base begin types, log_types); fun change_ignore (TSig {classes, default, types, log_types}) = make_tsig (classes, default, Name_Space.change_ignore types, log_types); fun build_tsig (classes, default, types) = let val log_types = Name_Space.fold_table (fn (c, LogicalType n) => cons (c, n) | _ => I) types [] |> Library.sort (int_ord o apply2 snd) |> map fst; in make_tsig (classes, default, types, log_types) end; fun map_tsig f (TSig {classes, default, types, log_types = _}) = build_tsig (f (classes, default, types)); val empty_tsig = build_tsig ((Name_Space.empty Markup.classN, Sorts.empty_algebra), [], Name_Space.empty_table Markup.type_nameN); (* classes and sorts *) val class_space = #1 o #classes o rep_tsig; fun defaultS (TSig {default, ...}) = default; fun logical_types (TSig {log_types, ...}) = log_types; fun eq_sort (TSig {classes, ...}) = Sorts.sort_eq (#2 classes); fun subsort (TSig {classes, ...}) = Sorts.sort_le (#2 classes); fun of_sort (TSig {classes, ...}) = Sorts.of_sort (#2 classes); fun inter_sort (TSig {classes, ...}) = Sorts.inter_sort (#2 classes); fun cert_class (TSig {classes = (_, algebra), ...}) c = if can (Graph.get_entry (Sorts.classes_of algebra)) c then c else raise TYPE ("Undeclared class: " ^ quote c, [], []); val cert_sort = map o cert_class; fun minimize_sort (TSig {classes, ...}) = Sorts.minimize_sort (#2 classes); fun witness_sorts (TSig {classes, log_types, ...}) = Sorts.witness_sorts (#2 classes) log_types; (* certification mode *) datatype mode = Mode of {normalize: bool, logical: bool}; val mode_default = Mode {normalize = true, logical = true}; val mode_syntax = Mode {normalize = true, logical = false}; val mode_abbrev = Mode {normalize = false, logical = false}; structure Mode = Proof_Data ( type T = mode; fun init _ = mode_default; ); val get_mode = Mode.get; fun set_mode mode = Mode.map (K mode); fun restore_mode ctxt = set_mode (get_mode ctxt); (* types *) val type_space = Name_Space.space_of_table o #types o rep_tsig; fun type_alias naming binding name = map_tsig (fn (classes, default, types) => (classes, default, (Name_Space.alias_table naming binding name types))); fun undecl_type c = "Undeclared type constructor: " ^ quote c; fun lookup_type (TSig {types, ...}) = Name_Space.lookup types; fun check_decl context (TSig {types, ...}) (c, pos) = Name_Space.check_reports context types (c, [pos]); fun the_decl tsig (c, pos) = (case lookup_type tsig c of NONE => error (undecl_type c ^ Position.here pos) | SOME decl => decl); (* certified types *) fun bad_nargs t = "Bad number of arguments for type constructor: " ^ quote t; local fun inst_typ env (Type (c, Ts)) = Type (c, map (inst_typ env) Ts) | inst_typ env (T as TFree (x, _)) = the_default T (AList.lookup (op =) env x) | inst_typ _ T = T; in fun cert_typ_mode (Mode {normalize, logical}) tsig ty = let fun err msg = raise TYPE (msg, [ty], []); val check_logical = if logical then fn c => err ("Illegal occurrence of syntactic type: " ^ quote c) else fn _ => (); fun cert (T as Type (c, Ts)) = let val Ts' = map cert Ts; fun nargs n = if length Ts <> n then err (bad_nargs c) else (); in (case the_decl tsig (c, Position.none) of LogicalType n => (nargs n; Type (c, Ts')) | Abbreviation (vs, U, syn) => (nargs (length vs); if syn then check_logical c else (); if normalize then inst_typ (vs ~~ Ts') U else Type (c, Ts')) | Nonterminal => (nargs 0; check_logical c; T)) end | cert (TFree (x, S)) = TFree (x, cert_sort tsig S) | cert (TVar (xi as (_, i), S)) = if i < 0 then err ("Malformed type variable: " ^ quote (Term.string_of_vname xi)) else TVar (xi, cert_sort tsig S); val ty' = cert ty; in if ty = ty' then ty else ty' end; (*avoid copying of already normal type*) val cert_typ = cert_typ_mode mode_default; end; (* type arities *) fun arity_number tsig a = (case lookup_type tsig a of SOME (LogicalType n) => n | _ => error (undecl_type a)); fun arity_sorts _ tsig a [] = replicate (arity_number tsig a) [] | arity_sorts context (TSig {classes, ...}) a S = Sorts.mg_domain (#2 classes) a S handle Sorts.CLASS_ERROR err => error (Sorts.class_error context err); (** special treatment of type vars **) (* sort_of_atyp *) fun sort_of_atyp (TFree (_, S)) = S | sort_of_atyp (TVar (_, S)) = S | sort_of_atyp T = raise TYPE ("sort_of_atyp", [T], []); (* strip_sorts *) val strip_sorts = map_atyps (fn TFree (x, _) => TFree (x, []) | TVar (xi, _) => TVar (xi, [])); val strip_sorts_dummy = map_atyps (fn TFree (x, _) => TFree (x, dummyS) | TVar (xi, _) => TVar (xi, dummyS)); (* no_tvars *) fun no_tvars T = (case Term.add_tvarsT T [] of [] => T | vs => raise TYPE ("Illegal schematic type variable(s): " ^ commas_quote (map (Term.string_of_vname o #1) (rev vs)), [T], [])); (* varify_global *) fun varify_global fixed t = let - val fs = Term.fold_types (Term.fold_atyps - (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t []; + val fs = + build (t |> (Term.fold_types o Term.fold_atyps) + (fn TFree v => if Term_Subst.TFrees.defined fixed v then I else insert (op =) v | _ => I)); val used = Name.context |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t; val fmap = fs ~~ map (rpair 0) (#1 (fold_map Name.variant (map fst fs) used)); fun thaw (f as (_, S)) = (case AList.lookup (op =) fmap f of NONE => TFree f | SOME xi => TVar (xi, S)); in (fmap, map_types (map_type_tfree thaw) t) end; (* freeze_thaw: freeze TVars in a term; return the "thaw" inverse *) local fun new_name ix (pairs, used) = let val v = singleton (Name.variant_list used) (string_of_indexname ix) in ((ix, v) :: pairs, v :: used) end; fun freeze_one alist (ix, sort) = TFree (the (AList.lookup (op =) alist ix), sort) handle Option.Option => raise TYPE ("Failure during freezing of ?" ^ string_of_indexname ix, [], []); fun thaw_one alist (a, sort) = TVar (the (AList.lookup (op =) alist a), sort) handle Option.Option => TFree (a, sort); in fun legacy_freeze_thaw_type T = let val used = Term.add_tfree_namesT T []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvarsT T [])) ([], used); in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end; val legacy_freeze_type = #1 o legacy_freeze_thaw_type; fun legacy_freeze_thaw t = let val used = Term.add_tfree_names t []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvars t [])) ([], used); in (case alist of [] => (t, fn x => x) (*nothing to do!*) | _ => (map_types (map_type_tvar (freeze_one alist)) t, map_types (map_type_tfree (thaw_one (map swap alist))))) end; val legacy_freeze = #1 o legacy_freeze_thaw; end; (** matching and unification of types **) type tyenv = (sort * typ) Vartab.table; fun tvar_clash ixn S S' = raise TYPE ("Type variable has two distinct sorts", [TVar (ixn, S), TVar (ixn, S')], []); fun lookup tye (ixn, S) = (case Vartab.lookup tye ixn of NONE => NONE | SOME (S', T) => if S = S' then SOME T else tvar_clash ixn S S'); (* matching *) exception TYPE_MATCH; fun typ_match tsig = let fun match (TVar (v, S), T) subs = (case lookup subs (v, S) of NONE => if of_sort tsig (T, S) then Vartab.update_new (v, (S, T)) subs else raise TYPE_MATCH | SOME U => if U = T then subs else raise TYPE_MATCH) | match (Type (a, Ts), Type (b, Us)) subs = if a <> b then raise TYPE_MATCH else matches (Ts, Us) subs | match (TFree x, TFree y) subs = if x = y then subs else raise TYPE_MATCH | match _ _ = raise TYPE_MATCH and matches (T :: Ts, U :: Us) subs = matches (Ts, Us) (match (T, U) subs) | matches _ subs = subs; in match end; fun typ_instance tsig (T, U) = - (typ_match tsig (U, T) Vartab.empty; true) handle TYPE_MATCH => false; + (Vartab.build (typ_match tsig (U, T)); true) handle TYPE_MATCH => false; (*purely structural matching*) fun raw_match (TVar (v, S), T) subs = (case lookup subs (v, S) of NONE => Vartab.update_new (v, (S, T)) subs | SOME U => if U = T then subs else raise TYPE_MATCH) | raw_match (Type (a, Ts), Type (b, Us)) subs = if a <> b then raise TYPE_MATCH else raw_matches (Ts, Us) subs | raw_match (TFree x, TFree y) subs = if x = y then subs else raise TYPE_MATCH | raw_match _ _ = raise TYPE_MATCH and raw_matches (T :: Ts, U :: Us) subs = raw_matches (Ts, Us) (raw_match (T, U) subs) | raw_matches ([], []) subs = subs | raw_matches _ _ = raise TYPE_MATCH; (*fast matching filter*) fun could_match (Type (a, Ts), Type (b, Us)) = a = b andalso could_matches (Ts, Us) | could_match (TFree (a, _), TFree (b, _)) = a = b | could_match (TVar _, _) = true | could_match _ = false and could_matches (T :: Ts, U :: Us) = could_match (T, U) andalso could_matches (Ts, Us) | could_matches ([], []) = true | could_matches _ = false; fun raw_instance (T, U) = if could_match (U, T) then - (raw_match (U, T) Vartab.empty; true) handle TYPE_MATCH => false + (Vartab.build (raw_match (U, T)); true) handle TYPE_MATCH => false else false; (* unification *) exception TUNIFY; (*occurs check*) fun occurs v tye = let fun occ (Type (_, Ts)) = exists occ Ts | occ (TFree _) = false | occ (TVar (w, S)) = Term.eq_ix (v, w) orelse (case lookup tye (w, S) of NONE => false | SOME U => occ U); in occ end; (*chase variable assignments; if devar returns a type var then it must be unassigned*) fun devar tye (T as TVar v) = (case lookup tye v of SOME U => devar tye U | NONE => T) | devar _ T = T; (*order-sorted unification*) fun unify (TSig {classes = (_, classes), ...}) TU (tyenv, maxidx) = let val tyvar_count = Unsynchronized.ref maxidx; fun gen_tyvar S = TVar ((Name.aT, Unsynchronized.inc tyvar_count), S); fun mg_domain a S = Sorts.mg_domain classes a S handle Sorts.CLASS_ERROR _ => raise TUNIFY; fun meet (_, []) tye = tye | meet (TVar (xi, S'), S) tye = if Sorts.sort_le classes (S', S) then tye else Vartab.update_new (xi, (S', gen_tyvar (Sorts.inter_sort classes (S', S)))) tye | meet (TFree (_, S'), S) tye = if Sorts.sort_le classes (S', S) then tye else raise TUNIFY | meet (Type (a, Ts), S) tye = meets (Ts, mg_domain a S) tye and meets (T :: Ts, S :: Ss) tye = meets (Ts, Ss) (meet (devar tye T, S) tye) | meets _ tye = tye; fun unif (ty1, ty2) tye = (case (devar tye ty1, devar tye ty2) of (T as TVar (v, S1), U as TVar (w, S2)) => if Term.eq_ix (v, w) then if S1 = S2 then tye else tvar_clash v S1 S2 else if Sorts.sort_le classes (S1, S2) then Vartab.update_new (w, (S2, T)) tye else if Sorts.sort_le classes (S2, S1) then Vartab.update_new (v, (S1, U)) tye else let val S = gen_tyvar (Sorts.inter_sort classes (S1, S2)) in Vartab.update_new (v, (S1, S)) (Vartab.update_new (w, (S2, S)) tye) end | (TVar (v, S), T) => if occurs v tye T then raise TUNIFY else meet (T, S) (Vartab.update_new (v, (S, T)) tye) | (T, TVar (v, S)) => if occurs v tye T then raise TUNIFY else meet (T, S) (Vartab.update_new (v, (S, T)) tye) | (Type (a, Ts), Type (b, Us)) => if a <> b then raise TUNIFY else unifs (Ts, Us) tye | (T, U) => if T = U then tye else raise TUNIFY) and unifs (T :: Ts, U :: Us) tye = unifs (Ts, Us) (unif (T, U) tye) | unifs _ tye = tye; in (unif TU tyenv, ! tyvar_count) end; (*purely structural unification*) fun raw_unify (ty1, ty2) tye = (case (devar tye ty1, devar tye ty2) of (T as TVar (v, S1), TVar (w, S2)) => if Term.eq_ix (v, w) then if S1 = S2 then tye else tvar_clash v S1 S2 else Vartab.update_new (w, (S2, T)) tye | (TVar (v, S), T) => if occurs v tye T then raise TUNIFY else Vartab.update_new (v, (S, T)) tye | (T, TVar (v, S)) => if occurs v tye T then raise TUNIFY else Vartab.update_new (v, (S, T)) tye | (Type (a, Ts), Type (b, Us)) => if a <> b then raise TUNIFY else raw_unifys (Ts, Us) tye | (T, U) => if T = U then tye else raise TUNIFY) and raw_unifys (T :: Ts, U :: Us) tye = raw_unifys (Ts, Us) (raw_unify (T, U) tye) | raw_unifys ([], []) tye = tye | raw_unifys _ _ = raise TUNIFY; (*fast unification filter*) fun could_unify (Type (a, Ts), Type (b, Us)) = a = b andalso could_unifys (Ts, Us) | could_unify (TFree (a, _), TFree (b, _)) = a = b | could_unify (TVar _, _) = true | could_unify (_, TVar _) = true | could_unify _ = false and could_unifys (T :: Ts, U :: Us) = could_unify (T, U) andalso could_unifys (Ts, Us) | could_unifys ([], []) = true | could_unifys _ = false; (*equality with respect to a type environment*) fun unified tye = let fun unif (T, T') = (case (devar tye T, devar tye T') of (Type (s, Ts), Type (s', Ts')) => s = s' andalso unifs (Ts, Ts') | (U, U') => U = U') and unifs ([], []) = true | unifs (T :: Ts, T' :: Ts') = unif (T', T') andalso unifs (Ts, Ts') | unifs _ = false; in if Vartab.is_empty tye then op = else unif end; (** extend and merge type signatures **) (* classes *) fun add_class context (c, cs) tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val cs' = map (cert_class tsig) cs handle TYPE (msg, _, _) => error msg; val _ = Binding.check c; val (c', space') = space |> Name_Space.declare context true c; val classes' = classes |> Sorts.add_class context (c', cs'); in ((space', classes'), default, types) end); fun hide_class fully c = map_tsig (fn ((space, classes), default, types) => ((Name_Space.hide fully c space, classes), default, types)); (* arities *) fun add_arity context (t, Ss, S) tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val _ = (case lookup_type tsig t of SOME (LogicalType n) => if length Ss <> n then error (bad_nargs t) else () | SOME _ => error ("Logical type constructor expected: " ^ quote t) | NONE => error (undecl_type t)); val (Ss', S') = (map (cert_sort tsig) Ss, cert_sort tsig S) handle TYPE (msg, _, _) => error msg; val classes' = classes |> Sorts.add_arities context ((t, map (fn c' => (c', Ss')) S')); in ((space, classes'), default, types) end); (* classrel *) fun add_classrel context rel tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val rel' = apply2 (cert_class tsig) rel handle TYPE (msg, _, _) => error msg; val classes' = classes |> Sorts.add_classrel context rel'; in ((space, classes'), default, types) end); (* default sort *) fun set_defsort S tsig = tsig |> map_tsig (fn (classes, _, types) => (classes, cert_sort tsig S handle TYPE (msg, _, _) => error msg, types)); (* types *) local fun new_decl context (c, decl) types = (Binding.check c; #2 (Name_Space.define context true (c, decl) types)); fun map_types f = map_tsig (fn (classes, default, types) => let val types' = f types; val _ = not (Name_Space.defined types' "dummy") orelse Name_Space.intern (Name_Space.space_of_table types') "dummy" = "dummy" orelse error "Illegal declaration of dummy type"; in (classes, default, types') end); fun syntactic tsig (Type (c, Ts)) = (case lookup_type tsig c of SOME Nonterminal => true | _ => false) orelse exists (syntactic tsig) Ts | syntactic _ _ = false; in fun add_type context (c, n) = if n < 0 then error ("Bad type constructor declaration " ^ Binding.print c) else map_types (new_decl context (c, LogicalType n)); fun add_abbrev context (a, vs, rhs) tsig = tsig |> map_types (fn types => let fun err msg = cat_error msg ("The error(s) above occurred in type abbreviation " ^ Binding.print a); val rhs' = strip_sorts (no_tvars (cert_typ_mode mode_syntax tsig rhs)) handle TYPE (msg, _, _) => err msg; val _ = (case duplicates (op =) vs of [] => [] | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups)); val _ = (case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of [] => [] | extras => err ("Extra variables on rhs: " ^ commas_quote extras)); in types |> new_decl context (a, Abbreviation (vs, rhs', syntactic tsig rhs')) end); fun add_nonterminal context = map_types o new_decl context o rpair Nonterminal; end; fun hide_type fully c = map_tsig (fn (classes, default, types) => (classes, default, Name_Space.hide_table fully c types)); (* merge type signatures *) fun merge_tsig context (tsig1, tsig2) = let val (TSig {classes = (space1, classes1), default = default1, types = types1, log_types = _}) = tsig1; val (TSig {classes = (space2, classes2), default = default2, types = types2, log_types = _}) = tsig2; val space' = Name_Space.merge (space1, space2); val classes' = Sorts.merge_algebra context (classes1, classes2); val default' = Sorts.inter_sort classes' (default1, default2); val types' = Name_Space.merge_tables (types1, types2); in build_tsig ((space', classes'), default', types') end; end; diff --git a/src/Pure/variable.ML b/src/Pure/variable.ML --- a/src/Pure/variable.ML +++ b/src/Pure/variable.ML @@ -1,748 +1,747 @@ (* Title: Pure/variable.ML Author: Makarius Fixed type/term variables and polymorphic term abbreviations. *) signature VARIABLE = sig val names_of: Proof.context -> Name.context val binds_of: Proof.context -> (typ * term) Vartab.table val maxidx_of: Proof.context -> int val constraints_of: Proof.context -> typ Vartab.table * sort Vartab.table val is_declared: Proof.context -> string -> bool val check_name: binding -> string val default_type: Proof.context -> string -> typ option val def_type: Proof.context -> bool -> indexname -> typ option val def_sort: Proof.context -> indexname -> sort option val declare_maxidx: int -> Proof.context -> Proof.context val declare_names: term -> Proof.context -> Proof.context val declare_constraints: term -> Proof.context -> Proof.context val declare_internal: term -> Proof.context -> Proof.context val declare_term: term -> Proof.context -> Proof.context val declare_typ: typ -> Proof.context -> Proof.context val declare_prf: Proofterm.proof -> Proof.context -> Proof.context val declare_thm: thm -> Proof.context -> Proof.context val variant_frees: Proof.context -> term list -> (string * 'a) list -> (string * 'a) list val bind_term: indexname * term -> Proof.context -> Proof.context val unbind_term: indexname -> Proof.context -> Proof.context val maybe_bind_term: indexname * term option -> Proof.context -> Proof.context val expand_binds: Proof.context -> term -> term val lookup_const: Proof.context -> string -> string option val is_const: Proof.context -> string -> bool val declare_const: string * string -> Proof.context -> Proof.context val next_bound: string * typ -> Proof.context -> term * Proof.context val revert_bounds: Proof.context -> term -> term val is_body: Proof.context -> bool val set_body: bool -> Proof.context -> Proof.context val restore_body: Proof.context -> Proof.context -> Proof.context val improper_fixes: Proof.context -> Proof.context val restore_proper_fixes: Proof.context -> Proof.context -> Proof.context val is_improper: Proof.context -> string -> bool val is_fixed: Proof.context -> string -> bool val is_newly_fixed: Proof.context -> Proof.context -> string -> bool val fixed_ord: Proof.context -> string ord val intern_fixed: Proof.context -> string -> string val lookup_fixed: Proof.context -> string -> string option val revert_fixed: Proof.context -> string -> string val markup_fixed: Proof.context -> string -> Markup.T val markup: Proof.context -> string -> Markup.T val markup_entity_def: Proof.context -> string -> Markup.T val dest_fixes: Proof.context -> (string * string) list val add_fixed_names: Proof.context -> term -> string list -> string list val add_fixed: Proof.context -> term -> (string * typ) list -> (string * typ) list val add_newly_fixed: Proof.context -> Proof.context -> term -> (string * typ) list -> (string * typ) list val add_free_names: Proof.context -> term -> string list -> string list val add_frees: Proof.context -> term -> (string * typ) list -> (string * typ) list val add_fixes_binding: binding list -> Proof.context -> string list * Proof.context val add_fixes: string list -> Proof.context -> string list * Proof.context val add_fixes_direct: string list -> Proof.context -> Proof.context val add_fixes_implicit: term -> Proof.context -> Proof.context val fix_dummy_patterns: term -> Proof.context -> term * Proof.context val variant_fixes: string list -> Proof.context -> string list * Proof.context val gen_all: Proof.context -> thm -> thm val export_terms: Proof.context -> Proof.context -> term list -> term list val exportT_terms: Proof.context -> Proof.context -> term list -> term list val exportT: Proof.context -> Proof.context -> thm list -> thm list val export_prf: Proof.context -> Proof.context -> Proofterm.proof -> Proofterm.proof val export: Proof.context -> Proof.context -> thm list -> thm list val export_morphism: Proof.context -> Proof.context -> morphism val invent_types: sort list -> Proof.context -> (string * sort) list * Proof.context val importT_inst: term list -> Proof.context -> typ Term_Subst.TVars.table * Proof.context val import_inst: bool -> term list -> Proof.context -> (typ Term_Subst.TVars.table * term Term_Subst.Vars.table) * Proof.context val importT_terms: term list -> Proof.context -> term list * Proof.context val import_terms: bool -> term list -> Proof.context -> term list * Proof.context - val importT: thm list -> Proof.context -> - (((indexname * sort) * ctyp) list * thm list) * Proof.context + val importT: thm list -> Proof.context -> (ctyp Term_Subst.TVars.table * thm list) * Proof.context val import_prf: bool -> Proofterm.proof -> Proof.context -> Proofterm.proof * Proof.context val import: bool -> thm list -> Proof.context -> - ((((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list) * thm list) * Proof.context + ((ctyp Term_Subst.TVars.table * cterm Term_Subst.Vars.table) * thm list) * Proof.context val import_vars: Proof.context -> thm -> thm val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list val is_bound_focus: Proof.context -> bool val set_bound_focus: bool -> Proof.context -> Proof.context val restore_bound_focus: Proof.context -> Proof.context -> Proof.context val focus_params: binding list option -> term -> Proof.context -> (string list * (string * typ) list) * Proof.context val focus: binding list option -> term -> Proof.context -> ((string * (string * typ)) list * term) * Proof.context val focus_cterm: binding list option -> cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context val focus_subgoal: binding list option -> int -> thm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context val warn_extra_tfrees: Proof.context -> Proof.context -> unit val polymorphic_types: Proof.context -> term list -> (indexname * sort) list * term list val polymorphic: Proof.context -> term list -> term list end; structure Variable: VARIABLE = struct (** local context data **) type fixes = (string * bool) Name_Space.table; val empty_fixes: fixes = Name_Space.empty_table Markup.fixedN; datatype data = Data of {names: Name.context, (*type/term variable names*) consts: string Symtab.table, (*consts within the local scope*) bounds: int * ((string * typ) * string) list, (*next index, internal name, type, external name*) fixes: fixes, (*term fixes -- global name space, intern ~> extern*) binds: (typ * term) Vartab.table, (*term bindings*) type_occs: string list Symtab.table, (*type variables -- possibly within term variables*) maxidx: int, (*maximum var index*) constraints: typ Vartab.table * (*type constraints*) sort Vartab.table}; (*default sorts*) fun make_data (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) = Data {names = names, consts = consts, bounds = bounds, fixes = fixes, binds = binds, type_occs = type_occs, maxidx = maxidx, constraints = constraints}; val empty_data = make_data (Name.context, Symtab.empty, (0, []), empty_fixes, Vartab.empty, Symtab.empty, ~1, (Vartab.empty, Vartab.empty)); structure Data = Proof_Data ( type T = data; fun init _ = empty_data; ); fun map_data f = Data.map (fn Data {names, consts, bounds, fixes, binds, type_occs, maxidx, constraints} => make_data (f (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints))); fun map_names f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (f names, consts, bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_consts f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, f consts, bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_bounds f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, f bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_fixes f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, f fixes, binds, type_occs, maxidx, constraints)); fun map_binds f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, f binds, type_occs, maxidx, constraints)); fun map_type_occs f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, f type_occs, maxidx, constraints)); fun map_maxidx f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, type_occs, f maxidx, constraints)); fun map_constraints f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, type_occs, maxidx, f constraints)); fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep); val names_of = #names o rep_data; val fixes_of = #fixes o rep_data; val fixes_space = Name_Space.space_of_table o fixes_of; val binds_of = #binds o rep_data; val type_occs_of = #type_occs o rep_data; val maxidx_of = #maxidx o rep_data; val constraints_of = #constraints o rep_data; val is_declared = Name.is_declared o names_of; val check_name = Name_Space.base_name o tap Binding.check; (** declarations **) (* default sorts and types *) fun default_type ctxt x = Vartab.lookup (#1 (constraints_of ctxt)) (x, ~1); fun def_type ctxt pattern xi = let val {binds, constraints = (types, _), ...} = rep_data ctxt in (case Vartab.lookup types xi of NONE => if pattern then NONE else Vartab.lookup binds xi |> Option.map (Type.mark_polymorphic o #1) | some => some) end; val def_sort = Vartab.lookup o #2 o constraints_of; (* maxidx *) val declare_maxidx = map_maxidx o Integer.max; (* names *) fun declare_type_names t = map_names (fold_types (fold_atyps Term.declare_typ_names) t) #> map_maxidx (fold_types Term.maxidx_typ t); fun declare_names t = declare_type_names t #> map_names (fold_aterms Term.declare_term_frees t) #> map_maxidx (Term.maxidx_term t); (* type occurrences *) fun decl_type_occsT T = fold_atyps (fn TFree (a, _) => Symtab.default (a, []) | _ => I) T; val decl_type_occs = fold_term_types (fn Free (x, _) => fold_atyps (fn TFree (a, _) => Symtab.insert_list (op =) (a, x) | _ => I) | _ => decl_type_occsT); val declare_type_occsT = map_type_occs o fold_types decl_type_occsT; val declare_type_occs = map_type_occs o decl_type_occs; (* constraints *) fun constrain_tvar (xi, raw_S) = let val S = #2 (Term_Position.decode_positionS raw_S) in if S = dummyS then Vartab.delete_safe xi else Vartab.update (xi, S) end; fun declare_constraints t = map_constraints (fn (types, sorts) => let val types' = fold_aterms (fn Free (x, T) => Vartab.update ((x, ~1), T) | Var v => Vartab.update v | _ => I) t types; val sorts' = (fold_types o fold_atyps) (fn TFree (x, S) => constrain_tvar ((x, ~1), S) | TVar v => constrain_tvar v | _ => I) t sorts; in (types', sorts') end) #> declare_type_occsT t #> declare_type_names t; (* common declarations *) fun declare_internal t = declare_names t #> declare_type_occs t #> Thm.declare_term_sorts t; fun declare_term t = declare_internal t #> declare_constraints t; val declare_typ = declare_term o Logic.mk_type; val declare_prf = Proofterm.fold_proof_terms_types declare_internal (declare_internal o Logic.mk_type); -val declare_thm = Thm.fold_terms declare_internal; +val declare_thm = Thm.fold_terms {hyps = true} declare_internal; (* renaming term/type frees *) fun variant_frees ctxt ts frees = let val names = names_of (fold declare_names ts ctxt); val xs = fst (fold_map Name.variant (map #1 frees) names); in xs ~~ map snd frees end; (** term bindings **) fun bind_term ((x, i), t) = let val u = Term.close_schematic_term t; val U = Term.fastype_of u; in declare_term u #> map_binds (Vartab.update ((x, i), (U, u))) end; val unbind_term = map_binds o Vartab.delete_safe; fun maybe_bind_term (xi, SOME t) = bind_term (xi, t) | maybe_bind_term (xi, NONE) = unbind_term xi; fun expand_binds ctxt = let val binds = binds_of ctxt; val get = fn Var (xi, _) => Vartab.lookup binds xi | _ => NONE; in Envir.beta_norm o Envir.expand_term get end; (** consts **) val lookup_const = Symtab.lookup o #consts o rep_data; val is_const = is_some oo lookup_const; val declare_fixed = map_consts o Symtab.delete_safe; val declare_const = map_consts o Symtab.update; (** bounds **) fun next_bound (a, T) ctxt = let val b = Name.bound (#1 (#bounds (rep_data ctxt))); val ctxt' = ctxt |> map_bounds (fn (next, bounds) => (next + 1, ((b, T), a) :: bounds)); in (Free (b, T), ctxt') end; fun revert_bounds ctxt t = (case #2 (#bounds (rep_data ctxt)) of [] => t | bounds => let val names = Term.declare_term_names t (names_of ctxt); val xs = rev (#1 (fold_map Name.variant (rev (map #2 bounds)) names)); fun subst ((b, T), _) x' = (Free (b, T), Syntax_Trans.mark_bound_abs (x', T)); in Term.subst_atomic (map2 subst bounds xs) t end); (** fixes **) (* inner body mode *) val inner_body = Config.declare_bool ("inner_body", \<^here>) (K false); val is_body = Config.apply inner_body; val set_body = Config.put inner_body; val restore_body = set_body o is_body; (* proper mode *) val proper_fixes = Config.declare_bool ("proper_fixes", \<^here>) (K true); val improper_fixes = Config.put proper_fixes false; val restore_proper_fixes = Config.put proper_fixes o Config.apply proper_fixes; fun is_improper ctxt x = (case Name_Space.lookup (fixes_of ctxt) x of SOME (_, proper) => not proper | NONE => false); (* specialized name space *) val is_fixed = Name_Space.defined o fixes_of; fun is_newly_fixed inner outer = is_fixed inner andf (not o is_fixed outer); val fixed_ord = Name_Space.entry_ord o fixes_space; val intern_fixed = Name_Space.intern o fixes_space; fun lookup_fixed ctxt x = let val x' = intern_fixed ctxt x in if is_fixed ctxt x' then SOME x' else NONE end; fun revert_fixed ctxt x = (case Name_Space.lookup (fixes_of ctxt) x of SOME (x', _) => if intern_fixed ctxt x' = x then x' else x | NONE => x); fun markup_fixed ctxt x = Name_Space.markup (fixes_space ctxt) x |> Markup.name (revert_fixed ctxt x); fun markup ctxt x = if is_improper ctxt x then Markup.improper else if Name.is_skolem x then Markup.skolem else Markup.free; val markup_entity_def = Name_Space.markup_def o fixes_space; fun dest_fixes ctxt = Name_Space.fold_table (fn (x, (y, _)) => cons (y, x)) (fixes_of ctxt) [] |> sort (Name_Space.entry_ord (fixes_space ctxt) o apply2 #2); (* collect variables *) fun add_free_names ctxt = fold_aterms (fn Free (x, _) => not (is_fixed ctxt x) ? insert (op =) x | _ => I); fun add_frees ctxt = fold_aterms (fn Free (x, T) => not (is_fixed ctxt x) ? insert (op =) (x, T) | _ => I); fun add_fixed_names ctxt = fold_aterms (fn Free (x, _) => is_fixed ctxt x ? insert (op =) x | _ => I); fun add_fixed ctxt = fold_aterms (fn Free (x, T) => is_fixed ctxt x ? insert (op =) (x, T) | _ => I); fun add_newly_fixed ctxt' ctxt = fold_aterms (fn Free (x, T) => is_newly_fixed ctxt' ctxt x ? insert (op =) (x, T) | _ => I); (* declarations *) local fun err_dups dups = error ("Duplicate fixed variable(s): " ^ commas (map Binding.print dups)); fun new_fixed ((x, x'), pos) ctxt = if is_some (lookup_fixed ctxt x') then err_dups [Binding.make (x, pos)] else let val proper = Config.get ctxt proper_fixes; val context = Context.Proof ctxt |> Name_Space.map_naming (K Name_Space.global_naming) |> Context_Position.set_visible_generic false; in ctxt |> map_fixes (Name_Space.define context true (Binding.make (x', pos), (x, proper)) #> snd #> Name_Space.alias_table Name_Space.global_naming (Binding.make (x, pos)) x') |> declare_fixed x |> declare_constraints (Syntax.free x') end; fun new_fixes names' args = map_names (K names') #> fold new_fixed args #> pair (map (#2 o #1) args); in fun add_fixes_binding bs ctxt = let val _ = (case filter (Name.is_skolem o Binding.name_of) bs of [] => () | bads => error ("Illegal internal Skolem constant(s): " ^ commas (map Binding.print bads))); val _ = (case duplicates (op = o apply2 Binding.name_of) bs of [] => () | dups => err_dups dups); val xs = map check_name bs; val names = names_of ctxt; val (xs', names') = if is_body ctxt then fold_map Name.variant xs names |>> map Name.skolem else (xs, fold Name.declare xs names); in ctxt |> new_fixes names' ((xs ~~ xs') ~~ map Binding.pos_of bs) end; fun variant_names ctxt raw_xs = let val names = names_of ctxt; val xs = map (fn x => Name.clean x |> Name.is_internal x ? Name.internal) raw_xs; val (xs', names') = fold_map Name.variant xs names |>> (is_body ctxt ? map Name.skolem); in (names', xs ~~ xs') end; fun variant_fixes xs ctxt = let val (names', vs) = variant_names ctxt xs; in ctxt |> new_fixes names' (map (rpair Position.none) vs) end; fun bound_fixes xs ctxt = let val (names', vs) = variant_names ctxt (map #1 xs); val (ys, ctxt') = fold_map next_bound (map2 (fn (x', _) => fn (_, T) => (x', T)) vs xs) ctxt; val fixes = map2 (fn (x, _) => fn Free (y, _) => ((x, y), Position.none)) vs ys; in ctxt' |> new_fixes names' fixes end; end; val add_fixes = add_fixes_binding o map Binding.name; fun add_fixes_direct xs ctxt = ctxt |> set_body false |> (snd o add_fixes xs) |> restore_body ctxt; fun add_fixes_implicit t ctxt = ctxt |> not (is_body ctxt) ? add_fixes_direct (rev (add_free_names ctxt t [])); (* dummy patterns *) fun fix_dummy_patterns (Const ("Pure.dummy_pattern", T)) ctxt = let val ([x], ctxt') = ctxt |> set_body true |> add_fixes [Name.uu_] ||> restore_body ctxt in (Free (x, T), ctxt') end | fix_dummy_patterns (Abs (x, T, b)) ctxt = let val (b', ctxt') = fix_dummy_patterns b ctxt in (Abs (x, T, b'), ctxt') end | fix_dummy_patterns (t $ u) ctxt = let val (t', ctxt') = fix_dummy_patterns t ctxt; val (u', ctxt'') = fix_dummy_patterns u ctxt'; in (t' $ u', ctxt'') end | fix_dummy_patterns a ctxt = (a, ctxt); (** export -- generalize type/term variables (beware of closure sizes) **) fun gen_all ctxt th = let val i = Thm.maxidx_thm th (maxidx_of ctxt) + 1; fun gen (x, T) = Thm.forall_elim (Thm.cterm_of ctxt (Var ((x, i), T))); in fold gen (Drule.outer_params (Thm.prop_of th)) th end; fun export_inst inner outer = let val declared_outer = is_declared outer; val still_fixed = not o is_newly_fixed inner outer; val gen_fixes = - Name_Space.fold_table (fn (y, _) => not (is_fixed outer y) ? Symtab.insert_set y) - (fixes_of inner) Symtab.empty; + Symtab.build (fixes_of inner |> Name_Space.fold_table (fn (y, _) => + not (is_fixed outer y) ? Symtab.insert_set y)); val type_occs_inner = type_occs_of inner; fun gen_fixesT ts = - Symtab.fold (fn (a, xs) => + Symtab.build (fold decl_type_occs ts type_occs_inner |> Symtab.fold (fn (a, xs) => if declared_outer a orelse exists still_fixed xs - then I else Symtab.insert_set a) (fold decl_type_occs ts type_occs_inner) Symtab.empty; + then I else Symtab.insert_set a)); in (gen_fixesT, gen_fixes) end; fun exportT_inst inner outer = #1 (export_inst inner outer); fun exportT_terms inner outer = let val mk_tfrees = exportT_inst inner outer; val maxidx = maxidx_of outer; in fn ts => ts |> map (Term_Subst.generalize (mk_tfrees ts, Symtab.empty) (fold (Term.fold_types Term.maxidx_typ) ts maxidx + 1)) end; fun export_terms inner outer = let val (mk_tfrees, tfrees) = export_inst inner outer; val maxidx = maxidx_of outer; in fn ts => ts |> map (Term_Subst.generalize (mk_tfrees ts, tfrees) (fold Term.maxidx_term ts maxidx + 1)) end; fun export_prf inner outer prf = let val (mk_tfrees, frees) = export_inst (declare_prf prf inner) outer; val tfrees = mk_tfrees []; val maxidx = maxidx_of outer; val idx = Proofterm.maxidx_proof prf maxidx + 1; val gen_term = Term_Subst.generalize_same (tfrees, frees) idx; val gen_typ = Term_Subst.generalizeT_same tfrees idx; in Same.commit (Proofterm.map_proof_terms_same gen_term gen_typ) prf end; fun gen_export (mk_tfrees, frees) maxidx ths = let val tfrees = mk_tfrees (map Thm.full_prop_of ths); val idx = fold Thm.maxidx_thm ths maxidx + 1; in map (Thm.generalize (tfrees, frees) idx) ths end; fun exportT inner outer = gen_export (exportT_inst inner outer, Symtab.empty) (maxidx_of outer); fun export inner outer = gen_export (export_inst inner outer) (maxidx_of outer); fun export_morphism inner outer = let val fact = export inner outer; val term = singleton (export_terms inner outer); val typ = Logic.type_map term; in Morphism.transfer_morphism' inner $> Morphism.transfer_morphism' outer $> Morphism.morphism "Variable.export" {binding = [], typ = [typ], term = [term], fact = [fact]} end; (** import -- fix schematic type/term variables **) fun invent_types Ss ctxt = let val tfrees = Name.invent (names_of ctxt) Name.aT (length Ss) ~~ Ss; val ctxt' = fold (declare_constraints o Logic.mk_type o TFree) tfrees ctxt; in (tfrees, ctxt') end; fun importT_inst ts ctxt = let - val tvars = rev (fold Term.add_tvars ts []); + val tvars = build_rev (fold Term.add_tvars ts); val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; val instT = - fold2 (fn a => fn b => Term_Subst.TVars.add (a, TFree b)) - tvars tfrees Term_Subst.TVars.empty; + Term_Subst.TVars.build (fold2 (fn a => fn b => + Term_Subst.TVars.add (a, TFree b)) tvars tfrees); in (instT, ctxt') end; fun import_inst is_open ts ctxt = let val ren = Name.clean #> (if is_open then I else Name.internal); val (instT, ctxt') = importT_inst ts ctxt; - val vars = map (apsnd (Term_Subst.instantiateT instT)) (rev (fold Term.add_vars ts [])); + val vars = map (apsnd (Term_Subst.instantiateT instT)) (build_rev (fold Term.add_vars ts)); val (ys, ctxt'') = variant_fixes (map (ren o #1 o #1) vars) ctxt'; val inst = - fold2 (fn (x, T) => fn y => Term_Subst.Vars.add ((x, T), Free (y, T))) - vars ys Term_Subst.Vars.empty; + Term_Subst.Vars.build (fold2 (fn (x, T) => fn y => + Term_Subst.Vars.add ((x, T), Free (y, T))) vars ys); in ((instT, inst), ctxt'') end; fun importT_terms ts ctxt = let val (instT, ctxt') = importT_inst ts ctxt in (map (Term_Subst.instantiate (instT, Term_Subst.Vars.empty)) ts, ctxt') end; fun import_terms is_open ts ctxt = let val (inst, ctxt') = import_inst is_open ts ctxt in (map (Term_Subst.instantiate inst) ts, ctxt') end; fun importT ths ctxt = let val (instT, ctxt') = importT_inst (map Thm.full_prop_of ths) ctxt; - val instT' = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt' T)) instT []; - val ths' = map (Thm.instantiate (instT', [])) ths; + val instT' = Term_Subst.TVars.map (K (Thm.ctyp_of ctxt')) instT; + val ths' = map (Thm.instantiate (Term_Subst.TVars.dest instT', [])) ths; in ((instT', ths'), ctxt') end; fun import_prf is_open prf ctxt = let val ts = rev (Proofterm.fold_proof_terms_types cons (cons o Logic.mk_type) prf []); val (insts, ctxt') = import_inst is_open ts ctxt; in (Proofterm.instantiate insts prf, ctxt') end; fun import is_open ths ctxt = let val ((instT, inst), ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; - val instT' = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt' T)) instT []; - val inst' = Term_Subst.Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt' t)) inst []; - val ths' = map (Thm.instantiate (instT', inst')) ths; + val instT' = Term_Subst.TVars.map (K (Thm.ctyp_of ctxt')) instT; + val inst' = Term_Subst.Vars.map (K (Thm.cterm_of ctxt')) inst; + val ths' = map (Thm.instantiate (Term_Subst.TVars.dest instT', Term_Subst.Vars.dest inst')) ths; in (((instT', inst'), ths'), ctxt') end; fun import_vars ctxt th = let val ((_, [th']), _) = ctxt |> set_body false |> import true [th]; in th' end; (* import/export *) fun gen_trade imp exp f ctxt ths = let val ((_, ths'), ctxt') = imp ths ctxt in exp ctxt' ctxt (f ctxt' ths') end; val tradeT = gen_trade importT exportT; val trade = gen_trade (import true) export; (* focus on outermost parameters: \x y z. B *) val bound_focus = Config.declare_bool ("bound_focus", \<^here>) (K false); val is_bound_focus = Config.apply bound_focus; val set_bound_focus = Config.put bound_focus; val restore_bound_focus = set_bound_focus o is_bound_focus; fun focus_params bindings t ctxt = let val ps = Term.variant_frees t (Term.strip_all_vars t); (*as they are printed :-*) val (xs, Ts) = split_list ps; val (xs', ctxt') = (case bindings of SOME bs => ctxt |> set_body true |> add_fixes_binding bs ||> restore_body ctxt | NONE => if is_bound_focus ctxt then bound_fixes ps ctxt else variant_fixes xs ctxt); val ps' = xs' ~~ Ts; val ctxt'' = ctxt' |> fold (declare_constraints o Free) ps'; in ((xs, ps'), ctxt'') end; fun focus bindings t ctxt = let val ((xs, ps), ctxt') = focus_params bindings t ctxt; val t' = Term.subst_bounds (rev (map Free ps), Term.strip_all_body t); in (((xs ~~ ps), t'), ctxt') end; fun forall_elim_prop t prop = Thm.beta_conversion false (Thm.apply (Thm.dest_arg prop) t) |> Thm.cprop_of |> Thm.dest_arg; fun focus_cterm bindings goal ctxt = let val ((xs, ps), ctxt') = focus_params bindings (Thm.term_of goal) ctxt; val ps' = map (Thm.cterm_of ctxt' o Free) ps; val goal' = fold forall_elim_prop ps' goal; in ((xs ~~ ps', goal'), ctxt') end; fun focus_subgoal bindings i st = let - val all_vars = Thm.fold_terms Term.add_vars st []; + val all_vars = Term_Subst.Vars.build (Thm.fold_terms {hyps = false} Term_Subst.add_vars st); in - fold (unbind_term o #1) all_vars #> - fold (declare_constraints o Var) all_vars #> + Term_Subst.Vars.fold (unbind_term o #1 o #1) all_vars #> + Term_Subst.Vars.fold (declare_constraints o Var o #1) all_vars #> focus_cterm bindings (Thm.cprem_of st i) end; (** implicit polymorphism **) (* warn_extra_tfrees *) fun warn_extra_tfrees ctxt1 ctxt2 = let fun occs_typ a = Term.exists_subtype (fn TFree (b, _) => a = b | _ => false); fun occs_free a x = (case def_type ctxt1 false (x, ~1) of SOME T => if occs_typ a T then I else cons (a, x) | NONE => cons (a, x)); val occs1 = type_occs_of ctxt1; val occs2 = type_occs_of ctxt2; val extras = Symtab.fold (fn (a, xs) => if Symtab.defined occs1 a then I else fold (occs_free a) xs) occs2 []; val tfrees = map #1 extras |> sort_distinct string_ord; val frees = map #2 extras |> sort_distinct string_ord; in if null extras orelse not (Context_Position.is_visible ctxt2) then () else warning ("Introduced fixed type variable(s): " ^ commas tfrees ^ " in " ^ space_implode " or " (map quote frees)) end; (* polymorphic terms *) fun polymorphic_types ctxt ts = let val ctxt' = fold declare_term ts ctxt; val occs = type_occs_of ctxt; val occs' = type_occs_of ctxt'; val types = - Symtab.fold (fn (a, _) => - if Symtab.defined occs a then I else Symtab.insert_set a) occs' Symtab.empty; + Symtab.build (occs' |> Symtab.fold (fn (a, _) => + if Symtab.defined occs a then I else Symtab.insert_set a)); val idx = maxidx_of ctxt' + 1; val Ts' = (fold o fold_types o fold_atyps) (fn T as TFree _ => (case Term_Subst.generalizeT types idx T of TVar v => insert (op =) v | _ => I) | _ => I) ts []; val ts' = map (Term_Subst.generalize (types, Symtab.empty) idx) ts; in (rev Ts', ts') end; fun polymorphic ctxt ts = snd (polymorphic_types ctxt ts); end; diff --git a/src/Tools/Code/code_printer.ML b/src/Tools/Code/code_printer.ML --- a/src/Tools/Code/code_printer.ML +++ b/src/Tools/Code/code_printer.ML @@ -1,449 +1,448 @@ (* Title: Tools/Code/code_printer.ML Author: Florian Haftmann, TU Muenchen Generic operations for pretty printing of target language code. *) signature CODE_PRINTER = sig type itype = Code_Thingol.itype type iterm = Code_Thingol.iterm type const = Code_Thingol.const type dict = Code_Thingol.dict val eqn_error: theory -> thm option -> string -> 'a val @@ : 'a * 'a -> 'a list val @| : 'a list * 'a -> 'a list val str: string -> Pretty.T val concat: Pretty.T list -> Pretty.T val brackets: Pretty.T list -> Pretty.T val enclose: string -> string -> Pretty.T list -> Pretty.T val commas: Pretty.T list -> Pretty.T list val enum: string -> string -> string -> Pretty.T list -> Pretty.T val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T val semicolon: Pretty.T list -> Pretty.T val doublesemicolon: Pretty.T list -> Pretty.T val indent: int -> Pretty.T -> Pretty.T val markup_stmt: Code_Symbol.T -> Pretty.T -> Pretty.T val format: Code_Symbol.T list -> int -> Pretty.T -> string type var_ctxt val make_vars: string list -> var_ctxt val intro_vars: string list -> var_ctxt -> var_ctxt val lookup_var: var_ctxt -> string -> string val intro_base_names: (string -> bool) -> (string -> string) -> string list -> var_ctxt -> var_ctxt val intro_base_names_for: (string -> bool) -> (Code_Symbol.T -> string) -> iterm list -> var_ctxt -> var_ctxt val aux_params: var_ctxt -> iterm list list -> string list type literals val Literals: { literal_string: string -> string, literal_numeral: int -> string, literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string } -> literals val literal_string: literals -> string -> string val literal_numeral: literals -> int -> string val literal_list: literals -> Pretty.T list -> Pretty.T val infix_cons: literals -> int * string type lrx val L: lrx val R: lrx val X: lrx type fixity val BR: fixity val NOBR: fixity val INFX: int * lrx -> fixity val APP: fixity val brackify: fixity -> Pretty.T list -> Pretty.T val brackify_infix: int * lrx -> fixity -> Pretty.T * Pretty.T * Pretty.T -> Pretty.T val brackify_block: fixity -> Pretty.T -> Pretty.T list -> Pretty.T -> Pretty.T val gen_applify: bool -> string -> string -> ('a -> Pretty.T) -> fixity -> Pretty.T -> 'a list -> Pretty.T val applify: string -> string -> ('a -> Pretty.T) -> fixity -> Pretty.T -> 'a list -> Pretty.T val tuplify: (fixity -> 'a -> Pretty.T) -> fixity -> 'a list -> Pretty.T option type raw_const_syntax val plain_const_syntax: string -> raw_const_syntax type simple_const_syntax val simple_const_syntax: simple_const_syntax -> raw_const_syntax type complex_const_syntax val complex_const_syntax: complex_const_syntax -> raw_const_syntax val parse_const_syntax: raw_const_syntax parser val requires_args: raw_const_syntax -> int datatype const_printer = Plain_printer of string | Complex_printer of (var_ctxt -> fixity -> iterm -> Pretty.T) -> thm option -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T type const_syntax = int * const_printer val prep_const_syntax: theory -> literals -> string -> raw_const_syntax -> const_syntax type tyco_syntax val parse_tyco_syntax: tyco_syntax parser val gen_print_app: (thm option -> var_ctxt -> const * iterm list -> Pretty.T list) -> (thm option -> var_ctxt -> fixity -> iterm -> Pretty.T) -> (string -> const_syntax option) -> thm option -> var_ctxt -> fixity -> const * iterm list -> Pretty.T val gen_print_bind: (thm option -> var_ctxt -> fixity -> iterm -> Pretty.T) -> thm option -> fixity -> iterm -> var_ctxt -> Pretty.T * var_ctxt type identifiers type printings type data val empty_data: data val map_data: (string list * identifiers * printings -> string list * identifiers * printings) -> data -> data val merge_data: data * data -> data val the_reserved: data -> string list; val the_identifiers: data -> identifiers; val the_printings: data -> printings; end; structure Code_Printer : CODE_PRINTER = struct open Basic_Code_Symbol; open Code_Thingol; (** generic nonsense *) fun eqn_error thy (SOME thm) s = error (s ^ ",\nin equation " ^ Thm.string_of_thm_global thy thm) | eqn_error _ NONE s = error s; val code_presentationN = "code_presentation"; val stmt_nameN = "stmt_name"; val _ = Markup.add_mode code_presentationN YXML.output_markup; (** assembling and printing text pieces **) infixr 5 @@; infixr 5 @|; fun x @@ y = [x, y]; fun xs @| y = xs @ [y]; fun with_no_print_mode f = Print_Mode.setmp [] f; val str = with_no_print_mode Pretty.str; val concat = Pretty.block o Pretty.breaks; val commas = with_no_print_mode Pretty.commas; fun enclose l r = with_no_print_mode (Pretty.enclose l r); val brackets = enclose "(" ")" o Pretty.breaks; fun enum sep l r = with_no_print_mode (Pretty.enum sep l r); fun enum_default default sep l r [] = str default | enum_default default sep l r xs = enum sep l r xs; fun semicolon ps = Pretty.block [concat ps, str ";"]; fun doublesemicolon ps = Pretty.block [concat ps, str ";;"]; fun indent i = with_no_print_mode (Pretty.indent i); fun with_presentation_marker f = Print_Mode.setmp [code_presentationN] f; fun markup_stmt sym = with_presentation_marker (Pretty.mark (code_presentationN, [(stmt_nameN, Code_Symbol.marker sym)])); -fun filter_presentation [] tree = - Buffer.empty - |> fold XML.add_content tree - | filter_presentation presentation_syms tree = +fun filter_presentation [] xml = + Buffer.build (fold XML.add_content xml) + | filter_presentation presentation_syms xml = let val presentation_idents = map Code_Symbol.marker presentation_syms fun is_selected (name, attrs) = name = code_presentationN andalso member (op =) presentation_idents (the (Properties.get attrs stmt_nameN)); fun add_content_with_space tree (is_first, buf) = buf |> not is_first ? Buffer.add "\n\n" |> XML.add_content tree |> pair false; fun filter (XML.Elem (name_attrs, xs)) = fold (if is_selected name_attrs then add_content_with_space else filter) xs | filter (XML.Text _) = I; - in snd (fold filter tree (true, Buffer.empty)) end; + in snd (fold filter xml (true, Buffer.empty)) end; fun format presentation_names width = with_presentation_marker (Pretty.string_of_margin width) #> YXML.parse_body #> filter_presentation presentation_names #> Buffer.add "\n" #> Buffer.content; (** names and variable name contexts **) type var_ctxt = string Symtab.table * Name.context; fun make_vars names = (fold (fn name => Symtab.update_new (name, name)) names Symtab.empty, Name.make_context names); fun intro_vars names (namemap, namectxt) = let val (names', namectxt') = fold_map Name.variant names namectxt; val namemap' = fold2 (curry Symtab.update) names names' namemap; in (namemap', namectxt') end; fun lookup_var (namemap, _) name = case Symtab.lookup namemap name of SOME name' => name' | NONE => error ("Invalid name in context: " ^ quote name); fun aux_params vars lhss = let fun fish_param _ (w as SOME _) = w | fish_param (IVar (SOME v)) NONE = SOME v | fish_param _ NONE = NONE; fun fillup_param _ (_, SOME v) = v | fillup_param x (i, NONE) = x ^ string_of_int i; val fished1 = fold (map2 fish_param) lhss (replicate (length (hd lhss)) NONE); val x = singleton (Name.variant_list (map_filter I fished1)) "x"; val fished2 = map_index (fillup_param x) fished1; val (fished3, _) = fold_map Name.variant fished2 Name.context; val vars' = intro_vars fished3 vars; in map (lookup_var vars') fished3 end; fun intro_base_names no_syntax deresolve = map_filter (fn name => if no_syntax name then let val name' = deresolve name in if Long_Name.is_qualified name' then NONE else SOME name' end else NONE) #> intro_vars; fun intro_base_names_for no_syntax deresolve ts = [] |> fold Code_Thingol.add_constsyms ts |> intro_base_names (fn Constant const => no_syntax const | _ => true) deresolve; (** pretty literals **) datatype literals = Literals of { literal_string: string -> string, literal_numeral: int -> string, literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string }; fun dest_Literals (Literals lits) = lits; val literal_string = #literal_string o dest_Literals; val literal_numeral = #literal_numeral o dest_Literals; val literal_list = #literal_list o dest_Literals; val infix_cons = #infix_cons o dest_Literals; (** syntax printer **) (* binding priorities *) datatype lrx = L | R | X; datatype fixity = BR | NOBR | INFX of (int * lrx); val APP = INFX (~1, L); fun fixity_lrx L L = false | fixity_lrx R R = false | fixity_lrx _ _ = true; fun fixity NOBR _ = false | fixity _ NOBR = false | fixity (INFX (pr, lr)) (INFX (pr_ctxt, lr_ctxt)) = pr < pr_ctxt orelse pr = pr_ctxt andalso fixity_lrx lr lr_ctxt orelse pr_ctxt = ~1 | fixity BR (INFX _) = false | fixity _ _ = true; fun gen_brackify _ [p] = p | gen_brackify true (ps as _::_) = enclose "(" ")" ps | gen_brackify false (ps as _::_) = Pretty.block ps; fun brackify fxy_ctxt = gen_brackify (fixity BR fxy_ctxt) o Pretty.breaks; fun brackify_infix infx fxy_ctxt (l, m, r) = gen_brackify (fixity (INFX infx) fxy_ctxt) [l, str " ", m, Pretty.brk 1, r]; fun brackify_block fxy_ctxt p1 ps p2 = let val p = Pretty.block_enclose (p1, p2) ps in if fixity BR fxy_ctxt then enclose "(" ")" [p] else p end; fun gen_applify strict opn cls f fxy_ctxt p [] = if strict then gen_brackify (fixity BR fxy_ctxt) [p, str (opn ^ cls)] else p | gen_applify strict opn cls f fxy_ctxt p ps = gen_brackify (fixity BR fxy_ctxt) (p @@ enum "," opn cls (map f ps)); fun applify opn = gen_applify false opn; fun tuplify _ _ [] = NONE | tuplify print fxy [x] = SOME (print fxy x) | tuplify print _ xs = SOME (enum "," "(" ")" (map (print NOBR) xs)); (* generic syntax *) type simple_const_syntax = int * ((fixity -> iterm -> Pretty.T) -> fixity -> (iterm * itype) list -> Pretty.T); type complex_const_syntax = int * (literals -> (var_ctxt -> fixity -> iterm -> Pretty.T) -> thm option -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T); datatype raw_const_syntax = plain_const_syntax of string | complex_const_syntax of complex_const_syntax; fun simple_const_syntax syn = complex_const_syntax (apsnd (fn f => fn _ => fn print => fn _ => fn vars => f (print vars)) syn); fun requires_args (plain_const_syntax _) = 0 | requires_args (complex_const_syntax (k, _)) = k; datatype const_printer = Plain_printer of string | Complex_printer of (var_ctxt -> fixity -> iterm -> Pretty.T) -> thm option -> var_ctxt -> fixity -> (iterm * itype) list -> Pretty.T; type const_syntax = int * const_printer; fun prep_const_syntax thy literals c (plain_const_syntax s) = (Code.args_number thy c, Plain_printer s) | prep_const_syntax thy literals c (complex_const_syntax (n, f))= (n, Complex_printer (f literals)); fun gen_print_app print_app_expr print_term const_syntax some_thm vars fxy (app as ({ sym, dom, ... }, ts)) = case sym of Constant const => (case const_syntax const of NONE => brackify fxy (print_app_expr some_thm vars app) | SOME (_, Plain_printer s) => brackify fxy (str s :: map (print_term some_thm vars BR) ts) | SOME (k, Complex_printer print) => let fun print' fxy ts = print (print_term some_thm) some_thm vars fxy (ts ~~ take k dom); in if k = length ts then print' fxy ts else if k < length ts then case chop k ts of (ts1, ts2) => brackify fxy (print' APP ts1 :: map (print_term some_thm vars BR) ts2) else print_term some_thm vars fxy (Code_Thingol.eta_expand k app) end) | _ => brackify fxy (print_app_expr some_thm vars app); fun gen_print_bind print_term thm (fxy : fixity) pat vars = let val vs = Code_Thingol.fold_varnames (insert (op =)) pat []; val vars' = intro_vars vs vars; in (print_term thm vars' fxy pat, vars') end; type tyco_syntax = int * ((fixity -> itype -> Pretty.T) -> fixity -> itype list -> Pretty.T); (* mixfix syntax *) datatype 'a mixfix = Arg of fixity | String of string | Break; fun printer_of_mixfix prep_arg (fixity_this, mfx) = let fun is_arg (Arg _) = true | is_arg _ = false; val i = (length o filter is_arg) mfx; fun fillin _ [] [] = [] | fillin print (Arg fxy :: mfx) (a :: args) = (print fxy o prep_arg) a :: fillin print mfx args | fillin print (String s :: mfx) args = str s :: fillin print mfx args | fillin print (Break :: mfx) args = Pretty.brk 1 :: fillin print mfx args; in (i, fn print => fn fixity_ctxt => fn args => gen_brackify (fixity fixity_this fixity_ctxt) (fillin print mfx args)) end; fun read_infix (fixity_this, i) s = let val l = case fixity_this of L => INFX (i, L) | _ => INFX (i, X); val r = case fixity_this of R => INFX (i, R) | _ => INFX (i, X); in (INFX (i, fixity_this), [Arg l, String " ", String s, Break, Arg r]) end; fun read_mixfix s = let val sym_any = Scan.one Symbol.not_eof; val parse = Scan.optional ($$ "!" >> K NOBR) BR -- Scan.repeat ( ($$ "(" -- $$ "_" -- $$ ")" >> K (Arg NOBR)) || ($$ "_" >> K (Arg BR)) || ($$ "/" |-- Scan.repeat ($$ " ") >> (K Break)) || (Scan.repeat1 ( $$ "'" |-- sym_any || Scan.unless ($$ "_" || $$ "/" || $$ "(" |-- $$ "_" |-- $$ ")") sym_any) >> (String o implode))); fun err s (_, NONE) = (fn () => "malformed mixfix annotation: " ^ quote s) | err _ (_, SOME msg) = msg; in case Scan.finite Symbol.stopper parse (Symbol.explode s) of (fixity_mixfix, []) => fixity_mixfix | _ => Scan.!! (err s) Scan.fail () end; val parse_fixity = (\<^keyword>\infix\ >> K X) || (\<^keyword>\infixl\ >> K L) || (\<^keyword>\infixr\ >> K R) fun parse_mixfix x = (Parse.string >> read_mixfix || parse_fixity -- Parse.nat -- Parse.string >> (fn ((fixity, i), s) => read_infix (fixity, i) s)) x; fun syntax_of_mixfix of_plain of_printer prep_arg (BR, [String s]) = of_plain s | syntax_of_mixfix of_plain of_printer prep_arg (fixity, mfx) = of_printer (printer_of_mixfix prep_arg (fixity, mfx)); fun parse_tyco_syntax x = (parse_mixfix >> syntax_of_mixfix (fn s => (0, (K o K o K o str) s)) I I) x; val parse_const_syntax = parse_mixfix >> syntax_of_mixfix plain_const_syntax simple_const_syntax fst; (** custom data structure **) type identifiers = (string list * string, string list * string, string list * string, string list * string, string list * string, string list * string) Code_Symbol.data; type printings = (const_syntax, tyco_syntax, string, unit, unit, (Pretty.T * Code_Symbol.T list)) Code_Symbol.data; datatype data = Data of { reserved: string list, identifiers: identifiers, printings: printings }; fun make_data (reserved, identifiers, printings) = Data { reserved = reserved, identifiers = identifiers, printings = printings }; val empty_data = make_data ([], Code_Symbol.empty_data, Code_Symbol.empty_data); fun map_data f (Data { reserved, identifiers, printings }) = make_data (f (reserved, identifiers, printings)); fun merge_data (Data { reserved = reserved1, identifiers = identifiers1, printings = printings1 }, Data { reserved = reserved2, identifiers = identifiers2, printings = printings2 }) = make_data (merge (op =) (reserved1, reserved2), Code_Symbol.merge_data (identifiers1, identifiers2), Code_Symbol.merge_data (printings1, printings2)); fun the_reserved (Data { reserved, ... }) = reserved; fun the_identifiers (Data { identifiers , ... }) = identifiers; fun the_printings (Data { printings, ... }) = printings; end; (*struct*) diff --git a/src/Tools/Haskell/Haskell.thy b/src/Tools/Haskell/Haskell.thy --- a/src/Tools/Haskell/Haskell.thy +++ b/src/Tools/Haskell/Haskell.thy @@ -1,3811 +1,3818 @@ (* Title: Tools/Haskell/Haskell.thy Author: Makarius Support for Isabelle tools in Haskell. *) theory Haskell imports Main begin generate_file "Isabelle/Bytes.hs" = \ {- Title: Isabelle/Bytes.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Compact byte strings. See \<^file>\$ISABELLE_HOME/src/Pure/General/bytes.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/General/bytes.scala\. -} {-# LANGUAGE TypeApplications #-} module Isabelle.Bytes ( Bytes, make, unmake, pack, unpack, empty, null, length, index, all, any, head, last, take, drop, isPrefixOf, isSuffixOf, try_unprefix, try_unsuffix, concat, space, spaces, char, all_char, any_char, byte, singleton ) where import Prelude hiding (null, length, all, any, head, last, take, drop, concat) import qualified Data.ByteString.Short as ShortByteString import Data.ByteString.Short (ShortByteString) import qualified Data.ByteString as ByteString import Data.ByteString (ByteString) import qualified Data.List as List import Data.Word (Word8) import Data.Array (Array, array, (!)) type Bytes = ShortByteString make :: ByteString -> Bytes make = ShortByteString.toShort unmake :: Bytes -> ByteString unmake = ShortByteString.fromShort pack :: [Word8] -> Bytes pack = ShortByteString.pack unpack :: Bytes -> [Word8] unpack = ShortByteString.unpack empty :: Bytes empty = ShortByteString.empty null :: Bytes -> Bool null = ShortByteString.null length :: Bytes -> Int length = ShortByteString.length index :: Bytes -> Int -> Word8 index = ShortByteString.index all :: (Word8 -> Bool) -> Bytes -> Bool all p = List.all p . unpack any :: (Word8 -> Bool) -> Bytes -> Bool any p = List.any p . unpack head :: Bytes -> Word8 head bytes = index bytes 0 last :: Bytes -> Word8 last bytes = index bytes (length bytes - 1) take :: Int -> Bytes -> Bytes take n bs | n == 0 = empty | n >= length bs = bs | otherwise = pack (List.take n (unpack bs)) drop :: Int -> Bytes -> Bytes drop n bs | n == 0 = bs | n >= length bs = empty | otherwise = pack (List.drop n (unpack bs)) isPrefixOf :: Bytes -> Bytes -> Bool isPrefixOf bs1 bs2 = n1 <= n2 && List.all (\i -> index bs1 i == index bs2 i) [0 .. n1 - 1] where n1 = length bs1; n2 = length bs2 isSuffixOf :: Bytes -> Bytes -> Bool isSuffixOf bs1 bs2 = n1 <= n2 && List.all (\i -> index bs1 i == index bs2 (i + k)) [0 .. n1 - 1] where n1 = length bs1; n2 = length bs2; k = n2 - n1 try_unprefix :: Bytes -> Bytes -> Maybe Bytes try_unprefix bs1 bs2 = if isPrefixOf bs1 bs2 then Just (drop (length bs1) bs2) else Nothing try_unsuffix :: Bytes -> Bytes -> Maybe Bytes try_unsuffix bs1 bs2 = if isSuffixOf bs1 bs2 then Just (take (length bs2 - length bs1) bs2) else Nothing concat :: [Bytes] -> Bytes concat = mconcat space :: Word8 space = 32 small_spaces :: Array Int Bytes small_spaces = array (0, 64) [(i, pack (replicate i space)) | i <- [0 .. 64]] spaces :: Int -> Bytes spaces n = if n < 64 then small_spaces ! n else concat ((small_spaces ! (n `mod` 64)) : replicate (n `div` 64) (small_spaces ! 64)) char :: Word8 -> Char char = toEnum . fromEnum all_char :: (Char -> Bool) -> Bytes -> Bool all_char pred = all (pred . char) any_char :: (Char -> Bool) -> Bytes -> Bool any_char pred = any (pred . char) byte :: Char -> Word8 byte = toEnum . fromEnum singletons :: Array Word8 Bytes singletons = array (minBound, maxBound) [(i, make (ByteString.singleton i)) | i <- [minBound .. maxBound]] singleton :: Word8 -> Bytes singleton b = singletons ! b \ generate_file "Isabelle/UTF8.hs" = \ {- Title: Isabelle/UTF8.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Variations on UTF-8. See \<^file>\$ISABELLE_HOME/src/Pure/General/utf8.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/General/utf8.scala\. -} {-# LANGUAGE MultiParamTypeClasses #-} {-# LANGUAGE TypeSynonymInstances #-} {-# LANGUAGE FlexibleInstances #-} {-# LANGUAGE InstanceSigs #-} module Isabelle.UTF8 ( setup, setup3, Recode (..) ) where import qualified System.IO as IO import Data.Text (Text) import qualified Data.Text as Text import qualified Data.Text.Encoding as Encoding import qualified Data.Text.Encoding.Error as Error import Data.ByteString (ByteString) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) setup :: IO.Handle -> IO () setup h = do IO.hSetEncoding h IO.utf8 IO.hSetNewlineMode h IO.noNewlineTranslation setup3 :: IO.Handle -> IO.Handle -> IO.Handle -> IO () setup3 h1 h2 h3 = do setup h1 setup h2 setup h3 class Recode a b where encode :: a -> b decode :: b -> a instance Recode Text ByteString where encode :: Text -> ByteString encode = Encoding.encodeUtf8 decode :: ByteString -> Text decode = Encoding.decodeUtf8With Error.lenientDecode instance Recode Text Bytes where encode :: Text -> Bytes encode = Bytes.make . encode decode :: Bytes -> Text decode = decode . Bytes.unmake instance Recode String ByteString where encode :: String -> ByteString encode = encode . Text.pack decode :: ByteString -> String decode = Text.unpack . decode instance Recode String Bytes where encode :: String -> Bytes encode = encode . Text.pack decode :: Bytes -> String decode = Text.unpack . decode \ generate_file "Isabelle/Library.hs" = \ {- Title: Isabelle/Library.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Basic library of Isabelle idioms. See \<^file>\$ISABELLE_HOME/src/Pure/General/basics.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/library.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# LANGUAGE TypeSynonymInstances #-} {-# LANGUAGE FlexibleInstances #-} {-# LANGUAGE InstanceSigs #-} module Isabelle.Library ( (|>), (|->), (#>), (#->), fold, fold_rev, single, the_single, singletonM, map_index, get_index, separate, StringLike, STRING (..), TEXT (..), BYTES (..), show_bytes, show_text, proper_string, enclose, quote, space_implode, commas, commas_quote, cat_lines, space_explode, split_lines, trim_line, trim_split_lines, getenv, getenv_strict) where import System.Environment (lookupEnv) import Data.Maybe (fromMaybe) import qualified Data.Text as Text import Data.Text (Text) import qualified Data.Text.Lazy as Lazy import Data.String (IsString) import qualified Data.List.Split as Split import qualified Isabelle.Symbol as Symbol import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.UTF8 as UTF8 {- functions -} (|>) :: a -> (a -> b) -> b x |> f = f x (|->) :: (a, b) -> (a -> b -> c) -> c (x, y) |-> f = f x y (#>) :: (a -> b) -> (b -> c) -> a -> c (f #> g) x = x |> f |> g (#->) :: (a -> (c, b)) -> (c -> b -> d) -> a -> d (f #-> g) x = x |> f |-> g {- lists -} fold :: (a -> b -> b) -> [a] -> b -> b fold _ [] y = y fold f (x : xs) y = fold f xs (f x y) fold_rev :: (a -> b -> b) -> [a] -> b -> b fold_rev _ [] y = y fold_rev f (x : xs) y = f x (fold_rev f xs y) single :: a -> [a] single x = [x] the_single :: [a] -> a the_single [x] = x the_single _ = undefined singletonM :: Monad m => ([a] -> m [b]) -> a -> m b singletonM f x = the_single <$> f [x] map_index :: ((Int, a) -> b) -> [a] -> [b] map_index f = map_aux 0 where map_aux _ [] = [] map_aux i (x : xs) = f (i, x) : map_aux (i + 1) xs get_index :: (a -> Maybe b) -> [a] -> Maybe (Int, b) get_index f = get_aux 0 where get_aux _ [] = Nothing get_aux i (x : xs) = case f x of Nothing -> get_aux (i + 1) xs Just y -> Just (i, y) separate :: a -> [a] -> [a] separate s (x : (xs @ (_ : _))) = x : s : separate s xs separate _ xs = xs; {- string-like interfaces -} class (IsString a, Monoid a, Eq a, Ord a) => StringLike a where space_explode :: Char -> a -> [a] trim_line :: a -> a gen_trim_line :: Int -> (Int -> Char) -> (Int -> a -> a) -> a -> a gen_trim_line n at trim s = if n >= 2 && at (n - 2) == '\r' && at (n - 1) == '\n' then trim (n - 2) s else if n >= 1 && Symbol.is_ascii_line_terminator (at (n - 1)) then trim (n - 1) s else s instance StringLike String where space_explode :: Char -> String -> [String] space_explode c = Split.split (Split.dropDelims (Split.whenElt (== c))) trim_line :: String -> String trim_line s = gen_trim_line (length s) ((!!) s) take s instance StringLike Text where space_explode :: Char -> Text -> [Text] space_explode c str = if Text.null str then [] else if Text.all (/= c) str then [str] else map Text.pack $ space_explode c $ Text.unpack str trim_line :: Text -> Text trim_line s = gen_trim_line (Text.length s) (Text.index s) Text.take s instance StringLike Lazy.Text where space_explode :: Char -> Lazy.Text -> [Lazy.Text] space_explode c str = if Lazy.null str then [] else if Lazy.all (/= c) str then [str] else map Lazy.pack $ space_explode c $ Lazy.unpack str trim_line :: Lazy.Text -> Lazy.Text trim_line = Lazy.fromStrict . trim_line . Lazy.toStrict instance StringLike Bytes where space_explode :: Char -> Bytes -> [Bytes] space_explode c str = if Bytes.null str then [] else if Bytes.all_char (/= c) str then [str] else explode (Bytes.unpack str) where explode rest = case span (/= (Bytes.byte c)) rest of (_, []) -> [Bytes.pack rest] (prfx, _ : rest') -> Bytes.pack prfx : explode rest' trim_line :: Bytes -> Bytes trim_line s = gen_trim_line (Bytes.length s) (Bytes.char . Bytes.index s) Bytes.take s class StringLike a => STRING a where make_string :: a -> String instance STRING String where make_string = id instance STRING Text where make_string = Text.unpack instance STRING Lazy.Text where make_string = Lazy.unpack instance STRING Bytes where make_string = UTF8.decode class StringLike a => TEXT a where make_text :: a -> Text instance TEXT String where make_text = Text.pack instance TEXT Text where make_text = id instance TEXT Lazy.Text where make_text = Lazy.toStrict instance TEXT Bytes where make_text = UTF8.decode class StringLike a => BYTES a where make_bytes :: a -> Bytes instance BYTES String where make_bytes = UTF8.encode instance BYTES Text where make_bytes = UTF8.encode instance BYTES Lazy.Text where make_bytes = UTF8.encode . Lazy.toStrict instance BYTES Bytes where make_bytes = id show_bytes :: Show a => a -> Bytes show_bytes = make_bytes . show show_text :: Show a => a -> Text show_text = make_text . show {- strings -} proper_string :: StringLike a => a -> Maybe a proper_string s = if s == "" then Nothing else Just s enclose :: StringLike a => a -> a -> a -> a enclose lpar rpar str = lpar <> str <> rpar quote :: StringLike a => a -> a quote = enclose "\"" "\"" space_implode :: StringLike a => a -> [a] -> a space_implode s = mconcat . separate s commas, commas_quote :: StringLike a => [a] -> a commas = space_implode ", " commas_quote = commas . map quote split_lines :: StringLike a => a -> [a] split_lines = space_explode '\n' cat_lines :: StringLike a => [a] -> a cat_lines = space_implode "\n" trim_split_lines :: StringLike a => a -> [a] trim_split_lines = trim_line #> split_lines #> map trim_line {- getenv -} getenv :: Bytes -> IO Bytes getenv x = do y <- lookupEnv (make_string x) return $ make_bytes $ fromMaybe "" y getenv_strict :: Bytes -> IO Bytes getenv_strict x = do y <- getenv x if Bytes.null y then errorWithoutStackTrace $ make_string ("Undefined Isabelle environment variable: " <> quote x) else return y \ generate_file "Isabelle/Symbol.hs" = \ {- Title: Isabelle/Symbols.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Isabelle text symbols. See \<^file>\$ISABELLE_HOME/src/Pure/General/symbol.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/General/symbol_explode.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Symbol ( Symbol, eof, is_eof, not_eof, is_ascii_letter, is_ascii_digit, is_ascii_hex, is_ascii_quasi, is_ascii_blank, is_ascii_line_terminator, is_ascii_letdig, is_ascii_identifier, explode ) where import Data.Word (Word8) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) {- type -} type Symbol = Bytes eof :: Symbol eof = "" is_eof, not_eof :: Symbol -> Bool is_eof = Bytes.null not_eof = not . is_eof {- ASCII characters -} is_ascii_letter :: Char -> Bool is_ascii_letter c = 'A' <= c && c <= 'Z' || 'a' <= c && c <= 'z' is_ascii_digit :: Char -> Bool is_ascii_digit c = '0' <= c && c <= '9' is_ascii_hex :: Char -> Bool is_ascii_hex c = '0' <= c && c <= '9' || 'A' <= c && c <= 'F' || 'a' <= c && c <= 'f' is_ascii_quasi :: Char -> Bool is_ascii_quasi c = c == '_' || c == '\'' is_ascii_blank :: Char -> Bool is_ascii_blank c = c `elem` (" \t\n\11\f\r" :: String) is_ascii_line_terminator :: Char -> Bool is_ascii_line_terminator c = c == '\r' || c == '\n' is_ascii_letdig :: Char -> Bool is_ascii_letdig c = is_ascii_letter c || is_ascii_digit c || is_ascii_quasi c is_ascii_identifier :: String -> Bool is_ascii_identifier s = not (null s) && is_ascii_letter (head s) && all is_ascii_letdig s {- explode symbols: ASCII, UTF8, named -} is_utf8 :: Word8 -> Bool is_utf8 b = b >= 128 is_utf8_trailer :: Word8 -> Bool is_utf8_trailer b = 128 <= b && b < 192 is_utf8_control :: Word8 -> Bool is_utf8_control b = 128 <= b && b < 160 (|>) :: a -> (a -> b) -> b x |> f = f x explode :: Bytes -> [Symbol] explode string = scan 0 where byte = Bytes.index string substring i j = if i == j - 1 then Bytes.singleton (byte i) else Bytes.pack (map byte [i .. j - 1]) n = Bytes.length string test pred i = i < n && pred (byte i) test_char pred i = i < n && pred (Bytes.char (byte i)) many pred i = if test pred i then many pred (i + 1) else i maybe_char c i = if test_char (== c) i then i + 1 else i maybe_ascii_id i = if test_char is_ascii_letter i then many (is_ascii_letdig . Bytes.char) (i + 1) else i scan i = if i < n then let b = byte i c = Bytes.char b in {-encoded newline-} if c == '\r' then "\n" : scan (maybe_char '\n' (i + 1)) {-pseudo utf8: encoded ascii control-} else if b == 192 && test is_utf8_control (i + 1) && not (test is_utf8 (i + 2)) then Bytes.singleton (byte (i + 1) - 128) : scan (i + 2) {-utf8-} else if is_utf8 b then let j = many is_utf8_trailer (i + 1) in substring i j : scan j {-named symbol-} else if c == '\\' && test_char (== '<') (i + 1) then let j = (i + 2) |> maybe_char '^' |> maybe_ascii_id |> maybe_char '>' in substring i j : scan j {-single character-} else Bytes.singleton b : scan (i + 1) else [] \ generate_file "Isabelle/Buffer.hs" = \ {- Title: Isabelle/Buffer.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Efficient buffer of byte strings. See \<^file>\$ISABELLE_HOME/src/Pure/General/buffer.ML\. -} -module Isabelle.Buffer (T, empty, add, content) +module Isabelle.Buffer (T, empty, add, content, build, build_content) where import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) +import Isabelle.Library newtype T = Buffer [Bytes] empty :: T empty = Buffer [] add :: Bytes -> T -> T add b (Buffer bs) = Buffer (if Bytes.null b then bs else b : bs) content :: T -> Bytes content (Buffer bs) = Bytes.concat (reverse bs) + +build :: (T -> T) -> T +build f = f empty + +build_content :: (T -> T) -> Bytes +build_content f = build f |> content \ generate_file "Isabelle/Value.hs" = \ {- Title: Isabelle/Value.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Plain values, represented as string. See \<^file>\$ISABELLE_HOME/src/Pure/General/value.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Value (print_bool, parse_bool, parse_nat, print_int, parse_int, print_real, parse_real) where import qualified Data.List as List import qualified Text.Read as Read import Isabelle.Bytes (Bytes) import Isabelle.Library {- bool -} print_bool :: Bool -> Bytes print_bool True = "true" print_bool False = "false" parse_bool :: Bytes -> Maybe Bool parse_bool "true" = Just True parse_bool "false" = Just False parse_bool _ = Nothing {- nat -} parse_nat :: Bytes -> Maybe Int parse_nat s = case Read.readMaybe (make_string s) of Just n | n >= 0 -> Just n _ -> Nothing {- int -} print_int :: Int -> Bytes print_int = show_bytes parse_int :: Bytes -> Maybe Int parse_int = Read.readMaybe . make_string {- real -} print_real :: Double -> Bytes print_real x = let s = show x in case span (/= '.') s of (a, '.' : b) | List.all (== '0') b -> make_bytes a _ -> make_bytes s parse_real :: Bytes -> Maybe Double parse_real = Read.readMaybe . make_string \ generate_file "Isabelle/Properties.hs" = \ {- Title: Isabelle/Properties.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Property lists. See \<^file>\$ISABELLE_HOME/src/Pure/General/properties.ML\. -} module Isabelle.Properties (Entry, T, defined, get, get_value, put, remove) where import qualified Data.List as List import Isabelle.Bytes (Bytes) type Entry = (Bytes, Bytes) type T = [Entry] defined :: T -> Bytes -> Bool defined props name = any (\(a, _) -> a == name) props get :: T -> Bytes -> Maybe Bytes get props name = List.lookup name props get_value :: (Bytes -> Maybe a) -> T -> Bytes -> Maybe a get_value parse props name = case get props name of Nothing -> Nothing Just s -> parse s put :: Entry -> T -> T put entry props = entry : remove (fst entry) props remove :: Bytes -> T -> T remove name props = if defined props name then filter (\(a, _) -> a /= name) props else props \ generate_file "Isabelle/Markup.hs" = \ {- Title: Isabelle/Markup.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Quasi-abstract markup elements. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/markup.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.Markup ( T, empty, is_empty, properties, nameN, name, xnameN, xname, kindN, bindingN, binding, entityN, entity, defN, refN, completionN, completion, no_completionN, no_completion, lineN, end_lineN, offsetN, end_offsetN, fileN, idN, positionN, position, position_properties, def_name, expressionN, expression, citationN, citation, pathN, path, urlN, url, docN, doc, markupN, consistentN, unbreakableN, indentN, widthN, blockN, block, breakN, break, fbreakN, fbreak, itemN, item, wordsN, words, tfreeN, tfree, tvarN, tvar, freeN, free, skolemN, skolem, boundN, bound, varN, var, numeralN, numeral, literalN, literal, delimiterN, delimiter, inner_stringN, inner_string, inner_cartoucheN, inner_cartouche, token_rangeN, token_range, sortingN, sorting, typingN, typing, class_parameterN, class_parameter, antiquotedN, antiquoted, antiquoteN, antiquote, paragraphN, paragraph, text_foldN, text_fold, keyword1N, keyword1, keyword2N, keyword2, keyword3N, keyword3, quasi_keywordN, quasi_keyword, improperN, improper, operatorN, operator, stringN, string, alt_stringN, alt_string, verbatimN, verbatim, cartoucheN, cartouche, commentN, comment, comment1N, comment1, comment2N, comment2, comment3N, comment3, forkedN, forked, joinedN, joined, runningN, running, finishedN, finished, failedN, failed, canceledN, canceled, initializedN, initialized, finalizedN, finalized, consolidatedN, consolidated, writelnN, writeln, stateN, state, informationN, information, tracingN, tracing, warningN, warning, legacyN, legacy, errorN, error, reportN, report, no_reportN, no_report, intensifyN, intensify, Output, no_output) where import Prelude hiding (words, error, break) import Data.Map.Strict (Map) import qualified Data.Map.Strict as Map import Isabelle.Library import qualified Isabelle.Properties as Properties import qualified Isabelle.Value as Value import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) {- basic markup -} type T = (Bytes, Properties.T) empty :: T empty = ("", []) is_empty :: T -> Bool is_empty ("", _) = True is_empty _ = False properties :: Properties.T -> T -> T properties more_props (elem, props) = (elem, fold_rev Properties.put more_props props) markup_elem :: Bytes -> T markup_elem name = (name, []) markup_string :: Bytes -> Bytes -> Bytes -> T markup_string name prop = \s -> (name, [(prop, s)]) {- misc properties -} nameN :: Bytes nameN = \Markup.nameN\ name :: Bytes -> T -> T name a = properties [(nameN, a)] xnameN :: Bytes xnameN = \Markup.xnameN\ xname :: Bytes -> T -> T xname a = properties [(xnameN, a)] kindN :: Bytes kindN = \Markup.kindN\ {- formal entities -} bindingN :: Bytes bindingN = \Markup.bindingN\ binding :: T binding = markup_elem bindingN entityN :: Bytes entityN = \Markup.entityN\ entity :: Bytes -> Bytes -> T entity kind name = (entityN, (if Bytes.null name then [] else [(nameN, name)]) <> (if Bytes.null kind then [] else [(kindN, kind)])) defN :: Bytes defN = \Markup.defN\ refN :: Bytes refN = \Markup.refN\ {- completion -} completionN :: Bytes completionN = \Markup.completionN\ completion :: T completion = markup_elem completionN no_completionN :: Bytes no_completionN = \Markup.no_completionN\ no_completion :: T no_completion = markup_elem no_completionN {- position -} lineN, end_lineN :: Bytes lineN = \Markup.lineN\ end_lineN = \Markup.end_lineN\ offsetN, end_offsetN :: Bytes offsetN = \Markup.offsetN\ end_offsetN = \Markup.end_offsetN\ fileN, idN :: Bytes fileN = \Markup.fileN\ idN = \Markup.idN\ positionN :: Bytes positionN = \Markup.positionN\ position :: T position = markup_elem positionN position_properties :: [Bytes] position_properties = [lineN, offsetN, end_offsetN, fileN, idN] {- position "def" names -} make_def :: Bytes -> Bytes make_def a = "def_" <> a def_names :: Map Bytes Bytes def_names = Map.fromList $ map (\a -> (a, make_def a)) position_properties def_name :: Bytes -> Bytes def_name a = case Map.lookup a def_names of Just b -> b Nothing -> make_def a {- expression -} expressionN :: Bytes expressionN = \Markup.expressionN\ expression :: Bytes -> T expression kind = (expressionN, if kind == "" then [] else [(kindN, kind)]) {- citation -} citationN :: Bytes citationN = \Markup.citationN\ citation :: Bytes -> T citation = markup_string nameN citationN {- external resources -} pathN :: Bytes pathN = \Markup.pathN\ path :: Bytes -> T path = markup_string pathN nameN urlN :: Bytes urlN = \Markup.urlN\ url :: Bytes -> T url = markup_string urlN nameN docN :: Bytes docN = \Markup.docN\ doc :: Bytes -> T doc = markup_string docN nameN {- pretty printing -} markupN, consistentN, unbreakableN, indentN :: Bytes markupN = \Markup.markupN\ consistentN = \Markup.consistentN\ unbreakableN = \Markup.unbreakableN\ indentN = \Markup.indentN\ widthN :: Bytes widthN = \Markup.widthN\ blockN :: Bytes blockN = \Markup.blockN\ block :: Bool -> Int -> T block c i = (blockN, (if c then [(consistentN, Value.print_bool c)] else []) <> (if i /= 0 then [(indentN, Value.print_int i)] else [])) breakN :: Bytes breakN = \Markup.breakN\ break :: Int -> Int -> T break w i = (breakN, (if w /= 0 then [(widthN, Value.print_int w)] else []) <> (if i /= 0 then [(indentN, Value.print_int i)] else [])) fbreakN :: Bytes fbreakN = \Markup.fbreakN\ fbreak :: T fbreak = markup_elem fbreakN itemN :: Bytes itemN = \Markup.itemN\ item :: T item = markup_elem itemN {- text properties -} wordsN :: Bytes wordsN = \Markup.wordsN\ words :: T words = markup_elem wordsN {- inner syntax -} tfreeN :: Bytes tfreeN = \Markup.tfreeN\ tfree :: T tfree = markup_elem tfreeN tvarN :: Bytes tvarN = \Markup.tvarN\ tvar :: T tvar = markup_elem tvarN freeN :: Bytes freeN = \Markup.freeN\ free :: T free = markup_elem freeN skolemN :: Bytes skolemN = \Markup.skolemN\ skolem :: T skolem = markup_elem skolemN boundN :: Bytes boundN = \Markup.boundN\ bound :: T bound = markup_elem boundN varN :: Bytes varN = \Markup.varN\ var :: T var = markup_elem varN numeralN :: Bytes numeralN = \Markup.numeralN\ numeral :: T numeral = markup_elem numeralN literalN :: Bytes literalN = \Markup.literalN\ literal :: T literal = markup_elem literalN delimiterN :: Bytes delimiterN = \Markup.delimiterN\ delimiter :: T delimiter = markup_elem delimiterN inner_stringN :: Bytes inner_stringN = \Markup.inner_stringN\ inner_string :: T inner_string = markup_elem inner_stringN inner_cartoucheN :: Bytes inner_cartoucheN = \Markup.inner_cartoucheN\ inner_cartouche :: T inner_cartouche = markup_elem inner_cartoucheN token_rangeN :: Bytes token_rangeN = \Markup.token_rangeN\ token_range :: T token_range = markup_elem token_rangeN sortingN :: Bytes sortingN = \Markup.sortingN\ sorting :: T sorting = markup_elem sortingN typingN :: Bytes typingN = \Markup.typingN\ typing :: T typing = markup_elem typingN class_parameterN :: Bytes class_parameterN = \Markup.class_parameterN\ class_parameter :: T class_parameter = markup_elem class_parameterN {- antiquotations -} antiquotedN :: Bytes antiquotedN = \Markup.antiquotedN\ antiquoted :: T antiquoted = markup_elem antiquotedN antiquoteN :: Bytes antiquoteN = \Markup.antiquoteN\ antiquote :: T antiquote = markup_elem antiquoteN {- text structure -} paragraphN :: Bytes paragraphN = \Markup.paragraphN\ paragraph :: T paragraph = markup_elem paragraphN text_foldN :: Bytes text_foldN = \Markup.text_foldN\ text_fold :: T text_fold = markup_elem text_foldN {- outer syntax -} keyword1N :: Bytes keyword1N = \Markup.keyword1N\ keyword1 :: T keyword1 = markup_elem keyword1N keyword2N :: Bytes keyword2N = \Markup.keyword2N\ keyword2 :: T keyword2 = markup_elem keyword2N keyword3N :: Bytes keyword3N = \Markup.keyword3N\ keyword3 :: T keyword3 = markup_elem keyword3N quasi_keywordN :: Bytes quasi_keywordN = \Markup.quasi_keywordN\ quasi_keyword :: T quasi_keyword = markup_elem quasi_keywordN improperN :: Bytes improperN = \Markup.improperN\ improper :: T improper = markup_elem improperN operatorN :: Bytes operatorN = \Markup.operatorN\ operator :: T operator = markup_elem operatorN stringN :: Bytes stringN = \Markup.stringN\ string :: T string = markup_elem stringN alt_stringN :: Bytes alt_stringN = \Markup.alt_stringN\ alt_string :: T alt_string = markup_elem alt_stringN verbatimN :: Bytes verbatimN = \Markup.verbatimN\ verbatim :: T verbatim = markup_elem verbatimN cartoucheN :: Bytes cartoucheN = \Markup.cartoucheN\ cartouche :: T cartouche = markup_elem cartoucheN commentN :: Bytes commentN = \Markup.commentN\ comment :: T comment = markup_elem commentN {- comments -} comment1N :: Bytes comment1N = \Markup.comment1N\ comment1 :: T comment1 = markup_elem comment1N comment2N :: Bytes comment2N = \Markup.comment2N\ comment2 :: T comment2 = markup_elem comment2N comment3N :: Bytes comment3N = \Markup.comment3N\ comment3 :: T comment3 = markup_elem comment3N {- command status -} forkedN, joinedN, runningN, finishedN, failedN, canceledN, initializedN, finalizedN, consolidatedN :: Bytes forkedN = \Markup.forkedN\ joinedN = \Markup.joinedN\ runningN = \Markup.runningN\ finishedN = \Markup.finishedN\ failedN = \Markup.failedN\ canceledN = \Markup.canceledN\ initializedN = \Markup.initializedN\ finalizedN = \Markup.finalizedN\ consolidatedN = \Markup.consolidatedN\ forked, joined, running, finished, failed, canceled, initialized, finalized, consolidated :: T forked = markup_elem forkedN joined = markup_elem joinedN running = markup_elem runningN finished = markup_elem finishedN failed = markup_elem failedN canceled = markup_elem canceledN initialized = markup_elem initializedN finalized = markup_elem finalizedN consolidated = markup_elem consolidatedN {- messages -} writelnN :: Bytes writelnN = \Markup.writelnN\ writeln :: T writeln = markup_elem writelnN stateN :: Bytes stateN = \Markup.stateN\ state :: T state = markup_elem stateN informationN :: Bytes informationN = \Markup.informationN\ information :: T information = markup_elem informationN tracingN :: Bytes tracingN = \Markup.tracingN\ tracing :: T tracing = markup_elem tracingN warningN :: Bytes warningN = \Markup.warningN\ warning :: T warning = markup_elem warningN legacyN :: Bytes legacyN = \Markup.legacyN\ legacy :: T legacy = markup_elem legacyN errorN :: Bytes errorN = \Markup.errorN\ error :: T error = markup_elem errorN reportN :: Bytes reportN = \Markup.reportN\ report :: T report = markup_elem reportN no_reportN :: Bytes no_reportN = \Markup.no_reportN\ no_report :: T no_report = markup_elem no_reportN intensifyN :: Bytes intensifyN = \Markup.intensifyN\ intensify :: T intensify = markup_elem intensifyN {- output -} type Output = (Bytes, Bytes) no_output :: Output no_output = ("", "") \ generate_file "Isabelle/Position.hs" = \ {- Title: Isabelle/Position.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Source positions starting from 1; values <= 0 mean "absent". Count Isabelle symbols, not UTF8 bytes nor UTF16 characters. Position range specifies a right-open interval offset .. end_offset (exclusive). See \<^file>\$ISABELLE_HOME/src/Pure/General/position.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Position ( T, line_of, column_of, offset_of, end_offset_of, file_of, id_of, start, none, put_file, file, file_only, put_id, id, id_only, symbol, symbol_explode, symbol_explode_string, shift_offsets, of_properties, properties_of, def_properties_of, entity_markup, make_entity_markup, Report, Report_Text, is_reported, is_reported_range, here, Range, no_range, no_range_position, range_position, range ) where import Prelude hiding (id) import Data.Maybe (isJust, fromMaybe) import Data.Bifunctor (first) import qualified Isabelle.Properties as Properties import qualified Isabelle.Bytes as Bytes import qualified Isabelle.Value as Value import Isabelle.Bytes (Bytes) import qualified Isabelle.Markup as Markup import qualified Isabelle.YXML as YXML import Isabelle.Library import qualified Isabelle.Symbol as Symbol import Isabelle.Symbol (Symbol) {- position -} data T = Position { _line :: Int, _column :: Int, _offset :: Int, _end_offset :: Int, _file :: Bytes, _id :: Bytes } deriving (Eq, Ord) valid, invalid :: Int -> Bool valid i = i > 0 invalid = not . valid maybe_valid :: Int -> Maybe Int maybe_valid i = if valid i then Just i else Nothing if_valid :: Int -> Int -> Int if_valid i i' = if valid i then i' else i {- fields -} line_of, column_of, offset_of, end_offset_of :: T -> Maybe Int line_of = maybe_valid . _line column_of = maybe_valid . _column offset_of = maybe_valid . _offset end_offset_of = maybe_valid . _end_offset file_of :: T -> Maybe Bytes file_of = proper_string . _file id_of :: T -> Maybe Bytes id_of = proper_string . _id {- make position -} start :: T start = Position 1 1 1 0 Bytes.empty Bytes.empty none :: T none = Position 0 0 0 0 Bytes.empty Bytes.empty put_file :: Bytes -> T -> T put_file file pos = pos { _file = file } file :: Bytes -> T file file = put_file file start file_only :: Bytes -> T file_only file = put_file file none put_id :: Bytes -> T -> T put_id id pos = pos { _id = id } id :: Bytes -> T id id = put_id id start id_only :: Bytes -> T id_only id = put_id id none {- count position -} count_line :: Symbol -> Int -> Int count_line "\n" line = if_valid line (line + 1) count_line _ line = line count_column :: Symbol -> Int -> Int count_column "\n" column = if_valid column 1 count_column s column = if Symbol.not_eof s then if_valid column (column + 1) else column count_offset :: Symbol -> Int -> Int count_offset s offset = if Symbol.not_eof s then if_valid offset (offset + 1) else offset symbol :: Symbol -> T -> T symbol s pos = pos { _line = count_line s (_line pos), _column = count_column s (_column pos), _offset = count_offset s (_offset pos) } symbol_explode :: BYTES a => a -> T -> T symbol_explode = fold symbol . Symbol.explode . make_bytes symbol_explode_string :: String -> T -> T symbol_explode_string = symbol_explode {- shift offsets -} shift_offsets :: Int -> T -> T shift_offsets shift pos = pos { _offset = offset', _end_offset = end_offset' } where offset = _offset pos end_offset = _end_offset pos offset' = if invalid offset || invalid shift then offset else offset + shift end_offset' = if invalid end_offset || invalid shift then end_offset else end_offset + shift {- markup properties -} get_string :: Properties.T -> Bytes -> Bytes get_string props name = fromMaybe "" (Properties.get_value Just props name) get_int :: Properties.T -> Bytes -> Int get_int props name = fromMaybe 0 (Properties.get_value Value.parse_int props name) of_properties :: Properties.T -> T of_properties props = none { _line = get_int props Markup.lineN, _offset = get_int props Markup.offsetN, _end_offset = get_int props Markup.end_offsetN, _file = get_string props Markup.fileN, _id = get_string props Markup.idN } string_entry :: Bytes -> Bytes -> Properties.T string_entry k s = if Bytes.null s then [] else [(k, s)] int_entry :: Bytes -> Int -> Properties.T int_entry k i = if invalid i then [] else [(k, Value.print_int i)] properties_of :: T -> Properties.T properties_of pos = int_entry Markup.lineN (_line pos) ++ int_entry Markup.offsetN (_offset pos) ++ int_entry Markup.end_offsetN (_end_offset pos) ++ string_entry Markup.fileN (_file pos) ++ string_entry Markup.idN (_id pos) def_properties_of :: T -> Properties.T def_properties_of = properties_of #> map (first Markup.def_name) entity_markup :: Bytes -> (Bytes, T) -> Markup.T entity_markup kind (name, pos) = Markup.entity kind name |> Markup.properties (def_properties_of pos) make_entity_markup :: Bool -> Int -> Bytes -> (Bytes, T) -> Markup.T make_entity_markup def serial kind (name, pos) = let props = if def then (Markup.defN, Value.print_int serial) : properties_of pos else (Markup.refN, Value.print_int serial) : def_properties_of pos in Markup.entity kind name |> Markup.properties props {- reports -} type Report = (T, Markup.T) type Report_Text = (Report, Bytes) is_reported :: T -> Bool is_reported pos = isJust (offset_of pos) && isJust (id_of pos) is_reported_range :: T -> Bool is_reported_range pos = is_reported pos && isJust (end_offset_of pos) {- here: user output -} here :: T -> Bytes here pos = if Bytes.null s2 then "" else s1 <> m1 <> s2 <> m2 where props = properties_of pos (m1, m2) = YXML.output_markup (Markup.properties props Markup.position) (s1, s2) = case (line_of pos, file_of pos) of (Just i, Nothing) -> (" ", "(line " <> Value.print_int i <> ")") (Just i, Just name) -> (" ", "(line " <> Value.print_int i <> " of " <> quote name <> ")") (Nothing, Just name) -> (" ", "(file " <> quote name <> ")") _ -> if is_reported pos then ("", "\092<^here>") else ("", "") {- range -} type Range = (T, T) no_range :: Range no_range = (none, none) no_range_position :: T -> T no_range_position pos = pos { _end_offset = 0 } range_position :: Range -> T range_position (pos, pos') = pos { _end_offset = _offset pos' } range :: Range -> Range range (pos, pos') = (range_position (pos, pos'), no_range_position pos') \ generate_file "Isabelle/XML.hs" = \ {- Title: Isabelle/XML.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Untyped XML trees and representation of ML values. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/xml.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.XML (Attributes, Body, Tree(..), wrap_elem, unwrap_elem, content_of) where import Isabelle.Library import qualified Isabelle.Properties as Properties import qualified Isabelle.Markup as Markup import qualified Isabelle.Buffer as Buffer import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) {- types -} type Attributes = Properties.T type Body = [Tree] data Tree = Elem (Markup.T, Body) | Text Bytes {- wrapped elements -} wrap_elem :: ((Markup.T, Body), [Tree]) -> Tree wrap_elem (((a, atts), body1), body2) = Elem ((\XML.xml_elemN\, (\XML.xml_nameN\, a) : atts), Elem ((\XML.xml_bodyN\, []), body1) : body2) unwrap_elem :: Tree -> Maybe ((Markup.T, Body), [Tree]) unwrap_elem (Elem ((\XML.xml_elemN\, (\XML.xml_nameN\, a) : atts), Elem ((\XML.xml_bodyN\, []), body1) : body2)) = Just (((a, atts), body1), body2) unwrap_elem _ = Nothing {- text content -} add_content :: Tree -> Buffer.T -> Buffer.T add_content tree = case unwrap_elem tree of Just (_, ts) -> fold add_content ts Nothing -> case tree of Elem (_, ts) -> fold add_content ts Text s -> Buffer.add s content_of :: Body -> Bytes -content_of body = Buffer.empty |> fold add_content body |> Buffer.content +content_of = Buffer.build_content . fold add_content {- string representation -} encode_char :: Char -> String encode_char '<' = "<" encode_char '>' = ">" encode_char '&' = "&" encode_char '\'' = "'" encode_char '\"' = """ encode_char c = [c] encode_text :: Bytes -> Bytes encode_text = make_bytes . concatMap (encode_char . Bytes.char) . Bytes.unpack instance Show Tree where show tree = - Buffer.empty |> show_tree tree |> Buffer.content |> make_string + make_string $ Buffer.build_content (show_tree tree) where show_tree (Elem ((name, atts), [])) = Buffer.add "<" #> Buffer.add (show_elem name atts) #> Buffer.add "/>" show_tree (Elem ((name, atts), ts)) = Buffer.add "<" #> Buffer.add (show_elem name atts) #> Buffer.add ">" #> fold show_tree ts #> Buffer.add " Buffer.add name #> Buffer.add ">" show_tree (Text s) = Buffer.add (encode_text s) show_elem name atts = space_implode " " (name : map (\(a, x) -> a <> "=\"" <> encode_text x <> "\"") atts) \ generate_file "Isabelle/XML/Encode.hs" = \ {- Title: Isabelle/XML/Encode.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) XML as data representation language. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/xml.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.XML.Encode ( A, T, V, P, int_atom, bool_atom, unit_atom, tree, properties, string, int, bool, unit, pair, triple, list, option, variant ) where import Data.Maybe (fromJust) import Isabelle.Library import Isabelle.Bytes (Bytes) import qualified Isabelle.Value as Value import qualified Isabelle.Properties as Properties import qualified Isabelle.XML as XML type A a = a -> Bytes type T a = a -> XML.Body type V a = a -> Maybe ([Bytes], XML.Body) type P a = a -> [Bytes] -- atomic values int_atom :: A Int int_atom = Value.print_int bool_atom :: A Bool bool_atom False = "0" bool_atom True = "1" unit_atom :: A () unit_atom () = "" -- structural nodes node ts = XML.Elem ((":", []), ts) vector = map_index (\(i, x) -> (int_atom i, x)) tagged (tag, (xs, ts)) = XML.Elem ((int_atom tag, vector xs), ts) -- representation of standard types tree :: T XML.Tree tree t = [t] properties :: T Properties.T properties props = [XML.Elem ((":", props), [])] string :: T Bytes string "" = [] string s = [XML.Text s] int :: T Int int = string . int_atom bool :: T Bool bool = string . bool_atom unit :: T () unit = string . unit_atom pair :: T a -> T b -> T (a, b) pair f g (x, y) = [node (f x), node (g y)] triple :: T a -> T b -> T c -> T (a, b, c) triple f g h (x, y, z) = [node (f x), node (g y), node (h z)] list :: T a -> T [a] list f xs = map (node . f) xs option :: T a -> T (Maybe a) option _ Nothing = [] option f (Just x) = [node (f x)] variant :: [V a] -> T a variant fs x = [tagged (fromJust (get_index (\f -> f x) fs))] \ generate_file "Isabelle/XML/Decode.hs" = \ {- Title: Isabelle/XML/Decode.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) XML as data representation language. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/xml.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.XML.Decode ( A, T, V, P, int_atom, bool_atom, unit_atom, tree, properties, string, int, bool, unit, pair, triple, list, option, variant ) where import Isabelle.Library import Isabelle.Bytes (Bytes) import qualified Isabelle.Value as Value import qualified Isabelle.Properties as Properties import qualified Isabelle.XML as XML type A a = Bytes -> a type T a = XML.Body -> a type V a = ([Bytes], XML.Body) -> a type P a = [Bytes] -> a err_atom = error "Malformed XML atom" err_body = error "Malformed XML body" {- atomic values -} int_atom :: A Int int_atom s = case Value.parse_int s of Just i -> i Nothing -> err_atom bool_atom :: A Bool bool_atom "0" = False bool_atom "1" = True bool_atom _ = err_atom unit_atom :: A () unit_atom "" = () unit_atom _ = err_atom {- structural nodes -} node (XML.Elem ((":", []), ts)) = ts node _ = err_body vector atts = map_index (\(i, (a, x)) -> if int_atom a == i then x else err_atom) atts tagged (XML.Elem ((name, atts), ts)) = (int_atom name, (vector atts, ts)) tagged _ = err_body {- representation of standard types -} tree :: T XML.Tree tree [t] = t tree _ = err_body properties :: T Properties.T properties [XML.Elem ((":", props), [])] = props properties _ = err_body string :: T Bytes string [] = "" string [XML.Text s] = s string _ = err_body int :: T Int int = int_atom . string bool :: T Bool bool = bool_atom . string unit :: T () unit = unit_atom . string pair :: T a -> T b -> T (a, b) pair f g [t1, t2] = (f (node t1), g (node t2)) pair _ _ _ = err_body triple :: T a -> T b -> T c -> T (a, b, c) triple f g h [t1, t2, t3] = (f (node t1), g (node t2), h (node t3)) triple _ _ _ _ = err_body list :: T a -> T [a] list f ts = map (f . node) ts option :: T a -> T (Maybe a) option _ [] = Nothing option f [t] = Just (f (node t)) option _ _ = err_body variant :: [V a] -> T a variant fs [t] = (fs !! tag) (xs, ts) where (tag, (xs, ts)) = tagged t variant _ _ = err_body \ generate_file "Isabelle/YXML.hs" = \ {- Title: Isabelle/YXML.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Efficient text representation of XML trees. Suitable for direct inlining into plain text. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/yxml.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures -fno-warn-incomplete-patterns #-} module Isabelle.YXML (charX, charY, strX, strY, detect, output_markup, buffer_body, buffer, string_of_body, string_of, parse_body, parse) where import qualified Data.List as List import Data.Word (Word8) import Isabelle.Library import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Markup as Markup import qualified Isabelle.XML as XML import qualified Isabelle.Buffer as Buffer {- markers -} charX, charY :: Word8 charX = 5 charY = 6 strX, strY, strXY, strXYX :: Bytes strX = Bytes.singleton charX strY = Bytes.singleton charY strXY = strX <> strY strXYX = strXY <> strX detect :: Bytes -> Bool detect = Bytes.any (\c -> c == charX || c == charY) {- output -} output_markup :: Markup.T -> Markup.Output output_markup markup@(name, atts) = if Markup.is_empty markup then Markup.no_output else (strXY <> name <> Bytes.concat (map (\(a, x) -> strY <> a <> "=" <> x) atts) <> strX, strXYX) buffer_attrib (a, x) = Buffer.add strY #> Buffer.add a #> Buffer.add "=" #> Buffer.add x buffer_body :: XML.Body -> Buffer.T -> Buffer.T buffer_body = fold buffer buffer :: XML.Tree -> Buffer.T -> Buffer.T buffer (XML.Elem ((name, atts), ts)) = Buffer.add strXY #> Buffer.add name #> fold buffer_attrib atts #> Buffer.add strX #> buffer_body ts #> Buffer.add strXYX buffer (XML.Text s) = Buffer.add s string_of_body :: XML.Body -> Bytes -string_of_body body = Buffer.empty |> buffer_body body |> Buffer.content +string_of_body = Buffer.build_content . buffer_body string_of :: XML.Tree -> Bytes string_of = string_of_body . single {- parse -} -- split: fields or non-empty tokens split :: Bool -> Word8 -> [Word8] -> [[Word8]] split _ _ [] = [] split fields sep str = splitting str where splitting rest = case span (/= sep) rest of (_, []) -> cons rest [] (prfx, _ : rest') -> cons prfx (splitting rest') cons item = if fields || not (null item) then (:) item else id -- structural errors err :: Bytes -> a err msg = error (make_string ("Malformed YXML: " <> msg)) err_attribute = err "bad attribute" err_element = err "bad element" err_unbalanced :: Bytes -> a err_unbalanced name = if Bytes.null name then err "unbalanced element" else err ("unbalanced element " <> quote name) -- stack operations add x ((elem, body) : pending) = (elem, x : body) : pending push name atts pending = if Bytes.null name then err_element else ((name, atts), []) : pending pop (((name, atts), body) : pending) = if Bytes.null name then err_unbalanced name else add (XML.Elem ((name, atts), reverse body)) pending -- parsing parse_attrib s = case List.elemIndex (Bytes.byte '=') s of Just i | i > 0 -> (Bytes.pack $ take i s, Bytes.pack $ drop (i + 1) s) _ -> err_attribute parse_chunk [[], []] = pop parse_chunk ([] : name : atts) = push (Bytes.pack name) (map parse_attrib atts) parse_chunk txts = fold (add . XML.Text . Bytes.pack) txts parse_body :: Bytes -> XML.Body parse_body source = case fold parse_chunk chunks [((Bytes.empty, []), [])] of [((name, _), result)] | Bytes.null name -> reverse result ((name, _), _) : _ -> err_unbalanced name where chunks = source |> Bytes.unpack |> split False charX |> map (split True charY) parse :: Bytes -> XML.Tree parse source = case parse_body source of [result] -> result [] -> XML.Text "" _ -> err "multiple results" \ generate_file "Isabelle/Completion.hs" = \ {- Title: Isabelle/Completion.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Completion of names. See \<^file>\$ISABELLE_HOME/src/Pure/General/completion.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Completion ( Name, T, names, none, make, markup_element, markup_report, make_report ) where import qualified Isabelle.Bytes as Bytes import qualified Isabelle.Name as Name import Isabelle.Name (Name) import qualified Isabelle.Properties as Properties import qualified Isabelle.Markup as Markup import Isabelle.XML.Classes import qualified Isabelle.XML as XML import qualified Isabelle.YXML as YXML type Names = [(Name, (Name, Name))] -- external name, kind, internal name data T = Completion Properties.T Int Names -- position, total length, names names :: Int -> Properties.T -> Names -> T names limit props names = Completion props (length names) (take limit names) none :: T none = names 0 [] [] make :: Int -> (Name, Properties.T) -> ((Name -> Bool) -> Names) -> T make limit (name, props) make_names = if name /= "" && name /= "_" then names limit props (make_names (Bytes.isPrefixOf (Name.clean name))) else none markup_element :: T -> (Markup.T, XML.Body) markup_element (Completion props total names) = if not (null names) then (Markup.properties props Markup.completion, encode (total, names)) else (Markup.empty, []) markup_report :: [T] -> Name markup_report [] = Bytes.empty markup_report elems = YXML.string_of $ XML.Elem (Markup.report, map (XML.Elem . markup_element) elems) make_report :: Int -> (Name, Properties.T) -> ((Name -> Bool) -> Names) -> Name make_report limit name_props make_names = markup_report [make limit name_props make_names] \ generate_file "Isabelle/File.hs" = \ {- Title: Isabelle/File.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) File-system operations. See \<^file>\$ISABELLE_HOME/src/Pure/General/file.ML\. -} module Isabelle.File (read, write, append) where import Prelude hiding (read) import qualified System.IO as IO import qualified Data.ByteString as ByteString import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) read :: IO.FilePath -> IO Bytes read path = Bytes.make <$> IO.withFile path IO.ReadMode ByteString.hGetContents write :: IO.FilePath -> Bytes -> IO () write path bs = IO.withFile path IO.WriteMode (\h -> ByteString.hPut h (Bytes.unmake bs)) append :: IO.FilePath -> Bytes -> IO () append path bs = IO.withFile path IO.AppendMode (\h -> ByteString.hPut h (Bytes.unmake bs)) \ generate_file "Isabelle/Pretty.hs" = \ {- Title: Isabelle/Pretty.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Generic pretty printing module. See \<^file>\$ISABELLE_HOME/src/Pure/General/pretty.ML\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-missing-signatures #-} module Isabelle.Pretty ( T, symbolic, formatted, unformatted, str, brk_indent, brk, fbrk, breaks, fbreaks, blk, block, strs, markup, mark, mark_str, marks_str, item, text_fold, keyword1, keyword2, text, paragraph, para, quote, cartouche, separate, commas, enclose, enum, list, str_list, big_list) where import qualified Data.List as List import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import Isabelle.Library hiding (enclose, quote, separate, commas) import qualified Isabelle.Buffer as Buffer import qualified Isabelle.Markup as Markup import qualified Isabelle.XML as XML import qualified Isabelle.YXML as YXML data T = Block Markup.T Bool Int [T] | Break Int Int | Str Bytes {- output -} symbolic_text s = if Bytes.null s then [] else [XML.Text s] symbolic_markup markup body = if Markup.is_empty markup then body else [XML.Elem (markup, body)] symbolic :: T -> XML.Body symbolic (Block markup consistent indent prts) = concatMap symbolic prts |> symbolic_markup block_markup |> symbolic_markup markup where block_markup = if null prts then Markup.empty else Markup.block consistent indent symbolic (Break wd ind) = [XML.Elem (Markup.break wd ind, symbolic_text (Bytes.spaces wd))] symbolic (Str s) = symbolic_text s formatted :: T -> Bytes formatted = YXML.string_of_body . symbolic unformatted :: T -> Bytes -unformatted prt = Buffer.empty |> out prt |> Buffer.content +unformatted = Buffer.build_content . out where out (Block markup _ _ prts) = let (bg, en) = YXML.output_markup markup in Buffer.add bg #> fold out prts #> Buffer.add en out (Break _ wd) = Buffer.add (Bytes.spaces wd) out (Str s) = Buffer.add s {- derived operations to create formatting expressions -} force_nat n | n < 0 = 0 force_nat n = n str :: BYTES a => a -> T str = Str . make_bytes brk_indent :: Int -> Int -> T brk_indent wd ind = Break (force_nat wd) ind brk :: Int -> T brk wd = brk_indent wd 0 fbrk :: T fbrk = Str "\n" breaks, fbreaks :: [T] -> [T] breaks = List.intersperse (brk 1) fbreaks = List.intersperse fbrk blk :: (Int, [T]) -> T blk (indent, es) = Block Markup.empty False (force_nat indent) es block :: [T] -> T block prts = blk (2, prts) strs :: BYTES a => [a] -> T strs = block . breaks . map str markup :: Markup.T -> [T] -> T markup m = Block m False 0 mark :: Markup.T -> T -> T mark m prt = if m == Markup.empty then prt else markup m [prt] mark_str :: BYTES a => (Markup.T, a) -> T mark_str (m, s) = mark m (str s) marks_str :: BYTES a => ([Markup.T], a) -> T marks_str (ms, s) = fold_rev mark ms (str s) item :: [T] -> T item = markup Markup.item text_fold :: [T] -> T text_fold = markup Markup.text_fold keyword1, keyword2 :: BYTES a => a -> T keyword1 name = mark_str (Markup.keyword1, name) keyword2 name = mark_str (Markup.keyword2, name) text :: BYTES a => a -> [T] text = breaks . map str . filter (not . Bytes.null) . space_explode ' ' . make_bytes paragraph :: [T] -> T paragraph = markup Markup.paragraph para :: BYTES a => a -> T para = paragraph . text quote :: T -> T quote prt = blk (1, [Str "\"", prt, Str "\""]) cartouche :: T -> T cartouche prt = blk (1, [Str "\92", prt, Str "\92"]) separate :: BYTES a => a -> [T] -> [T] separate sep = List.intercalate [str sep, brk 1] . map single commas :: [T] -> [T] commas = separate ("," :: Bytes) enclose :: BYTES a => a -> a -> [T] -> T enclose lpar rpar prts = block (str lpar : prts <> [str rpar]) enum :: BYTES a => a -> a -> a -> [T] -> T enum sep lpar rpar = enclose lpar rpar . separate sep list :: BYTES a => a -> a -> [T] -> T list = enum "," str_list :: BYTES a => a -> a -> [a] -> T str_list lpar rpar = list lpar rpar . map str big_list :: BYTES a => a -> [T] -> T big_list name prts = block (fbreaks (str name : prts)) \ generate_file "Isabelle/Name.hs" = \ {- Title: Isabelle/Name.hs Author: Makarius Names of basic logical entities (variables etc.). See \<^file>\$ISABELLE_HOME/src/Pure/name.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Name ( Name, uu, uu_, aT, clean_index, clean, internal, skolem, is_internal, is_skolem, dest_internal, dest_skolem, Context, declare, is_declared, context, make_context, variant ) where import Data.Word (Word8) import qualified Data.Set as Set import Data.Set (Set) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Symbol as Symbol import Isabelle.Library type Name = Bytes {- common defaults -} uu, uu_, aT :: Name uu = "uu" uu_ = "uu_" aT = "'a" {- internal names -- NB: internal subsumes skolem -} underscore :: Word8 underscore = Bytes.byte '_' internal, skolem :: Name -> Name internal x = x <> "_" skolem x = x <> "__" is_internal, is_skolem :: Name -> Bool is_internal = Bytes.isSuffixOf "_" is_skolem = Bytes.isSuffixOf "__" dest_internal, dest_skolem :: Name -> Maybe Name dest_internal = Bytes.try_unsuffix "_" dest_skolem = Bytes.try_unsuffix "__" clean_index :: Name -> (Name, Int) clean_index x = if Bytes.null x || Bytes.last x /= underscore then (x, 0) else let rev = reverse (Bytes.unpack x) n = length (takeWhile (== underscore) rev) y = Bytes.pack (reverse (drop n rev)) in (y, n) clean :: Name -> Name clean = fst . clean_index {- context for used names -} newtype Context = Context (Set Name) declare :: Name -> Context -> Context declare x (Context names) = Context (Set.insert x names) is_declared :: Context -> Name -> Bool is_declared (Context names) x = Set.member x names context :: Context context = Context (Set.fromList ["", "'"]) make_context :: [Name] -> Context make_context used = fold declare used context {- generating fresh names -} bump_init :: Name -> Name bump_init str = str <> "a" bump_string :: Name -> Name bump_string str = let a = Bytes.byte 'a' z = Bytes.byte 'z' bump (b : bs) | b == z = a : bump bs bump (b : bs) | a <= b && b < z = b + 1 : bs bump bs = a : bs rev = reverse (Bytes.unpack str) part2 = reverse (takeWhile (Symbol.is_ascii_quasi . Bytes.char) rev) part1 = reverse (bump (drop (length part2) rev)) in Bytes.pack (part1 <> part2) variant :: Name -> Context -> (Name, Context) variant name names = let vary bump x = if is_declared names x then bump x |> vary bump_string else x (x, n) = clean_index name x' = x |> vary bump_init names' = declare x' names; in (x' <> Bytes.pack (replicate n underscore), names') \ generate_file "Isabelle/Term.hs" = \ {- Title: Isabelle/Term.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Lambda terms, types, sorts. See \<^file>\$ISABELLE_HOME/src/Pure/term.scala\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Term ( Indexname, Sort, Typ(..), Term(..), Free, lambda, declare_frees, incr_boundvars, subst_bound, dest_lambda, strip_lambda, type_op0, type_op1, op0, op1, op2, typed_op0, typed_op1, typed_op2, binder, dummyS, dummyT, is_dummyT, propT, is_propT, (-->), dest_funT, (--->), aconv, list_comb, strip_comb, head_of ) where import Isabelle.Library import qualified Isabelle.Name as Name import Isabelle.Name (Name) infixr 5 --> infixr ---> {- types and terms -} type Indexname = (Name, Int) type Sort = [Name] data Typ = Type (Name, [Typ]) | TFree (Name, Sort) | TVar (Indexname, Sort) deriving (Show, Eq, Ord) data Term = Const (Name, [Typ]) | Free (Name, Typ) | Var (Indexname, Typ) | Bound Int | Abs (Name, Typ, Term) | App (Term, Term) deriving (Show, Eq, Ord) {- free and bound variables -} type Free = (Name, Typ) lambda :: Free -> Term -> Term lambda (name, typ) body = Abs (name, typ, abstract 0 body) where abstract lev (Free (x, ty)) | name == x && typ == ty = Bound lev abstract lev (Abs (a, ty, t)) = Abs (a, ty, abstract (lev + 1) t) abstract lev (App (t, u)) = App (abstract lev t, abstract lev u) abstract _ t = t declare_frees :: Term -> Name.Context -> Name.Context declare_frees (Free (x, _)) = Name.declare x declare_frees (Abs (_, _, b)) = declare_frees b declare_frees (App (t, u)) = declare_frees t #> declare_frees u declare_frees _ = id incr_boundvars :: Int -> Term -> Term incr_boundvars inc = if inc == 0 then id else incr 0 where incr lev (Bound i) = if i >= lev then Bound (i + inc) else Bound i incr lev (Abs (a, ty, b)) = Abs (a, ty, incr (lev + 1) b) incr lev (App (t, u)) = App (incr lev t, incr lev u) incr _ t = t subst_bound :: Term -> Term -> Term subst_bound arg = subst 0 where subst lev (Bound i) = if i < lev then Bound i else if i == lev then incr_boundvars lev arg else Bound (i - 1) subst lev (Abs (a, ty, b)) = Abs (a, ty, subst (lev + 1) b) subst lev (App (t, u)) = App (subst lev t, subst lev u) subst _ t = t dest_lambda :: Name.Context -> Term -> Maybe (Free, Term) dest_lambda names (Abs (x, ty, b)) = let (x', _) = Name.variant x (declare_frees b names) v = (x', ty) in Just (v, subst_bound (Free v) b) dest_lambda _ _ = Nothing strip_lambda :: Name.Context -> Term -> ([Free], Term) strip_lambda names tm = case dest_lambda names tm of Just (v, t) -> let (vs, t') = strip_lambda names t' in (v : vs, t') Nothing -> ([], tm) {- type and term operators -} type_op0 :: Name -> (Typ, Typ -> Bool) type_op0 name = (mk, is) where mk = Type (name, []) is (Type (c, _)) = c == name is _ = False type_op1 :: Name -> (Typ -> Typ, Typ -> Maybe Typ) type_op1 name = (mk, dest) where mk ty = Type (name, [ty]) dest (Type (c, [ty])) | c == name = Just ty dest _ = Nothing type_op2 :: Name -> (Typ -> Typ -> Typ, Typ -> Maybe (Typ, Typ)) type_op2 name = (mk, dest) where mk ty1 ty2 = Type (name, [ty1, ty2]) dest (Type (c, [ty1, ty2])) | c == name = Just (ty1, ty2) dest _ = Nothing op0 :: Name -> (Term, Term -> Bool) op0 name = (mk, is) where mk = Const (name, []) is (Const (c, _)) = c == name is _ = False op1 :: Name -> (Term -> Term, Term -> Maybe Term) op1 name = (mk, dest) where mk t = App (Const (name, []), t) dest (App (Const (c, _), t)) | c == name = Just t dest _ = Nothing op2 :: Name -> (Term -> Term -> Term, Term -> Maybe (Term, Term)) op2 name = (mk, dest) where mk t u = App (App (Const (name, []), t), u) dest (App (App (Const (c, _), t), u)) | c == name = Just (t, u) dest _ = Nothing typed_op0 :: Name -> (Typ -> Term, Term -> Maybe Typ) typed_op0 name = (mk, dest) where mk ty = Const (name, [ty]) dest (Const (c, [ty])) | c == name = Just ty dest _ = Nothing typed_op1 :: Name -> (Typ -> Term -> Term, Term -> Maybe (Typ, Term)) typed_op1 name = (mk, dest) where mk ty t = App (Const (name, [ty]), t) dest (App (Const (c, [ty]), t)) | c == name = Just (ty, t) dest _ = Nothing typed_op2 :: Name -> (Typ -> Term -> Term -> Term, Term -> Maybe (Typ, Term, Term)) typed_op2 name = (mk, dest) where mk ty t u = App (App (Const (name, [ty]), t), u) dest (App (App (Const (c, [ty]), t), u)) | c == name = Just (ty, t, u) dest _ = Nothing binder :: Name -> (Free -> Term -> Term, Name.Context -> Term -> Maybe (Free, Term)) binder name = (mk, dest) where mk (a, ty) b = App (Const (name, [ty]), lambda (a, ty) b) dest names (App (Const (c, _), t)) | c == name = dest_lambda names t dest _ _ = Nothing {- type operations -} dummyS :: Sort dummyS = [""] dummyT :: Typ; is_dummyT :: Typ -> Bool (dummyT, is_dummyT) = type_op0 \\<^type_name>\dummy\\ propT :: Typ; is_propT :: Typ -> Bool (propT, is_propT) = type_op0 \\<^type_name>\prop\\ (-->) :: Typ -> Typ -> Typ; dest_funT :: Typ -> Maybe (Typ, Typ) ((-->), dest_funT) = type_op2 \\<^type_name>\fun\\ (--->) :: [Typ] -> Typ -> Typ [] ---> b = b (a : as) ---> b = a --> (as ---> b) {- term operations -} aconv :: Term -> Term -> Bool aconv (App (t1, u1)) (App (t2, u2)) = aconv t1 t2 && aconv u1 u2 aconv (Abs (_, ty1, t1)) (Abs (_, ty2, t2)) = aconv t1 t2 && ty1 == ty2 aconv a1 a2 = a1 == a2 list_comb :: Term -> [Term] -> Term list_comb f [] = f list_comb f (t : ts) = list_comb (App (f, t)) ts strip_comb :: Term -> (Term, [Term]) strip_comb tm = strip (tm, []) where strip (App (f, t), ts) = strip (f, t : ts) strip x = x head_of :: Term -> Term head_of (App (f, _)) = head_of f head_of u = u \ generate_file "Isabelle/Pure.hs" = \ {- Title: Isabelle/Term.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Support for Isabelle/Pure logic. See \<^file>\$ISABELLE_HOME/src/Pure/logic.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Pure ( mk_forall_op, dest_forall_op, mk_forall, dest_forall, mk_equals, dest_equals, mk_implies, dest_implies ) where import qualified Isabelle.Name as Name import Isabelle.Term mk_forall_op :: Typ -> Term -> Term; dest_forall_op :: Term -> Maybe (Typ, Term) (mk_forall_op, dest_forall_op) = typed_op1 \\<^const_name>\Pure.all\\ mk_forall :: Free -> Term -> Term; dest_forall :: Name.Context -> Term -> Maybe (Free, Term) (mk_forall, dest_forall) = binder \\<^const_name>\Pure.all\\ mk_equals :: Typ -> Term -> Term -> Term; dest_equals :: Term -> Maybe (Typ, Term, Term) (mk_equals, dest_equals) = typed_op2 \\<^const_name>\Pure.eq\\ mk_implies :: Term -> Term -> Term; dest_implies :: Term -> Maybe (Term, Term) (mk_implies, dest_implies) = op2 \\<^const_name>\Pure.imp\\ \ generate_file "Isabelle/HOL.hs" = \ {- Title: Isabelle/Term.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Support for Isabelle/HOL logic. See \<^file>\$ISABELLE_HOME/src/HOL/Tools/hologic.ML\. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.HOL ( boolT, is_boolT, mk_trueprop, dest_trueprop, mk_setT, dest_setT, mk_mem, dest_mem, mk_eq, dest_eq, true, is_true, false, is_false, mk_not, dest_not, mk_conj, dest_conj, mk_disj, dest_disj, mk_imp, dest_imp, mk_iff, dest_iff, mk_all_op, dest_all_op, mk_ex_op, dest_ex_op, mk_all, dest_all, mk_ex, dest_ex, mk_undefined, dest_undefined ) where import qualified Isabelle.Name as Name import Isabelle.Term boolT :: Typ; is_boolT :: Typ -> Bool (boolT, is_boolT) = type_op0 \\<^type_name>\bool\\ mk_trueprop :: Term -> Term; dest_trueprop :: Term -> Maybe Term (mk_trueprop, dest_trueprop) = op1 \\<^const_name>\Trueprop\\ mk_setT :: Typ -> Typ; dest_setT :: Typ -> Maybe Typ (mk_setT, dest_setT) = type_op1 \\<^type_name>\set\\ mk_mem :: Typ -> Term -> Term -> Term; dest_mem :: Term -> Maybe (Typ, Term, Term) (mk_mem, dest_mem) = typed_op2 \\<^const_name>\Set.member\\ mk_eq :: Typ -> Term -> Term -> Term; dest_eq :: Term -> Maybe (Typ, Term, Term) (mk_eq, dest_eq) = typed_op2 \\<^const_name>\HOL.eq\\ true :: Term; is_true :: Term -> Bool (true, is_true) = op0 \\<^const_name>\True\\ false :: Term; is_false :: Term -> Bool (false, is_false) = op0 \\<^const_name>\False\\ mk_not :: Term -> Term; dest_not :: Term -> Maybe Term (mk_not, dest_not) = op1 \\<^const_name>\Not\\ mk_conj :: Term -> Term -> Term; dest_conj :: Term -> Maybe (Term, Term) (mk_conj, dest_conj) = op2 \\<^const_name>\conj\\ mk_disj :: Term -> Term -> Term; dest_disj :: Term -> Maybe (Term, Term) (mk_disj, dest_disj) = op2 \\<^const_name>\disj\\ mk_imp :: Term -> Term -> Term; dest_imp :: Term -> Maybe (Term, Term) (mk_imp, dest_imp) = op2 \\<^const_name>\implies\\ mk_iff :: Term -> Term -> Term mk_iff = mk_eq boolT dest_iff :: Term -> Maybe (Term, Term) dest_iff tm = case dest_eq tm of Just (ty, t, u) | ty == boolT -> Just (t, u) _ -> Nothing mk_all_op :: Typ -> Term -> Term; dest_all_op :: Term -> Maybe (Typ, Term) (mk_all_op, dest_all_op) = typed_op1 \\<^const_name>\All\\ mk_ex_op :: Typ -> Term -> Term; dest_ex_op :: Term -> Maybe (Typ, Term) (mk_ex_op, dest_ex_op) = typed_op1 \\<^const_name>\Ex\\ mk_all :: Free -> Term -> Term; dest_all :: Name.Context -> Term -> Maybe (Free, Term) (mk_all, dest_all) = binder \\<^const_name>\All\\ mk_ex :: Free -> Term -> Term; dest_ex :: Name.Context -> Term -> Maybe (Free, Term) (mk_ex, dest_ex) = binder \\<^const_name>\Ex\\ mk_undefined :: Typ -> Term; dest_undefined :: Term -> Maybe Typ (mk_undefined, dest_undefined) = typed_op0 \\<^const_name>\undefined\\ \ generate_file "Isabelle/Term_XML/Encode.hs" = \ {- Title: Isabelle/Term_XML/Encode.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) XML data representation of lambda terms. See \<^file>\$ISABELLE_HOME/src/Pure/term_xml.ML\. -} {-# LANGUAGE LambdaCase #-} module Isabelle.Term_XML.Encode (indexname, sort, typ, typ_body, term) where import Isabelle.Library import Isabelle.XML.Encode import Isabelle.Term indexname :: P Indexname indexname (a, b) = if b == 0 then [a] else [a, int_atom b] sort :: T Sort sort = list string typ :: T Typ typ ty = ty |> variant [\case { Type (a, b) -> Just ([a], list typ b); _ -> Nothing }, \case { TFree (a, b) -> Just ([a], sort b); _ -> Nothing }, \case { TVar (a, b) -> Just (indexname a, sort b); _ -> Nothing }] typ_body :: T Typ typ_body ty = if is_dummyT ty then [] else typ ty term :: T Term term t = t |> variant [\case { Const (a, b) -> Just ([a], list typ b); _ -> Nothing }, \case { Free (a, b) -> Just ([a], typ_body b); _ -> Nothing }, \case { Var (a, b) -> Just (indexname a, typ_body b); _ -> Nothing }, \case { Bound a -> Just ([], int a); _ -> Nothing }, \case { Abs (a, b, c) -> Just ([a], pair typ term (b, c)); _ -> Nothing }, \case { App a -> Just ([], pair term term a); _ -> Nothing }] \ generate_file "Isabelle/Term_XML/Decode.hs" = \ {- Title: Isabelle/Term_XML/Decode.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) XML data representation of lambda terms. See \<^file>\$ISABELLE_HOME/src/Pure/term_xml.ML\. -} {-# OPTIONS_GHC -fno-warn-incomplete-patterns #-} module Isabelle.Term_XML.Decode (indexname, sort, typ, typ_body, term) where import Isabelle.Library import Isabelle.XML.Decode import Isabelle.Term indexname :: P Indexname indexname [a] = (a, 0) indexname [a, b] = (a, int_atom b) sort :: T Sort sort = list string typ :: T Typ typ ty = ty |> variant [\([a], b) -> Type (a, list typ b), \([a], b) -> TFree (a, sort b), \(a, b) -> TVar (indexname a, sort b)] typ_body :: T Typ typ_body [] = dummyT typ_body body = typ body term :: T Term term t = t |> variant [\([a], b) -> Const (a, list typ b), \([a], b) -> Free (a, typ_body b), \(a, b) -> Var (indexname a, typ_body b), \([], a) -> Bound (int a), \([a], b) -> let (c, d) = pair typ term b in Abs (a, c, d), \([], a) -> App (pair term term a)] \ generate_file "Isabelle/XML/Classes.hs" = \ {- generated by Isabelle -} {- Title: Isabelle/XML/Classes.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Type classes for XML data representation. -} {-# LANGUAGE TypeSynonymInstances #-} {-# LANGUAGE FlexibleInstances #-} module Isabelle.XML.Classes (Encode_Atom(..), Decode_Atom(..), Encode (..), Decode (..)) where import qualified Isabelle.XML as XML import qualified Isabelle.XML.Encode as Encode import qualified Isabelle.XML.Decode as Decode import qualified Isabelle.Term_XML.Encode as Encode import qualified Isabelle.Term_XML.Decode as Decode import qualified Isabelle.Properties as Properties import Isabelle.Bytes (Bytes) import Isabelle.Term (Typ, Term) class Encode_Atom a where encode_atom :: Encode.A a class Decode_Atom a where decode_atom :: Decode.A a instance Encode_Atom Int where encode_atom = Encode.int_atom instance Decode_Atom Int where decode_atom = Decode.int_atom instance Encode_Atom Bool where encode_atom = Encode.bool_atom instance Decode_Atom Bool where decode_atom = Decode.bool_atom instance Encode_Atom () where encode_atom = Encode.unit_atom instance Decode_Atom () where decode_atom = Decode.unit_atom class Encode a where encode :: Encode.T a class Decode a where decode :: Decode.T a instance Encode Bytes where encode = Encode.string instance Decode Bytes where decode = Decode.string instance Encode Int where encode = Encode.int instance Decode Int where decode = Decode.int instance Encode Bool where encode = Encode.bool instance Decode Bool where decode = Decode.bool instance Encode () where encode = Encode.unit instance Decode () where decode = Decode.unit instance (Encode a, Encode b) => Encode (a, b) where encode = Encode.pair encode encode instance (Decode a, Decode b) => Decode (a, b) where decode = Decode.pair decode decode instance (Encode a, Encode b, Encode c) => Encode (a, b, c) where encode = Encode.triple encode encode encode instance (Decode a, Decode b, Decode c) => Decode (a, b, c) where decode = Decode.triple decode decode decode instance Encode a => Encode [a] where encode = Encode.list encode instance Decode a => Decode [a] where decode = Decode.list decode instance Encode a => Encode (Maybe a) where encode = Encode.option encode instance Decode a => Decode (Maybe a) where decode = Decode.option decode instance Encode XML.Tree where encode = Encode.tree instance Decode XML.Tree where decode = Decode.tree instance Encode Properties.T where encode = Encode.properties instance Decode Properties.T where decode = Decode.properties instance Encode Typ where encode = Encode.typ instance Decode Typ where decode = Decode.typ instance Encode Term where encode = Encode.term instance Decode Term where decode = Decode.term \ generate_file "Isabelle/UUID.hs" = \ {- Title: Isabelle/UUID.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Universally unique identifiers. See \<^file>\$ISABELLE_HOME/src/Pure/General/uuid.scala\. -} module Isabelle.UUID (T, random, print, parse) where import Prelude hiding (print) import Data.UUID (UUID) import qualified Data.UUID as UUID import Data.UUID.V4 (nextRandom) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) type T = UUID random :: IO T random = nextRandom print :: T -> Bytes print = Bytes.make . UUID.toASCIIBytes parse :: Bytes -> Maybe T parse = UUID.fromASCIIBytes . Bytes.unmake \ generate_file "Isabelle/Byte_Message.hs" = \ {- Title: Isabelle/Byte_Message.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Byte-oriented messages. See \<^file>\$ISABELLE_HOME/src/Pure/PIDE/byte_message.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/PIDE/byte_message.scala\. -} {-# LANGUAGE OverloadedStrings #-} {-# OPTIONS_GHC -fno-warn-incomplete-patterns #-} module Isabelle.Byte_Message ( write, write_line, read, read_block, read_line, make_message, write_message, read_message, exchange_message, exchange_message0, make_line_message, write_line_message, read_line_message, read_yxml, write_yxml ) where import Prelude hiding (read) import Data.Maybe import qualified Data.ByteString as ByteString import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Symbol as Symbol import qualified Isabelle.UTF8 as UTF8 import qualified Isabelle.XML as XML import qualified Isabelle.YXML as YXML import Network.Socket (Socket) import qualified Network.Socket.ByteString as Socket import Isabelle.Library import qualified Isabelle.Value as Value {- output operations -} write :: Socket -> [Bytes] -> IO () write socket = Socket.sendMany socket . map Bytes.unmake write_line :: Socket -> Bytes -> IO () write_line socket s = write socket [s, "\n"] {- input operations -} read :: Socket -> Int -> IO Bytes read socket n = read_body 0 [] where result = Bytes.concat . reverse read_body len ss = if len >= n then return (result ss) else (do s <- Socket.recv socket (min (n - len) 8192) case ByteString.length s of 0 -> return (result ss) m -> read_body (len + m) (Bytes.make s : ss)) read_block :: Socket -> Int -> IO (Maybe Bytes, Int) read_block socket n = do msg <- read socket n let len = Bytes.length msg return (if len == n then Just msg else Nothing, len) read_line :: Socket -> IO (Maybe Bytes) read_line socket = read_body [] where result = trim_line . Bytes.pack . reverse read_body bs = do s <- Socket.recv socket 1 case ByteString.length s of 0 -> return (if null bs then Nothing else Just (result bs)) 1 -> case ByteString.head s of 10 -> return (Just (result bs)) b -> read_body (b : bs) {- messages with multiple chunks (arbitrary content) -} make_header :: [Int] -> [Bytes] make_header ns = [space_implode "," (map Value.print_int ns), "\n"] make_message :: [Bytes] -> [Bytes] make_message chunks = make_header (map Bytes.length chunks) <> chunks write_message :: Socket -> [Bytes] -> IO () write_message socket = write socket . make_message parse_header :: Bytes -> [Int] parse_header line = let res = map Value.parse_nat (space_explode ',' line) in if all isJust res then map fromJust res else error ("Malformed message header: " <> quote (UTF8.decode line)) read_chunk :: Socket -> Int -> IO Bytes read_chunk socket n = do res <- read_block socket n return $ case res of (Just chunk, _) -> chunk (Nothing, len) -> error ("Malformed message chunk: unexpected EOF after " <> show len <> " of " <> show n <> " bytes") read_message :: Socket -> IO (Maybe [Bytes]) read_message socket = do res <- read_line socket case res of Just line -> Just <$> mapM (read_chunk socket) (parse_header line) Nothing -> return Nothing exchange_message :: Socket -> [Bytes] -> IO (Maybe [Bytes]) exchange_message socket msg = do write_message socket msg read_message socket exchange_message0 :: Socket -> [Bytes] -> IO () exchange_message0 socket msg = do _ <- exchange_message socket msg return () -- hybrid messages: line or length+block (with content restriction) is_length :: Bytes -> Bool is_length msg = not (Bytes.null msg) && Bytes.all_char (\c -> '0' <= c && c <= '9') msg is_terminated :: Bytes -> Bool is_terminated msg = not (Bytes.null msg) && Symbol.is_ascii_line_terminator (Bytes.char $ Bytes.last msg) make_line_message :: Bytes -> [Bytes] make_line_message msg = let n = Bytes.length msg in if is_length msg || is_terminated msg then error ("Bad content for line message:\n" <> take 100 (UTF8.decode msg)) else (if n > 100 || Bytes.any_char (== '\n') msg then make_header [n + 1] else []) <> [msg, "\n"] write_line_message :: Socket -> Bytes -> IO () write_line_message socket = write socket . make_line_message read_line_message :: Socket -> IO (Maybe Bytes) read_line_message socket = do opt_line <- read_line socket case opt_line of Nothing -> return Nothing Just line -> case Value.parse_nat line of Nothing -> return $ Just line Just n -> fmap trim_line . fst <$> read_block socket n read_yxml :: Socket -> IO (Maybe XML.Body) read_yxml socket = do res <- read_line_message socket return (YXML.parse_body <$> res) write_yxml :: Socket -> XML.Body -> IO () write_yxml socket body = write_line_message socket (YXML.string_of_body body) \ generate_file "Isabelle/Isabelle_Thread.hs" = \ {- Title: Isabelle/Isabelle_Thread.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Isabelle-specific thread management. See \<^file>\$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.ML\ and \<^file>\$ISABELLE_HOME/src/Pure/Concurrent/isabelle_thread.scala\. -} {-# LANGUAGE NamedFieldPuns #-} module Isabelle.Isabelle_Thread ( ThreadId, Result, find_id, properties, change_properties, add_resource, del_resource, bracket_resource, is_stopped, expose_stopped, stop, my_uuid, stop_uuid, Fork, fork_finally, fork) where import Data.Unique import Data.IORef import System.IO.Unsafe import qualified Data.List as List import Control.Monad (when, forM_) import Data.Map.Strict (Map) import qualified Data.Map.Strict as Map import Control.Exception as Exception import Control.Concurrent (ThreadId) import qualified Control.Concurrent as Concurrent import Control.Concurrent.Thread (Result) import qualified Control.Concurrent.Thread as Thread import qualified Isabelle.UUID as UUID import qualified Isabelle.Properties as Properties {- thread info -} type Resources = Map Unique (IO ()) data Info = Info {uuid :: UUID.T, props :: Properties.T, stopped :: Bool, resources :: Resources} type Infos = Map ThreadId Info lookup_info :: Infos -> ThreadId -> Maybe Info lookup_info infos id = Map.lookup id infos init_info :: ThreadId -> UUID.T -> Infos -> (Infos, ()) init_info id uuid infos = (Map.insert id (Info uuid [] False Map.empty) infos, ()) {- global state -} {-# NOINLINE global_state #-} global_state :: IORef Infos global_state = unsafePerformIO (newIORef Map.empty) find_id :: UUID.T -> IO (Maybe ThreadId) find_id uuid = do state <- readIORef global_state return $ fst <$> List.find (\(_, Info{uuid = uuid'}) -> uuid == uuid') (Map.assocs state) get_info :: ThreadId -> IO (Maybe Info) get_info id = do state <- readIORef global_state return $ lookup_info state id map_info :: ThreadId -> (Info -> Info) -> IO (Maybe Info) map_info id f = atomicModifyIORef' global_state (\infos -> case lookup_info infos id of Nothing -> (infos, Nothing) Just info -> let info' = f info in (Map.insert id info' infos, Just info')) delete_info :: ThreadId -> IO () delete_info id = atomicModifyIORef' global_state (\infos -> (Map.delete id infos, ())) {- thread properties -} my_info :: IO (Maybe Info) my_info = do id <- Concurrent.myThreadId get_info id properties :: IO Properties.T properties = maybe [] props <$> my_info change_properties :: (Properties.T -> Properties.T) -> IO () change_properties f = do id <- Concurrent.myThreadId map_info id (\info -> info {props = f (props info)}) return () {- managed resources -} add_resource :: IO () -> IO Unique add_resource resource = do id <- Concurrent.myThreadId u <- newUnique map_info id (\info -> info {resources = Map.insert u resource (resources info)}) return u del_resource :: Unique -> IO () del_resource u = do id <- Concurrent.myThreadId map_info id (\info -> info {resources = Map.delete u (resources info)}) return () bracket_resource :: IO () -> IO a -> IO a bracket_resource resource body = Exception.bracket (add_resource resource) del_resource (const body) {- stop -} is_stopped :: IO Bool is_stopped = maybe False stopped <$> my_info expose_stopped :: IO () expose_stopped = do stopped <- is_stopped when stopped $ throw ThreadKilled stop :: ThreadId -> IO () stop id = do info <- map_info id (\info -> info {stopped = True}) let ops = case info of Nothing -> []; Just Info{resources} -> map snd (Map.toDescList resources) sequence_ ops {- UUID -} my_uuid :: IO (Maybe UUID.T) my_uuid = fmap uuid <$> my_info stop_uuid :: UUID.T -> IO () stop_uuid uuid = do id <- find_id uuid forM_ id stop {- fork -} type Fork a = (ThreadId, UUID.T, IO (Result a)) fork_finally :: IO a -> (Either SomeException a -> IO b) -> IO (Fork b) fork_finally body finally = do uuid <- UUID.random (id, result) <- Exception.mask (\restore -> Thread.forkIO (Exception.try (do id <- Concurrent.myThreadId atomicModifyIORef' global_state (init_info id uuid) restore body) >>= (\res -> do id <- Concurrent.myThreadId; delete_info id; finally res))) return (id, uuid, result) fork :: IO a -> IO (Fork a) fork body = fork_finally body Thread.result \ generate_file "Isabelle/Server.hs" = \ {- Title: Isabelle/Server.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) TCP server on localhost. -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Server ( localhost_name, localhost_prefix, localhost, publish_text, publish_stdout, server, connection ) where import Control.Monad (forever, when) import qualified Control.Exception as Exception import Network.Socket (Socket) import qualified Network.Socket as Socket import qualified System.IO as IO import qualified Data.ByteString.Char8 as Char8 import Isabelle.Library import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.UUID as UUID import qualified Isabelle.Byte_Message as Byte_Message import qualified Isabelle.Isabelle_Thread as Isabelle_Thread {- server address -} localhost_name :: Bytes localhost_name = "127.0.0.1" localhost_prefix :: Bytes localhost_prefix = localhost_name <> ":" localhost :: Socket.HostAddress localhost = Socket.tupleToHostAddress (127, 0, 0, 1) publish_text :: Bytes -> Bytes -> UUID.T -> Bytes publish_text name address password = "server " <> quote name <> " = " <> address <> " (password " <> quote (show_bytes password) <> ")" publish_stdout :: Bytes -> Bytes -> UUID.T -> IO () publish_stdout name address password = Char8.putStrLn (Bytes.unmake $ publish_text name address password) {- server -} server :: (Bytes -> UUID.T -> IO ()) -> (Socket -> IO ()) -> IO () server publish handle = Socket.withSocketsDo $ Exception.bracket open (Socket.close . fst) (uncurry loop) where open :: IO (Socket, Bytes) open = do server_socket <- Socket.socket Socket.AF_INET Socket.Stream Socket.defaultProtocol Socket.bind server_socket (Socket.SockAddrInet 0 localhost) Socket.listen server_socket 50 port <- Socket.socketPort server_socket let address = localhost_name <> ":" <> show_bytes port password <- UUID.random publish address password return (server_socket, UUID.print password) loop :: Socket -> Bytes -> IO () loop server_socket password = forever $ do (connection, _) <- Socket.accept server_socket Isabelle_Thread.fork_finally (do line <- Byte_Message.read_line connection when (line == Just password) $ handle connection) (\finally -> do Socket.close connection case finally of Left exn -> IO.hPutStrLn IO.stderr $ Exception.displayException exn Right () -> return ()) return () {- client connection -} connection :: String -> Bytes -> (Socket -> IO a) -> IO a connection port password client = Socket.withSocketsDo $ do addr <- resolve Exception.bracket (open addr) Socket.close body where resolve = do let hints = Socket.defaultHints { Socket.addrFlags = [Socket.AI_NUMERICHOST, Socket.AI_NUMERICSERV], Socket.addrSocketType = Socket.Stream } head <$> Socket.getAddrInfo (Just hints) (Just $ make_string localhost_name) (Just port) open addr = do socket <- Socket.socket (Socket.addrFamily addr) (Socket.addrSocketType addr) (Socket.addrProtocol addr) Socket.connect socket $ Socket.addrAddress addr return socket body socket = do Byte_Message.write_line socket password client socket \ generate_file "Isabelle/Time.hs" = \ {- Title: Isabelle/Time.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Time based on milliseconds. See \<^file>\~~/src/Pure/General/time.scala\ -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Time ( Time, seconds, minutes, ms, zero, is_zero, is_relevant, get_seconds, get_minutes, get_ms, message, now ) where import Text.Printf (printf) import Data.Time.Clock.POSIX (getPOSIXTime) import Isabelle.Bytes (Bytes) import Isabelle.Library newtype Time = Time Int instance Eq Time where Time ms1 == Time ms2 = ms1 == ms2 instance Ord Time where compare (Time ms1) (Time ms2) = compare ms1 ms2 seconds :: Double -> Time seconds s = Time (round (s * 1000.0)) minutes :: Double -> Time minutes m = Time (round (m * 60000.0)) ms :: Int -> Time ms = Time zero :: Time zero = ms 0 is_zero :: Time -> Bool is_zero (Time ms) = ms == 0 is_relevant :: Time -> Bool is_relevant (Time ms) = ms >= 1 get_seconds :: Time -> Double get_seconds (Time ms) = fromIntegral ms / 1000.0 get_minutes :: Time -> Double get_minutes (Time ms) = fromIntegral ms / 60000.0 get_ms :: Time -> Int get_ms (Time ms) = ms instance Show Time where show t = printf "%.3f" (get_seconds t) message :: Time -> Bytes message t = make_bytes (show t) <> "s" now :: IO Time now = do t <- getPOSIXTime return $ Time (round (realToFrac t * 1000.0 :: Double)) \ generate_file "Isabelle/Timing.hs" = \ {- Title: Isabelle/Timing.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Support for time measurement. See \<^file>\~~/src/Pure/General/timing.ML\ and \<^file>\~~/src/Pure/General/timing.scala\ -} module Isabelle.Timing ( Timing (..), zero, is_zero, is_relevant ) where import qualified Isabelle.Time as Time import Isabelle.Time (Time) data Timing = Timing {elapsed :: Time, cpu :: Time, gc :: Time} deriving (Show, Eq) zero :: Timing zero = Timing Time.zero Time.zero Time.zero is_zero :: Timing -> Bool is_zero t = Time.is_zero (elapsed t) && Time.is_zero (cpu t) && Time.is_zero (gc t) is_relevant :: Timing -> Bool is_relevant t = Time.is_relevant (elapsed t) || Time.is_relevant (cpu t) || Time.is_relevant (gc t) \ generate_file "Isabelle/Bash.hs" = \ {- Title: Isabelle/Bash.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Support for GNU bash. See \<^file>\$ISABELLE_HOME/src/Pure/System/bash.ML\ -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Bash ( string, strings, Params, get_script, get_input, get_cwd, get_putenv, get_redirect, get_timeout, get_description, script, input, cwd, putenv, redirect, timeout, description, server_run, server_kill, server_uuid, server_interrupt, server_failure, server_result ) where import Text.Printf (printf) import qualified Isabelle.Symbol as Symbol import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Time as Time import Isabelle.Time (Time) import Isabelle.Library {- concrete syntax -} string :: Bytes -> Bytes string str = if Bytes.null str then "\"\"" else str |> Bytes.unpack |> map trans |> Bytes.concat where trans b = case Bytes.char b of '\t' -> "$'\\t'" '\n' -> "$'\\n'" '\f' -> "$'\\f'" '\r' -> "$'\\r'" c -> if Symbol.is_ascii_letter c || Symbol.is_ascii_digit c || c `elem` ("+,-./:_" :: String) then Bytes.singleton b else if b < 32 || b >= 127 then make_bytes (printf "$'\\x%02x'" b :: String) else "\\" <> Bytes.singleton b strings :: [Bytes] -> Bytes strings = space_implode " " . map string {- server parameters -} data Params = Params { _script :: Bytes, _input :: Bytes, _cwd :: Maybe Bytes, _putenv :: [(Bytes, Bytes)], _redirect :: Bool, _timeout :: Time, _description :: Bytes} deriving (Show, Eq) get_script :: Params -> Bytes get_script = _script get_input :: Params -> Bytes get_input = _input get_cwd :: Params -> Maybe Bytes get_cwd = _cwd get_putenv :: Params -> [(Bytes, Bytes)] get_putenv = _putenv get_redirect :: Params -> Bool get_redirect = _redirect get_timeout :: Params -> Time get_timeout = _timeout get_description :: Params -> Bytes get_description = _description script :: Bytes -> Params script script = Params script "" Nothing [] False Time.zero "" input :: Bytes -> Params -> Params input input params = params { _input = input } cwd :: Bytes -> Params -> Params cwd cwd params = params { _cwd = Just cwd } putenv :: [(Bytes, Bytes)] -> Params -> Params putenv putenv params = params { _putenv = putenv } redirect :: Params -> Params redirect params = params { _redirect = True } timeout :: Time -> Params -> Params timeout timeout params = params { _timeout = timeout } description :: Bytes -> Params -> Params description description params = params { _description = description } {- server messages -} server_run, server_kill :: Bytes server_run = \Bash.server_run\; server_kill = \Bash.server_kill\; server_uuid, server_interrupt, server_failure, server_result :: Bytes server_uuid = \Bash.server_uuid\; server_interrupt = \Bash.server_interrupt\; server_failure = \Bash.server_failure\; server_result = \Bash.server_result\; \ generate_file "Isabelle/Process_Result.hs" = \ {- Title: Isabelle/Process_Result.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Result of system process. See \<^file>\~~/src/Pure/System/process_result.ML\ and \<^file>\~~/src/Pure/System/process_result.scala\ -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Process_Result ( interrupt_rc, timeout_rc, T, make, rc, out_lines, err_lines, timing, timing_elapsed, out, err, ok, check ) where import Isabelle.Time (Time) import qualified Isabelle.Timing as Timing import Isabelle.Timing (Timing) import Isabelle.Bytes (Bytes) import Isabelle.Library interrupt_rc :: Int interrupt_rc = 130 timeout_rc :: Int timeout_rc = 142 data T = Process_Result { _rc :: Int, _out_lines :: [Bytes], _err_lines :: [Bytes], _timing :: Timing} deriving (Show, Eq) make :: Int -> [Bytes] -> [Bytes] -> Timing -> T make = Process_Result rc :: T -> Int rc = _rc out_lines :: T -> [Bytes] out_lines = _out_lines err_lines :: T -> [Bytes] err_lines = _err_lines timing :: T -> Timing timing = _timing timing_elapsed :: T -> Time timing_elapsed = Timing.elapsed . timing out :: T -> Bytes out = trim_line . cat_lines . out_lines err :: T -> Bytes err = trim_line . cat_lines . err_lines ok :: T -> Bool ok result = rc result == 0 check :: T -> T check result = if ok result then result else error (make_string $ err result) \ generate_file "Isabelle/Options.hs" = \ {- Title: Isabelle/Options.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) System options with external string representation. See \<^file>\~~/src/Pure/System/options.ML\ and \<^file>\~~/src/Pure/System/options.scala\ -} {-# LANGUAGE OverloadedStrings #-} {-# LANGUAGE InstanceSigs #-} module Isabelle.Options ( boolT, intT, realT, stringT, unknownT, T, typ, bool, int, real, seconds, string, decode ) where import qualified Data.Map.Strict as Map import Data.Map.Strict (Map) import qualified Isabelle.Properties as Properties import Isabelle.Bytes (Bytes) import qualified Isabelle.Value as Value import qualified Isabelle.Time as Time import Isabelle.Time (Time) import Isabelle.Library import qualified Isabelle.XML.Decode as Decode import Isabelle.XML.Classes (Decode (..)) {- representation -} boolT :: Bytes boolT = "bool" intT :: Bytes intT = "int" realT :: Bytes realT = "real" stringT :: Bytes stringT = "string" unknownT :: Bytes unknownT = "unknown" data Opt = Opt { _pos :: Properties.T, _name :: Bytes, _typ :: Bytes, _value :: Bytes } data T = Options (Map Bytes Opt) {- check -} check_name :: T -> Bytes -> Opt check_name (Options map) name = case Map.lookup name map of Just opt | _typ opt /= unknownT -> opt _ -> error (make_string ("Unknown system option " <> quote name)) check_type :: T -> Bytes -> Bytes -> Opt check_type options name typ = let opt = check_name options name t = _typ opt in if t == typ then opt else error (make_string ("Ill-typed system option " <> quote name <> " : " <> t <> " vs. " <> typ)) {- get typ -} typ :: T -> Bytes -> Bytes typ options name = _typ (check_name options name) {- get value -} get :: Bytes -> (Bytes -> Maybe a) -> T -> Bytes -> a get typ parse options name = let opt = check_type options name typ in case parse (_value opt) of Just x -> x Nothing -> error (make_string ("Malformed value for system option " <> quote name <> " : " <> typ <> " =\n" <> quote (_value opt))) bool :: T -> Bytes -> Bool bool = get boolT Value.parse_bool int :: T -> Bytes -> Int int = get intT Value.parse_int real :: T -> Bytes -> Double real = get realT Value.parse_real seconds :: T -> Bytes -> Time seconds options = Time.seconds . real options string :: T -> Bytes -> Bytes string = get stringT Just {- decode -} instance Decode T where decode :: Decode.T T decode = let decode_entry :: Decode.T (Bytes, Opt) decode_entry body = let (pos, (name, (typ, value))) = Decode.pair Decode.properties (Decode.pair Decode.string (Decode.pair Decode.string Decode.string)) body in (name, Opt { _pos = pos, _name = name, _typ = typ, _value = value }) in Options . Map.fromList . Decode.list decode_entry \ generate_file "Isabelle/Isabelle_System.hs" = \ {- Title: Isabelle/Isabelle_System.hs Author: Makarius LICENSE: BSD 3-clause (Isabelle) Isabelle system support. See \<^file>\~~/src/Pure/System/isabelle_system.ML\ and \<^file>\~~/src/Pure/System/isabelle_system.scala\ -} {-# LANGUAGE OverloadedStrings #-} module Isabelle.Isabelle_System ( bash_process, bash_process0 ) where import Data.Maybe (fromMaybe) import Control.Exception (throw, AsyncException (UserInterrupt)) import Network.Socket (Socket) import qualified Isabelle.Bytes as Bytes import Isabelle.Bytes (Bytes) import qualified Isabelle.Byte_Message as Byte_Message import qualified Isabelle.Time as Time import Isabelle.Timing (Timing (..)) import qualified Isabelle.Options as Options import qualified Isabelle.Bash as Bash import qualified Isabelle.Process_Result as Process_Result import qualified Isabelle.XML.Encode as Encode import qualified Isabelle.YXML as YXML import qualified Isabelle.Value as Value import qualified Isabelle.Server as Server import qualified Isabelle.Isabelle_Thread as Isabelle_Thread import Isabelle.Library {- bash_process -} absolute_path :: Bytes -> Bytes -- FIXME dummy absolute_path = id bash_process :: Options.T -> Bash.Params -> IO Process_Result.T bash_process options = bash_process0 address password where address = Options.string options "bash_process_address" password = Options.string options "bash_process_password" bash_process0 :: Bytes -> Bytes -> Bash.Params -> IO Process_Result.T bash_process0 address password params = do Server.connection port password (\socket -> do isabelle_tmp <- getenv "ISABELLE_TMP" Byte_Message.write_message socket (run isabelle_tmp) loop Nothing socket) where port = case Bytes.try_unprefix Server.localhost_prefix address of Just port -> make_string port Nothing -> errorWithoutStackTrace "Bad bash_process server address" script = Bash.get_script params input = Bash.get_input params cwd = Bash.get_cwd params putenv = Bash.get_putenv params redirect = Bash.get_redirect params timeout = Bash.get_timeout params description = Bash.get_description params run :: Bytes -> [Bytes] run isabelle_tmp = [Bash.server_run, script, input, YXML.string_of_body (Encode.option (Encode.string . absolute_path) cwd), YXML.string_of_body (Encode.list (Encode.pair Encode.string Encode.string) (("ISABELLE_TMP", isabelle_tmp) : putenv)), Value.print_bool redirect, Value.print_real (Time.get_seconds timeout), description] kill :: Maybe Bytes -> IO () kill maybe_uuid = do case maybe_uuid of Just uuid -> Server.connection port password (\socket -> Byte_Message.write_message socket [Bash.server_kill, uuid]) Nothing -> return () err = errorWithoutStackTrace "Malformed result from bash_process server" the = fromMaybe err loop :: Maybe Bytes -> Socket -> IO Process_Result.T loop maybe_uuid socket = do result <- Isabelle_Thread.bracket_resource (kill maybe_uuid) (Byte_Message.read_message socket) case result of Just [head, uuid] | head == Bash.server_uuid -> loop (Just uuid) socket Just [head] | head == Bash.server_interrupt -> throw UserInterrupt Just [head, msg] | head == Bash.server_failure -> errorWithoutStackTrace $ make_string msg Just (head : a : b : c : d : lines) | head == Bash.server_result -> let rc = the $ Value.parse_int a elapsed = Time.ms $ the $ Value.parse_int b cpu = Time.ms $ the $ Value.parse_int c timing = Timing elapsed cpu Time.zero n = the $ Value.parse_int d out_lines = take n lines err_lines = drop n lines in return $ Process_Result.make rc out_lines err_lines timing _ -> err \ export_generated_files _ end diff --git a/src/Tools/IsaPlanner/rw_inst.ML b/src/Tools/IsaPlanner/rw_inst.ML --- a/src/Tools/IsaPlanner/rw_inst.ML +++ b/src/Tools/IsaPlanner/rw_inst.ML @@ -1,243 +1,243 @@ (* Title: Tools/IsaPlanner/rw_inst.ML Author: Lucas Dixon, University of Edinburgh Rewriting using a conditional meta-equality theorem which supports schematic variable instantiation. *) signature RW_INST = sig val rw: Proof.context -> ((indexname * (sort * typ)) list * (* type var instantiations *) (indexname * (typ * term)) list) (* schematic var instantiations *) * (string * typ) list (* Fake named bounds + types *) * (string * typ) list (* names of bound + types *) * term -> (* outer term for instantiation *) thm -> (* rule with indexes lifted *) thm -> (* target thm *) thm (* rewritten theorem possibly with additional premises for rule conditions *) end; structure RW_Inst: RW_INST = struct (* Given (string,type) pairs capturing the free vars that need to be allified in the assumption, and a theorem with assumptions possibly containing the free vars, then we give back the assumptions allified as hidden hyps. Given: x th: A vs ==> B vs Results in: "B vs" [!!x. A x] *) fun allify_conditions ctxt Ts th = let fun allify (x, T) t = Logic.all_const T $ Abs (x, T, Term.abstract_over (Free (x, T), t)); val cTs = map (Thm.cterm_of ctxt o Free) Ts; val cterm_asms = map (Thm.cterm_of ctxt o fold_rev allify Ts) (Thm.prems_of th); val allifyied_asm_thms = map (Drule.forall_elim_list cTs o Thm.assume) cterm_asms; in (fold (curry op COMP) allifyied_asm_thms th, cterm_asms) end; (* Given a list of variables that were bound, and a that has been instantiated with free variable placeholders for the bound vars, it creates an abstracted version of the theorem, with local bound vars as lambda-params: Ts: ("x", ty) rule:: C :x ==> P :x = Q :x results in: ("!! x. C x", (%x. p x = %y. p y) [!! x. C x]) note: assumes rule is instantiated *) (* Note, we take abstraction in the order of last abstraction first *) fun mk_abstractedrule ctxt TsFake Ts rule = let (* now we change the names of temporary free vars that represent bound vars with binders outside the redex *) val ns = IsaND.variant_names ctxt (Thm.full_prop_of rule :: Thm.hyps_of rule) (map fst Ts); val (fromnames, tonames, Ts') = fold (fn (((faken, _), (n, ty)), n2) => fn (rnf, rnt, Ts'') => (Thm.cterm_of ctxt (Free(faken,ty)) :: rnf, Thm.cterm_of ctxt (Free(n2,ty)) :: rnt, (n2,ty) :: Ts'')) (TsFake ~~ Ts ~~ ns) ([], [], []); (* rename conflicting free's in the rule to avoid cconflicts with introduced vars from bounds outside in redex *) val rule' = rule |> Drule.forall_intr_list fromnames |> Drule.forall_elim_list tonames; (* make unconditional rule and prems *) val (uncond_rule, cprems) = allify_conditions ctxt (rev Ts') rule'; (* using these names create lambda-abstracted version of the rule *) val abstractions = rev (Ts' ~~ tonames); val abstract_rule = fold (fn ((n, ty), ct) => Thm.abstract_rule n ct) abstractions uncond_rule; in (cprems, abstract_rule) end; (* given names to avoid, and vars that need to be fixed, it gives unique new names to the vars so that they can be fixed as free variables *) (* make fixed unique free variable instantiations for non-ground vars *) (* Create a table of vars to be renamed after instantiation - ie other uninstantiated vars in the hyps of the rule ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *) fun mk_renamings ctxt tgt rule_inst = let val rule_conds = Thm.prems_of rule_inst; val (_, cond_vs) = fold (fn t => fn (tyvs, vs) => (union (op =) (Misc_Legacy.term_tvars t) tyvs, union (op =) (map Term.dest_Var (Misc_Legacy.term_vars t)) vs)) rule_conds ([], []); val termvars = map Term.dest_Var (Misc_Legacy.term_vars tgt); val vars_to_fix = union (op =) termvars cond_vs; val ys = IsaND.variant_names ctxt (tgt :: rule_conds) (map (fst o fst) vars_to_fix); in map2 (fn (xi, T) => fn y => ((xi, T), Free (y, T))) vars_to_fix ys end; (* make a new fresh typefree instantiation for the given tvar *) fun new_tfree (tv as (ix,sort)) (pairs, used) = let val v = singleton (Name.variant_list used) (string_of_indexname ix) in ((ix,(sort,TFree(v,sort)))::pairs, v::used) end; (* make instantiations to fix type variables that are not already instantiated (in ignore_ixs) from the list of terms. *) fun mk_fixtvar_tyinsts ignore_insts ts = let val ignore_ixs = map fst ignore_insts; val (tvars, tfrees) = fold_rev (fn t => fn (varixs, tfrees) => (Misc_Legacy.add_term_tvars (t,varixs), Misc_Legacy.add_term_tfrees (t,tfrees))) ts ([], []); val unfixed_tvars = filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars; val (fixtyinsts, _) = fold_rev new_tfree unfixed_tvars ([], map fst tfrees) in (fixtyinsts, tfrees) end; (* cross-instantiate the instantiations - ie for each instantiation replace all occurrences in other instantiations - no loops are possible and thus only one-parsing of the instantiations is necessary. *) fun cross_inst insts = let fun instL (ix, (ty,t)) = map (fn (ix2,(ty2,t2)) => (ix2, (ty2,Term.subst_vars ([], [(ix, t)]) t2))); fun cross_instL ([], l) = rev l | cross_instL ((ix, t) :: insts, l) = cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l)); in cross_instL (insts, []) end; (* as above but for types -- I don't know if this is needed, will we ever incur mixed up types? *) fun cross_inst_typs insts = let fun instL (ix, (srt,ty)) = map (fn (ix2,(srt2,ty2)) => (ix2, (srt2,Term.typ_subst_TVars [(ix, ty)] ty2))); fun cross_instL ([], l) = rev l | cross_instL ((ix, t) :: insts, l) = cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l)); in cross_instL (insts, []) end; (* assume that rule and target_thm have distinct var names. THINK: efficient version with tables for vars for: target vars, introduced vars, and rule vars, for quicker instantiation? The outerterm defines which part of the target_thm was modified. Note: we take Ts in the upterm order, ie last abstraction first., and with an outeterm where the abstracted subterm has the arguments in the revered order, ie first abstraction first. FakeTs has abstractions using the fake name - ie the name distinct from all other abstractions. *) fun rw ctxt ((nonfixed_typinsts, unprepinsts), FakeTs, Ts, outerterm) rule target_thm = let (* fix all non-instantiated tvars *) val (fixtyinsts, othertfrees) = (* FIXME proper context!? *) mk_fixtvar_tyinsts nonfixed_typinsts [Thm.prop_of rule, Thm.prop_of target_thm]; val typinsts = cross_inst_typs (nonfixed_typinsts @ fixtyinsts); (* certified instantiations for types *) val ctyp_insts = map (fn (ix, (s, ty)) => ((ix, s), Thm.ctyp_of ctxt ty)) typinsts; (* type instantiated versions *) val tgt_th_tyinst = Thm.instantiate (ctyp_insts,[]) target_thm; val rule_tyinst = Thm.instantiate (ctyp_insts,[]) rule; val term_typ_inst = map (fn (ix,(_,ty)) => (ix,ty)) typinsts; (* type instanitated outer term *) val outerterm_tyinst = Term.subst_TVars term_typ_inst outerterm; val FakeTs_tyinst = map (apsnd (Term.typ_subst_TVars term_typ_inst)) FakeTs; val Ts_tyinst = map (apsnd (Term.typ_subst_TVars term_typ_inst)) Ts; (* type-instantiate the var instantiations *) val insts_tyinst = fold_rev (fn (ix, (ty, t)) => fn insts_tyinst => (ix, (Term.typ_subst_TVars term_typ_inst ty, Term.subst_TVars term_typ_inst t)) :: insts_tyinst) unprepinsts []; (* cross-instantiate *) val insts_tyinst_inst = cross_inst insts_tyinst; (* create certms of instantiations *) val cinsts_tyinst = map (fn (ix, (ty, t)) => ((ix, ty), Thm.cterm_of ctxt t)) insts_tyinst_inst; (* The instantiated rule *) val rule_inst = rule_tyinst |> Thm.instantiate ([], cinsts_tyinst); (* Create a table of vars to be renamed after instantiation - ie other uninstantiated vars in the hyps the *instantiated* rule ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *) val renamings = mk_renamings ctxt (Thm.prop_of tgt_th_tyinst) rule_inst; val cterm_renamings = map (fn (x, y) => apply2 (Thm.cterm_of ctxt) (Var x, y)) renamings; (* Create the specific version of the rule for this target application *) val outerterm_inst = outerterm_tyinst |> Term.subst_Vars (map (fn (ix, (ty, t)) => (ix, t)) insts_tyinst_inst) |> Term.subst_Vars (map (fn ((ix, ty), t) => (ix, t)) renamings); val couter_inst = Thm.reflexive (Thm.cterm_of ctxt outerterm_inst); val (cprems, abstract_rule_inst) = rule_inst |> Thm.instantiate ([], map (apfst (dest_Var o Thm.term_of)) cterm_renamings) |> mk_abstractedrule ctxt FakeTs_tyinst Ts_tyinst; val specific_tgt_rule = Conv.fconv_rule Drule.beta_eta_conversion (Thm.combination couter_inst abstract_rule_inst); (* create an instantiated version of the target thm *) val tgt_th_inst = tgt_th_tyinst |> Thm.instantiate ([], cinsts_tyinst) |> Thm.instantiate ([], map (apfst (dest_Var o Thm.term_of)) cterm_renamings); val (vars,frees_of_fixed_vars) = Library.split_list cterm_renamings; in Conv.fconv_rule Drule.beta_eta_conversion tgt_th_inst |> Thm.equal_elim specific_tgt_rule |> Drule.implies_intr_list cprems |> Drule.forall_intr_list frees_of_fixed_vars |> Drule.forall_elim_list vars - |> Thm.varifyT_global' othertfrees + |> Thm.varifyT_global' (Term_Subst.TFrees.make_set othertfrees) |-> K Drule.zero_var_indexes end; end; diff --git a/src/Tools/misc_legacy.ML b/src/Tools/misc_legacy.ML --- a/src/Tools/misc_legacy.ML +++ b/src/Tools/misc_legacy.ML @@ -1,253 +1,267 @@ (* Title: Tools/misc_legacy.ML Misc legacy stuff -- to be phased out eventually. *) signature MISC_LEGACY = sig + val thm_frees: thm -> cterm list + val cterm_frees: cterm -> cterm list val add_term_names: term * string list -> string list val add_term_tvars: term * (indexname * sort) list -> (indexname * sort) list val add_term_tfrees: term * (string * sort) list -> (string * sort) list val typ_tvars: typ -> (indexname * sort) list val term_tfrees: term -> (string * sort) list val term_tvars: term -> (indexname * sort) list val add_term_vars: term * term list -> term list val term_vars: term -> term list val add_term_frees: term * term list -> term list val term_frees: term -> term list val mk_defpair: term * term -> string * term val get_def: theory -> xstring -> thm val METAHYPS: Proof.context -> (thm list -> tactic) -> int -> tactic val freeze_thaw_robust: Proof.context -> thm -> thm * (int -> thm -> thm) end; structure Misc_Legacy: MISC_LEGACY = struct +fun thm_frees th = + (th, (Term_Subst.Frees.empty, [])) |-> Thm.fold_atomic_cterms {hyps = true} Term.is_Free + (fn ct => fn (set, list) => + let val v = Term.dest_Free (Thm.term_of ct) in + if not (Term_Subst.Frees.defined set v) + then (Term_Subst.Frees.add (v, ()) set, ct :: list) + else (set, list) + end) + |> #2; + +val cterm_frees = thm_frees o Drule.mk_term; + (*iterate a function over all types in a term*) fun it_term_types f = let fun iter(Const(_,T), a) = f(T,a) | iter(Free(_,T), a) = f(T,a) | iter(Var(_,T), a) = f(T,a) | iter(Abs(_,T,t), a) = iter(t,f(T,a)) | iter(f$u, a) = iter(f, iter(u, a)) | iter(Bound _, a) = a in iter end (*Accumulates the names in the term, suppressing duplicates. Includes Frees and Consts. For choosing unambiguous bound var names.*) fun add_term_names (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs | add_term_names (Free(a,_), bs) = insert (op =) a bs | add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs)) | add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs) | add_term_names (_, bs) = bs; (*Accumulates the TVars in a type, suppressing duplicates.*) fun add_typ_tvars(Type(_,Ts),vs) = List.foldr add_typ_tvars vs Ts | add_typ_tvars(TFree(_),vs) = vs | add_typ_tvars(TVar(v),vs) = insert (op =) v vs; (*Accumulates the TFrees in a type, suppressing duplicates.*) fun add_typ_tfree_names(Type(_,Ts),fs) = List.foldr add_typ_tfree_names fs Ts | add_typ_tfree_names(TFree(f,_),fs) = insert (op =) f fs | add_typ_tfree_names(TVar(_),fs) = fs; fun add_typ_tfrees(Type(_,Ts),fs) = List.foldr add_typ_tfrees fs Ts | add_typ_tfrees(TFree(f),fs) = insert (op =) f fs | add_typ_tfrees(TVar(_),fs) = fs; (*Accumulates the TVars in a term, suppressing duplicates.*) val add_term_tvars = it_term_types add_typ_tvars; (*Accumulates the TFrees in a term, suppressing duplicates.*) val add_term_tfrees = it_term_types add_typ_tfrees; val add_term_tfree_names = it_term_types add_typ_tfree_names; (*Non-list versions*) fun typ_tfrees T = add_typ_tfrees(T,[]); fun typ_tvars T = add_typ_tvars(T,[]); fun term_tfrees t = add_term_tfrees(t,[]); fun term_tvars t = add_term_tvars(t,[]); (*Accumulates the Vars in the term, suppressing duplicates.*) fun add_term_vars (t, vars: term list) = case t of Var _ => Ord_List.insert Term_Ord.term_ord t vars | Abs (_,_,body) => add_term_vars(body,vars) | f$t => add_term_vars (f, add_term_vars(t, vars)) | _ => vars; fun term_vars t = add_term_vars(t,[]); (*Accumulates the Frees in the term, suppressing duplicates.*) fun add_term_frees (t, frees: term list) = case t of Free _ => Ord_List.insert Term_Ord.term_ord t frees | Abs (_,_,body) => add_term_frees(body,frees) | f$t => add_term_frees (f, add_term_frees(t, frees)) | _ => frees; fun term_frees t = add_term_frees(t,[]); fun mk_defpair (lhs, rhs) = (case Term.head_of lhs of Const (name, _) => (Thm.def_name (Long_Name.base_name name), Logic.mk_equals (lhs, rhs)) | _ => raise TERM ("Malformed definition: head of lhs not a constant", [lhs, rhs])); fun get_def thy = Thm.axiom thy o Name_Space.intern (Theory.axiom_space thy) o Thm.def_name; (**** METAHYPS -- tactical for using hypotheses as meta-level assumptions METAHYPS (fn prems => tac prems) i converts subgoal i, of the form !!x1...xm. [| A1;...;An] ==> A into a new proof state A==>A, supplying A1,...,An as meta-level assumptions (in "prems"). The parameters x1,...,xm become free variables. If the resulting proof state is [| B1;...;Bk] ==> C (possibly assuming A1,...,An) then it is lifted back into the original context, yielding k subgoals. Replaces unknowns in the context by Frees having the prefix METAHYP_ New unknowns in [| B1;...;Bk] ==> C are lifted over x1,...,xm. DOES NOT HANDLE TYPE UNKNOWNS. NOTE: This version does not observe the proof context, and thus cannot work reliably. See also Subgoal.SUBPROOF and Subgoal.FOCUS for properly localized variants of the same idea. ****) local (*Strips assumptions in goal yielding ( [x1,...,xm], [H1,...,Hn], B ) H1,...,Hn are the hypotheses; x1...xm are variants of the parameters. Main difference from strip_assums concerns parameters: it replaces the bound variables by free variables. *) fun strip_context_aux (params, Hs, Const (\<^const_name>\Pure.imp\, _) $ H $ B) = strip_context_aux (params, H :: Hs, B) | strip_context_aux (params, Hs, Const (\<^const_name>\Pure.all\,_) $ Abs (a, T, t)) = let val (b, u) = Syntax_Trans.variant_abs (a, T, t) in strip_context_aux ((b, T) :: params, Hs, u) end | strip_context_aux (params, Hs, B) = (rev params, rev Hs, B); fun strip_context A = strip_context_aux ([], [], A); (*Left-to-right replacements: ctpairs = [...,(vi,ti),...]. Instantiates distinct free variables by terms of same type.*) fun free_instantiate ctpairs = forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs); fun free_of s ((a, i), T) = Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T) fun mk_inst v = (Var v, free_of "METAHYP1_" v) fun metahyps_split_prem prem = let (*find all vars in the hyps -- should find tvars also!*) val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) [] val insts = map mk_inst hyps_vars (*replace the hyps_vars by Frees*) val prem' = subst_atomic insts prem val (params,hyps,concl) = strip_context prem' in (insts,params,hyps,concl) end; fun metahyps_aux_tac ctxt tacf (prem,gno) state = let val (insts,params,hyps,concl) = metahyps_split_prem prem val maxidx = Thm.maxidx_of state val chyps = map (Thm.cterm_of ctxt) hyps val hypths = map Thm.assume chyps val subprems = map (Thm.forall_elim_vars 0) hypths val fparams = map Free params val cparams = map (Thm.cterm_of ctxt) fparams fun swap_ctpair (t, u) = apply2 (Thm.cterm_of ctxt) (u, t) (*Subgoal variables: make Free; lift type over params*) fun mk_subgoal_inst concl_vars (v, T) = if member (op =) concl_vars (v, T) then ((v, T), true, free_of "METAHYP2_" (v, T)) else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T)) (*Instantiate subgoal vars by Free applied to params*) fun mk_inst (v, in_concl, u) = if in_concl then (v, Thm.cterm_of ctxt u) else (v, Thm.cterm_of ctxt (list_comb (u, fparams))) (*Restore Vars with higher type and index*) fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) = if in_concl then apply2 (Thm.cterm_of ctxt) (u, Var ((a, i), T)) else apply2 (Thm.cterm_of ctxt) (u, Var ((a, i + maxidx), U)) (*Embed B in the original context of params and hyps*) fun embed B = fold_rev Logic.all fparams (Logic.list_implies (hyps, B)) (*Strip the context using elimination rules*) fun elim Bhyp = implies_elim_list (forall_elim_list cparams Bhyp) hypths (*A form of lifting that discharges assumptions.*) fun relift st = let val prop = Thm.prop_of st val subgoal_vars = (*Vars introduced in the subgoals*) fold Term.add_vars (Logic.strip_imp_prems prop) [] and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) [] val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars val st' = Thm.instantiate ([], map mk_inst subgoal_insts) st val emBs = map (Thm.cterm_of ctxt o embed) (Thm.prems_of st') val Cth = implies_elim_list st' (map (elim o Thm.assume) emBs) in (*restore the unknowns to the hypotheses*) free_instantiate (map swap_ctpair insts @ map mk_subgoal_swap_ctpair subgoal_insts) (*discharge assumptions from state in same order*) (implies_intr_list emBs (forall_intr_list cparams (implies_intr_list chyps Cth))) end (*function to replace the current subgoal*) fun next st = Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false} (false, relift st, Thm.nprems_of st) gno state in Seq.maps next (tacf subprems (Thm.trivial (Thm.cterm_of ctxt concl))) end; in fun METAHYPS ctxt tacf n thm = SUBGOAL (metahyps_aux_tac ctxt tacf) n thm handle THM ("assume: variables", _, _) => Seq.empty end; (* generating identifiers -- often fresh *) local (*Maps 0-61 to A-Z, a-z, 0-9; exclude _ or ' to avoid clash with internal/unusual indentifiers*) fun gensym_char i = if i<26 then chr (ord "A" + i) else if i<52 then chr (ord "a" + i - 26) else chr (ord "0" + i - 52); val char_vec = Vector.tabulate (62, gensym_char); fun newid n = implode (map (fn i => Vector.sub (char_vec, i)) (radixpand (62, n))); val gensym_seed = Synchronized.var "gensym_seed" (0: int); in fun gensym pre = Synchronized.change_result gensym_seed (fn i => (pre ^ newid i, i + 1)); end; (*Convert all Vars in a theorem to Frees. Also return a function for reversing that operation. DOES NOT WORK FOR TYPE VARIABLES.*) fun freeze_thaw_robust ctxt th = let val fth = Thm.legacy_freezeT th in - case Thm.fold_terms Term.add_vars fth [] of + case Thm.fold_terms {hyps = false} Term.add_vars fth [] of [] => (fth, fn _ => fn x => x) (*No vars: nothing to do!*) | vars => let fun newName (ix,_) = (ix, gensym (string_of_indexname ix)) val alist = map newName vars fun mk_inst (v,T) = apply2 (Thm.cterm_of ctxt) (Var (v, T), Free (the (AList.lookup (op =) alist v), T)) val insts = map mk_inst vars fun thaw i th' = (*i is non-negative increment for Var indexes*) th' |> forall_intr_list (map #2 insts) |> forall_elim_list (map (Thm.incr_indexes_cterm i o #1) insts) in (Thm.instantiate ([], map (apfst (dest_Var o Thm.term_of)) insts) fth, thaw) end end; end;