diff --git a/src/Doc/Classes/Classes.thy b/src/Doc/Classes/Classes.thy --- a/src/Doc/Classes/Classes.thy +++ b/src/Doc/Classes/Classes.thy @@ -1,631 +1,631 @@ theory Classes imports Main Setup begin section \Introduction\ text \ - Type classes were introduced by Wadler and Blott @{cite wadler89how} + Type classes were introduced by Wadler and Blott \<^cite>\wadler89how\ into the Haskell language to allow for a reasonable implementation of overloading\footnote{throughout this tutorial, we are referring to classical Haskell 1.0 type classes, not considering later additions in expressiveness}. As a canonical example, a polymorphic equality function \eq :: \ \ \ \ bool\ which is overloaded on different types for \\\, which is achieved by splitting introduction of the \eq\ function from its overloaded definitions by means of \class\ and \instance\ declarations: \footnote{syntax here is a kind of isabellized Haskell} \begin{quote} \<^noindent> \class eq where\ \\ \hspace*{2ex}\eq :: \ \ \ \ bool\ \<^medskip>\<^noindent> \instance nat :: eq where\ \\ \hspace*{2ex}\eq 0 0 = True\ \\ \hspace*{2ex}\eq 0 _ = False\ \\ \hspace*{2ex}\eq _ 0 = False\ \\ \hspace*{2ex}\eq (Suc n) (Suc m) = eq n m\ \<^medskip>\<^noindent> \instance (\::eq, \::eq) pair :: eq where\ \\ \hspace*{2ex}\eq (x1, y1) (x2, y2) = eq x1 x2 \ eq y1 y2\ \<^medskip>\<^noindent> \class ord extends eq where\ \\ \hspace*{2ex}\less_eq :: \ \ \ \ bool\ \\ \hspace*{2ex}\less :: \ \ \ \ bool\ \end{quote} \<^noindent> Type variables are annotated with (finitely many) classes; these annotations are assertions that a particular polymorphic type provides definitions for overloaded functions. Indeed, type classes not only allow for simple overloading but form a generic calculus, an instance of order-sorted algebra - @{cite "nipkow-sorts93" and "Nipkow-Prehofer:1993" and "Wenzel:1997:TPHOL"}. + \<^cite>\"nipkow-sorts93" and "Nipkow-Prehofer:1993" and "Wenzel:1997:TPHOL"\. From a software engineering point of view, type classes roughly correspond to interfaces in object-oriented languages like Java; so, it is naturally desirable that type classes do not only provide functions (class parameters) but also state specifications implementations must obey. For example, the \class eq\ above could be given the following specification, demanding that \class eq\ is an equivalence relation obeying reflexivity, symmetry and transitivity: \begin{quote} \<^noindent> \class eq where\ \\ \hspace*{2ex}\eq :: \ \ \ \ bool\ \\ \satisfying\ \\ \hspace*{2ex}\refl: eq x x\ \\ \hspace*{2ex}\sym: eq x y \ eq x y\ \\ \hspace*{2ex}\trans: eq x y \ eq y z \ eq x z\ \end{quote} \<^noindent> From a theoretical point of view, type classes are lightweight modules; Haskell type classes may be emulated by SML - functors @{cite classes_modules}. Isabelle/Isar offers a discipline + functors \<^cite>\classes_modules\. Isabelle/Isar offers a discipline of type classes which brings all those aspects together: \begin{enumerate} \item specifying abstract parameters together with corresponding specifications, \item instantiating those abstract parameters by a particular type \item in connection with a ``less ad-hoc'' approach to overloading, \item with a direct link to the Isabelle module system: - locales @{cite "kammueller-locales"}. + locales \<^cite>\"kammueller-locales"\. \end{enumerate} \<^noindent> Isar type classes also directly support code generation in a Haskell like fashion. Internally, they are mapped to more - primitive Isabelle concepts @{cite "Haftmann-Wenzel:2006:classes"}. + primitive Isabelle concepts \<^cite>\"Haftmann-Wenzel:2006:classes"\. This tutorial demonstrates common elements of structured specifications and abstract reasoning with type classes by the algebraic hierarchy of semigroups, monoids and groups. Our - background theory is that of Isabelle/HOL @{cite "isa-tutorial"}, for + background theory is that of Isabelle/HOL \<^cite>\"isa-tutorial"\, for which some familiarity is assumed. \ section \A simple algebra example \label{sec:example}\ subsection \Class definition\ text \ Depending on an arbitrary type \\\, class \semigroup\ introduces a binary operator \(\)\ that is assumed to be associative: \ class %quote semigroup = fixes mult :: "\ \ \ \ \" (infixl "\" 70) assumes assoc: "(x \ y) \ z = x \ (y \ z)" text \ \<^noindent> This @{command class} specification consists of two parts: the \qn{operational} part names the class parameter (@{element "fixes"}), the \qn{logical} part specifies properties on them (@{element "assumes"}). The local @{element "fixes"} and @{element "assumes"} are lifted to the theory toplevel, yielding the global parameter @{term [source] "mult :: \::semigroup \ \ \ \"} and the global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\x y z :: \::semigroup. (x \ y) \ z = x \ (y \ z)"}. \ subsection \Class instantiation \label{sec:class_inst}\ text \ The concrete type \<^typ>\int\ is made a \<^class>\semigroup\ instance by providing a suitable definition for the class parameter \(\)\ and a proof for the specification of @{fact assoc}. This is accomplished by the @{command instantiation} target: \ instantiation %quote int :: semigroup begin definition %quote mult_int_def: "i \ j = i + (j::int)" instance %quote proof fix i j k :: int have "(i + j) + k = i + (j + k)" by simp then show "(i \ j) \ k = i \ (j \ k)" unfolding mult_int_def . qed end %quote text \ \<^noindent> @{command instantiation} defines class parameters at a particular instance using common specification tools (here, @{command definition}). The concluding @{command instance} opens a proof that the given parameters actually conform to the class specification. Note that the first proof step is the @{method standard} method, which for such instance proofs maps to the @{method intro_classes} method. This reduces an instance judgement to the relevant primitive proof goals; typically it is the first method applied in an instantiation proof. From now on, the type-checker will consider \<^typ>\int\ as a \<^class>\semigroup\ automatically, i.e.\ any general results are immediately available on concrete instances. \<^medskip> Another instance of \<^class>\semigroup\ yields the natural numbers: \ instantiation %quote nat :: semigroup begin primrec %quote mult_nat where "(0::nat) \ n = n" | "Suc m \ n = Suc (m \ n)" instance %quote proof fix m n q :: nat show "m \ n \ q = m \ (n \ q)" by (induct m) auto qed end %quote text \ \<^noindent> Note the occurrence of the name \mult_nat\ in the primrec declaration; by default, the local name of a class operation \f\ to be instantiated on type constructor \\\ is mangled as \f_\\. In case of uncertainty, these names may be inspected using the @{command "print_context"} command. \ subsection \Lifting and parametric types\ text \ Overloaded definitions given at a class instantiation may include recursion over the syntactic structure of types. As a canonical example, we model product semigroups using our simple algebra: \ instantiation %quote prod :: (semigroup, semigroup) semigroup begin definition %quote mult_prod_def: "p\<^sub>1 \ p\<^sub>2 = (fst p\<^sub>1 \ fst p\<^sub>2, snd p\<^sub>1 \ snd p\<^sub>2)" instance %quote proof fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "\::semigroup \ \::semigroup" show "p\<^sub>1 \ p\<^sub>2 \ p\<^sub>3 = p\<^sub>1 \ (p\<^sub>2 \ p\<^sub>3)" unfolding mult_prod_def by (simp add: assoc) qed end %quote text \ \<^noindent> Associativity of product semigroups is established using the definition of \(\)\ on products and the hypothetical associativity of the type components; these hypotheses are legitimate due to the \<^class>\semigroup\ constraints imposed on the type components by the @{command instance} proposition. Indeed, this pattern often occurs with parametric types and type classes. \ subsection \Subclassing\ text \ We define a subclass \monoidl\ (a semigroup with a left-hand neutral) by extending \<^class>\semigroup\ with one additional parameter \neutral\ together with its characteristic property: \ class %quote monoidl = semigroup + fixes neutral :: "\" ("\") assumes neutl: "\ \ x = x" text \ \<^noindent> Again, we prove some instances, by providing suitable parameter definitions and proofs for the additional specifications. Observe that instantiations for types with the same arity may be simultaneous: \ instantiation %quote nat and int :: monoidl begin definition %quote neutral_nat_def: "\ = (0::nat)" definition %quote neutral_int_def: "\ = (0::int)" instance %quote proof fix n :: nat show "\ \ n = n" unfolding neutral_nat_def by simp next fix k :: int show "\ \ k = k" unfolding neutral_int_def mult_int_def by simp qed end %quote instantiation %quote prod :: (monoidl, monoidl) monoidl begin definition %quote neutral_prod_def: "\ = (\, \)" instance %quote proof fix p :: "\::monoidl \ \::monoidl" show "\ \ p = p" unfolding neutral_prod_def mult_prod_def by (simp add: neutl) qed end %quote text \ \<^noindent> Fully-fledged monoids are modelled by another subclass, which does not add new parameters but tightens the specification: \ class %quote monoid = monoidl + assumes neutr: "x \ \ = x" instantiation %quote nat and int :: monoid begin instance %quote proof fix n :: nat show "n \ \ = n" unfolding neutral_nat_def by (induct n) simp_all next fix k :: int show "k \ \ = k" unfolding neutral_int_def mult_int_def by simp qed end %quote instantiation %quote prod :: (monoid, monoid) monoid begin instance %quote proof fix p :: "\::monoid \ \::monoid" show "p \ \ = p" unfolding neutral_prod_def mult_prod_def by (simp add: neutr) qed end %quote text \ \<^noindent> To finish our small algebra example, we add a \group\ class with a corresponding instance: \ class %quote group = monoidl + fixes inverse :: "\ \ \" ("(_\
)" [1000] 999) assumes invl: "x\
\ x = \" instantiation %quote int :: group begin definition %quote inverse_int_def: "i\
= - (i::int)" instance %quote proof fix i :: int have "-i + i = 0" by simp then show "i\
\ i = \" unfolding mult_int_def neutral_int_def inverse_int_def . qed end %quote section \Type classes as locales\ subsection \A look behind the scenes\ text \ The example above gives an impression how Isar type classes work in practice. As stated in the introduction, classes also provide a link to Isar's locale system. Indeed, the logical core of a class is nothing other than a locale: \ class %quote idem = fixes f :: "\ \ \" assumes idem: "f (f x) = f x" text \ \<^noindent> essentially introduces the locale \ (*<*)setup %invisible \Sign.add_path "foo"\ (*>*) locale %quote idem = fixes f :: "\ \ \" assumes idem: "f (f x) = f x" text \\<^noindent> together with corresponding constant(s):\ consts %quote f :: "\ \ \" text \ \<^noindent> The connection to the type system is done by means of a primitive type class \idem\, together with a corresponding interpretation: \ interpretation %quote idem_class: idem "f :: (\::idem) \ \" (*<*)sorry(*>*) text \ \<^noindent> This gives you the full power of the Isabelle module system; conclusions in locale \idem\ are implicitly propagated to class \idem\. \ (*<*)setup %invisible \Sign.parent_path\ (*>*) subsection \Abstract reasoning\ text \ Isabelle locales enable reasoning at a general level, while results are implicitly transferred to all instances. For example, we can now establish the \left_cancel\ lemma for groups, which states that the function \(x \)\ is injective: \ lemma %quote (in group) left_cancel: "x \ y = x \ z \ y = z" proof assume "x \ y = x \ z" then have "x\
\ (x \ y) = x\
\ (x \ z)" by simp then have "(x\
\ x) \ y = (x\
\ x) \ z" using assoc by simp then show "y = z" using neutl and invl by simp next assume "y = z" then show "x \ y = x \ z" by simp qed text \ \<^noindent> Here the \qt{@{keyword "in"} \<^class>\group\} target specification indicates that the result is recorded within that context for later use. This local theorem is also lifted to the global one @{fact "group.left_cancel:"} @{prop [source] "\x y z :: \::group. x \ y = x \ z \ y = z"}. Since type \int\ has been made an instance of \group\ before, we may refer to that fact as well: @{prop [source] "\x y z :: int. x \ y = x \ z \ y = z"}. \ subsection \Derived definitions\ text \ Isabelle locales are targets which support local definitions: \ primrec %quote (in monoid) pow_nat :: "nat \ \ \ \" where "pow_nat 0 x = \" | "pow_nat (Suc n) x = x \ pow_nat n x" text \ \<^noindent> If the locale \group\ is also a class, this local definition is propagated onto a global definition of @{term [source] "pow_nat :: nat \ \::monoid \ \::monoid"} with corresponding theorems @{thm pow_nat.simps [no_vars]}. \<^noindent> As you can see from this example, for local definitions you may use any specification tool which works together with locales, such as Krauss's recursive function package - @{cite krauss2006}. + \<^cite>\krauss2006\. \ subsection \A functor analogy\ text \ We introduced Isar classes by analogy to type classes in functional programming; if we reconsider this in the context of what has been said about type classes and locales, we can drive this analogy further by stating that type classes essentially correspond to functors that have a canonical interpretation as type classes. There is also the possibility of other interpretations. For example, \list\s also form a monoid with \append\ and \<^term>\[]\ as operations, but it seems inappropriate to apply to lists the same operations as for genuinely algebraic types. In such a case, we can simply make a particular interpretation of monoids for lists: \ interpretation %quote list_monoid: monoid append "[]" proof qed auto text \ \<^noindent> This enables us to apply facts on monoids to lists, e.g. @{thm list_monoid.neutl [no_vars]}. When using this interpretation pattern, it may also be appropriate to map derived definitions accordingly: \ primrec %quote replicate :: "nat \ \ list \ \ list" where "replicate 0 _ = []" | "replicate (Suc n) xs = xs @ replicate n xs" interpretation %quote list_monoid: monoid append "[]" rewrites "monoid.pow_nat append [] = replicate" proof - interpret monoid append "[]" .. show "monoid.pow_nat append [] = replicate" proof fix n show "monoid.pow_nat append [] n = replicate n" by (induct n) auto qed qed intro_locales text \ \<^noindent> This pattern is also helpful to reuse abstract specifications on the \emph{same} type. For example, think of a class \preorder\; for type \<^typ>\nat\, there are at least two possible instances: the natural order or the order induced by the divides relation. But only one of these instances can be used for @{command instantiation}; using the locale behind the class \preorder\, it is still possible to utilise the same abstract specification again using @{command interpretation}. \ subsection \Additional subclass relations\ text \ Any \group\ is also a \monoid\; this can be made explicit by claiming an additional subclass relation, together with a proof of the logical difference: \ subclass %quote (in group) monoid proof fix x from invl have "x\
\ x = \" by simp with assoc [symmetric] neutl invl have "x\
\ (x \ \) = x\
\ x" by simp with left_cancel show "x \ \ = x" by simp qed text \ The logical proof is carried out on the locale level. Afterwards it is propagated to the type system, making \group\ an instance of \monoid\ by adding an additional edge to the graph of subclass relations (\figref{fig:subclass}). \begin{figure}[htbp] \begin{center} \small \unitlength 0.6mm \begin{picture}(40,60)(0,0) \put(20,60){\makebox(0,0){\semigroup\}} \put(20,40){\makebox(0,0){\monoidl\}} \put(00,20){\makebox(0,0){\monoid\}} \put(40,00){\makebox(0,0){\group\}} \put(20,55){\vector(0,-1){10}} \put(15,35){\vector(-1,-1){10}} \put(25,35){\vector(1,-3){10}} \end{picture} \hspace{8em} \begin{picture}(40,60)(0,0) \put(20,60){\makebox(0,0){\semigroup\}} \put(20,40){\makebox(0,0){\monoidl\}} \put(00,20){\makebox(0,0){\monoid\}} \put(40,00){\makebox(0,0){\group\}} \put(20,55){\vector(0,-1){10}} \put(15,35){\vector(-1,-1){10}} \put(05,15){\vector(3,-1){30}} \end{picture} \caption{Subclass relationship of monoids and groups: before and after establishing the relationship \group \ monoid\; transitive edges are left out.} \label{fig:subclass} \end{center} \end{figure} For illustration, a derived definition in \group\ using \pow_nat\ \ definition %quote (in group) pow_int :: "int \ \ \ \" where "pow_int k x = (if k >= 0 then pow_nat (nat k) x else (pow_nat (nat (- k)) x)\
)" text \ \<^noindent> yields the global definition of @{term [source] "pow_int :: int \ \::group \ \::group"} with the corresponding theorem @{thm pow_int_def [no_vars]}. \ subsection \A note on syntax\ text \ As a convenience, class context syntax allows references to local class operations and their global counterparts uniformly; type inference resolves ambiguities. For example: \ context %quote semigroup begin term %quote "x \ y" \ \example 1\ term %quote "(x::nat) \ y" \ \example 2\ end %quote term %quote "x \ y" \ \example 3\ text \ \<^noindent> Here in example 1, the term refers to the local class operation \mult [\]\, whereas in example 2 the type constraint enforces the global class operation \mult [nat]\. In the global context in example 3, the reference is to the polymorphic global class operation \mult [?\ :: semigroup]\. \ section \Further issues\ subsection \Type classes and code generation\ text \ Turning back to the first motivation for type classes, namely overloading, it is obvious that overloading stemming from @{command class} statements and @{command instantiation} targets naturally maps to Haskell type classes. The code generator framework - @{cite "isabelle-codegen"} takes this into account. If the target + \<^cite>\"isabelle-codegen"\ takes this into account. If the target language (e.g.~SML) lacks type classes, then they are implemented by an explicit dictionary construction. As example, let's go back to the power function: \ definition %quote example :: int where "example = pow_int 10 (-2)" text \ \<^noindent> This maps to Haskell as follows: \ text %quote \ @{code_stmts example (Haskell)} \ text \ \<^noindent> The code in SML has explicit dictionary passing: \ text %quote \ @{code_stmts example (SML)} \ text \ \<^noindent> In Scala, implicits are used as dictionaries: \ text %quote \ @{code_stmts example (Scala)} \ subsection \Inspecting the type class universe\ text \ To facilitate orientation in complex subclass structures, two diagnostics commands are provided: \begin{description} \item[@{command "print_classes"}] print a list of all classes together with associated operations etc. \item[@{command "class_deps"}] visualizes the subclass relation between all classes as a Hasse diagram. An optional first sort argument constrains the set of classes to all subclasses of this sort, an optional second sort argument to all superclasses of this sort. \end{description} \ end diff --git a/src/Doc/Codegen/Computations.thy b/src/Doc/Codegen/Computations.thy --- a/src/Doc/Codegen/Computations.thy +++ b/src/Doc/Codegen/Computations.thy @@ -1,545 +1,545 @@ theory Computations imports Setup "HOL-Library.Code_Target_Int" begin section \Computations \label{sec:computations}\ subsection \Prelude -- The \code\ antiquotation \label{sec:code_antiq}\ text \ The @{ML_antiquotation_def code} antiquotation allows to include constants from generated code directly into ML system code, as in the following toy example: \ datatype %quote form = T | F | And form form | Or form form (*<*) (*>*) ML %quote \ fun eval_form @{code T} = true | eval_form @{code F} = false | eval_form (@{code And} (p, q)) = eval_form p andalso eval_form q | eval_form (@{code Or} (p, q)) = eval_form p orelse eval_form q; \ text \ \noindent The antiquotation @{ML_antiquotation code} takes the name of a constant as argument; the required code is generated transparently and the corresponding constant names are inserted for the given antiquotations. This technique allows to use pattern matching on constructors stemming from compiled datatypes. Note that the @{ML_antiquotation code} antiquotation may not refer to constants which carry adaptations; here you have to refer to the corresponding adapted code directly. \ subsection \The concept of computations\ text \ Computations embody the simple idea that for each monomorphic Isabelle/HOL term of type \\\ by virtue of code generation there exists an corresponding ML type \T\ and a morphism \\ :: \ \ T\ satisfying \\ (t\<^sub>1 \ t\<^sub>2) = \ t\<^sub>1 \ \ t\<^sub>2\, with \\\ denoting term application. For a given Isabelle/HOL type \\\, parts of \\\ can be implemented by a corresponding ML function \\\<^sub>\ :: term \ T\. How? \<^descr> \Let input be a constant C :: \.\ \\ Then \\\<^sub>\ C = f\<^sub>C\ with \f\<^sub>C\ being the image of \C\ under code generation. \<^descr> \Let input be an application (t\<^sub>1 \ t\<^sub>2) :: \.\ \\ Then \\\<^sub>\ (t\<^sub>1 \ t\<^sub>2) = \\<^sub>\ t\<^sub>1 (\\<^sub>\ t\<^sub>2)\. \noindent Using these trivial properties, each monomorphic constant \C : \<^vec>\\<^sub>n \ \\ yields the following equations: \ text %quote \ \\\<^bsub>(\\<^sub>1 \ \\<^sub>2 \ \ \ \\<^sub>n \ \)\<^esub> C = f\<^sub>C\ \\ \\\<^bsub>(\\<^sub>2 \ \ \ \\<^sub>n \ \)\<^esub> (C \ t\<^sub>1) = f\<^sub>C (\\<^bsub>\\<^sub>1\<^esub> t\<^sub>1)\ \\ \\\ \\ \\\<^bsub>\\<^esub> (C \ t\<^sub>1 \ \ \ t\<^sub>n) = f\<^sub>C (\\<^bsub>\\<^sub>1\<^esub> t\<^sub>1) \ (\\<^bsub>\\<^sub>n\<^esub> t\<^sub>n)\ \ text \ \noindent Hence a computation is characterized as follows: \<^item> Let \input constants\ denote a set of monomorphic constants. \<^item> Let \\\ denote a monomorphic type and \'ml\ be a schematic placeholder for its corresponding type in ML under code generation. \<^item> Then the corresponding computation is an ML function of type \<^ML_type>\Proof.context -> term -> 'ml\ partially implementing the morphism \\ :: \ \ T\ for all \<^emph>\input terms\ consisting only of input constants and applications. \noindent The charming idea is that all required code is automatically generated by the code generator for givens input constants and types; that code is directly inserted into the Isabelle/ML runtime system by means of antiquotations. \ subsection \The \computation\ antiquotation\ text \ The following example illustrates its basic usage: \ ML %quote \ local fun int_of_nat @{code "0 :: nat"} = 0 | int_of_nat (@{code Suc} n) = int_of_nat n + 1; in val comp_nat = @{computation nat terms: "plus :: nat \ nat \ nat" "times :: nat \ nat \ nat" "sum_list :: nat list \ nat" "prod_list :: nat list \ nat" datatypes: nat "nat list"} (fn post => post o HOLogic.mk_nat o int_of_nat o the); end \ text \ \<^item> Antiquotations occurring in the same ML block always refer to the same transparently generated code; particularly, they share the same transparently generated datatype declarations. \<^item> The type of a computation is specified as first argument. \<^item> Input constants are specified the following ways: \<^item> Each term following \terms:\ specifies all constants it contains as input constants. \<^item> Each type following \datatypes:\ specifies all constructors of the corresponding code datatype as input constants. Note that this does not increase expressiveness but succinctness for datatypes with many constructors. Abstract type constructors are skipped silently. \<^item> The code generated by a \computation\ antiquotation takes a functional argument which describes how to conclude the computation. What's the rationale behind this? \<^item> There is no automated way to generate a reconstruction function from the resulting ML type to a Isabelle term -- this is in the responsibility of the implementor. One possible approach for robust term reconstruction is the \code\ antiquotation. \<^item> Both statically specified input constants and dynamically provided input terms are subject to preprocessing. Likewise the result is supposed to be subject to postprocessing; the implementor is expected to take care for this explicitly. \<^item> Computations follow the partiality principle (cf.~\secref{sec:partiality_principle}): failures due to pattern matching or unspecified functions are interpreted as partiality; therefore resulting ML values are optional. Hence the functional argument accepts the following parameters \<^item> A postprocessor function \<^ML_type>\term -> term\. \<^item> The resulting value as optional argument. The functional argument is supposed to compose the final result from these in a suitable manner. \noindent A simple application: \ ML_val %quote \ comp_nat \<^context> \<^term>\sum_list [Suc 0, Suc (Suc 0)] * Suc (Suc 0)\ \ text \ \noindent A single ML block may contain an arbitrary number of computation antiquotations. These share the \<^emph>\same\ set of possible input constants. In other words, it does not matter in which antiquotation constants are specified; in the following example, \<^emph>\all\ constants are specified by the first antiquotation once and for all: \ ML %quote \ local fun int_of_nat @{code "0 :: nat"} = 0 | int_of_nat (@{code Suc} n) = int_of_nat n + 1; in val comp_nat = @{computation nat terms: "plus :: nat \ nat \ nat" "times :: nat \ nat \ nat" "sum_list :: nat list \ nat" "prod_list :: nat list \ nat" "replicate :: nat \ nat \ nat list" datatypes: nat "nat list"} (fn post => post o HOLogic.mk_nat o int_of_nat o the); val comp_nat_list = @{computation "nat list"} (fn post => post o HOLogic.mk_list \<^typ>\nat\ o map (HOLogic.mk_nat o int_of_nat) o the); end \ subsection \Pitfalls when specifying input constants \label{sec:input_constants_pitfalls}\ text \ \<^descr> \Complete type coverage.\ Specified input constants must be \<^emph>\complete\ in the sense that for each required type \\\ there is at least one corresponding input constant which can actually \<^emph>\construct\ a concrete value of type \\\, potentially requiring more types recursively; otherwise the system of equations cannot be generated properly. Hence such incomplete input constants sets are rejected immediately. \<^descr> \Unsuitful right hand sides.\ The generated code for a computation must compile in the strict ML runtime environment. This imposes the technical restriction that each compiled input constant \f\<^sub>C\ on the right hand side of a generated equations must compile without throwing an exception. That rules out pathological examples like @{term [source] "undefined :: nat"} as input constants, as well as abstract constructors (cf. \secref{sec:invariant}). \<^descr> \Preprocessing.\ For consistency, input constants are subject to preprocessing; however, the overall approach requires to operate on constants \C\ and their respective compiled images \f\<^sub>C\.\footnote{Technical restrictions of the implementation enforce this, although those could be lifted in the future.} This is a problem whenever preprocessing maps an input constant to a non-constant. To circumvent these situations, the computation machinery has a dedicated preprocessor which is applied \<^emph>\before\ the regular preprocessing, both to input constants as well as input terms. This can be used to map problematic constants to other ones that are not subject to regular preprocessing. Rewrite rules are added using attribute @{attribute code_computation_unfold}. There should rarely be a need to use this beyond the few default setups in HOL, which deal with literals (see also \secref{sec:literal_input}). \ subsection \Pitfalls concerning input terms\ text \ \<^descr> \No polymorphism.\ Input terms must be monomorphic: compilation to ML requires dedicated choice of monomorphic types. \<^descr> \No abstractions.\ Only constants and applications are admissible as input; abstractions are not possible. In theory, the compilation schema could be extended to cover abstractions also, but this would increase the trusted code base. If abstractions are required, they can always be eliminated by a dedicated preprocessing step, e.g.~using combinatorial logic. \<^descr> \Potential interfering of the preprocessor.\ As described in \secref{sec:input_constants_pitfalls} regular preprocessing can have a disruptive effect for input constants. The same also applies to input terms; however the same measures to circumvent that problem for input constants apply to input terms also. \ subsection \Computations using the \computation_conv\ antiquotation\ text \ Computations are a device to implement fast proof procedures. Then results of a computation are often assumed to be trustable - and thus wrapped in oracles (see @{cite "isabelle-isar-ref"}), + and thus wrapped in oracles (see \<^cite>\"isabelle-isar-ref"\), as in the following example:\footnote{ The technical details how numerals are dealt with are given later in \secref{sec:literal_input}.} \ ML %quote \ local fun raw_dvd (b, ct) = \<^instantiate>\x = ct and y = \if b then \<^cterm>\True\ else \<^cterm>\False\\ in cterm \x \ y\ for x y :: bool\; val (_, dvd_oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\dvd\, raw_dvd))); in val conv_dvd = @{computation_conv bool terms: "Rings.dvd :: int \ int \ bool" "plus :: int \ int \ int" "minus :: int \ int \ int" "times :: int \ int \ int" "0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int" } (K (curry dvd_oracle)) end \ text \ \<^item> Antiquotation @{ML_antiquotation computation_conv} basically yields a conversion of type \<^ML_type>\Proof.context -> cterm -> thm\ - (see further @{cite "isabelle-implementation"}). + (see further \<^cite>\"isabelle-implementation"\). \<^item> The antiquotation expects one functional argument to bridge the \qt{untrusted gap}; this can be done e.g.~using an oracle. Since that oracle will only yield \qt{valid} results in the context of that particular computation, the implementor must make sure that it does not leave the local ML scope; in this example, this is achieved using an explicit \local\ ML block. The very presence of the oracle in the code acknowledges that each computation requires explicit thinking before it can be considered trustworthy! \<^item> Postprocessing just operates as further conversion after normalization. \<^item> Partiality makes the whole conversion fall back to reflexivity. \ (*<*) (*>*) ML_val %quote \ conv_dvd \<^context> \<^cterm>\7 dvd ( 62437867527846782 :: int)\; conv_dvd \<^context> \<^cterm>\7 dvd (-62437867527846783 :: int)\; \ text \ \noindent An oracle is not the only way to construct a valid theorem. A computation result can be used to construct an appropriate certificate: \ lemma %quote conv_div_cert: "(Code_Target_Int.positive r * Code_Target_Int.positive s) div Code_Target_Int.positive s \ (numeral r :: int)" (is "?lhs \ ?rhs") proof (rule eq_reflection) have "?lhs = numeral r * numeral s div numeral s" by simp also have "numeral r * numeral s div numeral s = ?rhs" by (rule nonzero_mult_div_cancel_right) simp finally show "?lhs = ?rhs" . qed lemma %quote positive_mult: "Code_Target_Int.positive r * Code_Target_Int.positive s = Code_Target_Int.positive (r * s)" by simp ML %quote \ local fun integer_of_int (@{code int_of_integer} k) = k val cterm_of_int = Thm.cterm_of \<^context> o HOLogic.mk_numeral o integer_of_int; val divisor = Thm.dest_arg o Thm.dest_arg; val evaluate_simps = map mk_eq @{thms arith_simps positive_mult}; fun evaluate ctxt = Simplifier.rewrite_rule ctxt evaluate_simps; fun revaluate ctxt k ct = @{thm conv_div_cert} |> Thm.instantiate' [] [SOME (cterm_of_int k), SOME (divisor ct)] |> evaluate ctxt; in val conv_div = @{computation_conv int terms: "divide :: int \ int \ int" "0 :: int" "1 :: int" "2 :: int" "3 :: int" } revaluate end \ ML_val %quote \ conv_div \<^context> \<^cterm>\46782454343499999992777742432342242323423425 div (7 :: int)\ \ text \ \noindent The example is intentionally kept simple: only positive integers are covered, only remainder-free divisions are feasible, and the input term is expected to have a particular shape. This exhibits the idea more clearly: the result of the computation is used as a mere hint how to instantiate @{fact conv_div_cert}, from which the required theorem is obtained by performing multiplication using the simplifier; hence that theorem is built of proper inferences, with no oracles involved. \ subsection \Computations using the \computation_check\ antiquotation\ text \ The \computation_check\ antiquotation is convenient if only a positive checking of propositions is desired, because then the result type is fixed (\<^typ>\prop\) and all the technical matter concerning postprocessing and oracles is done in the framework once and for all: \ ML %quote \ val check_nat = @{computation_check terms: Trueprop "less :: nat \ nat \ bool" "plus :: nat \ nat \ nat" "times :: nat \ nat \ nat" "0 :: nat" Suc } \ text \ \noindent The HOL judgement \<^term>\Trueprop\ embeds an expression of type \<^typ>\bool\ into \<^typ>\prop\. \ ML_val %quote \ check_nat \<^context> \<^cprop>\less (Suc (Suc 0)) (Suc (Suc (Suc 0)))\ \ text \ \noindent Note that such computations can only \<^emph>\check\ for \<^typ>\prop\s to hold but not \<^emph>\decide\. \ subsection \Some practical hints\ subsubsection \Inspecting generated code\ text \ The antiquotations for computations attempt to produce meaningful error messages. If these are not sufficient, it might by useful to inspect the generated code, which can be achieved using \ declare %quote [[ML_source_trace]] (*<*) declare %quote [[ML_source_trace = false]] (*>*) subsubsection \Literals as input constants \label{sec:literal_input}\ text \ Literals (characters, numerals) in computations cannot be processed naively: the compilation pattern for computations fails whenever target-language literals are involved; since various common code generator setups (see \secref{sec:common_adaptation}) implement \<^typ>\nat\ and \<^typ>\int\ by target-language literals, this problem manifests whenever numeric types are involved. In practice, this is circumvented with a dedicated preprocessor setup for literals (see also \secref{sec:input_constants_pitfalls}). The following examples illustrate the pattern how to specify input constants when literals are involved, without going into too much detail: \ paragraph \An example for \<^typ>\nat\\ ML %quote \ val check_nat = @{computation_check terms: Trueprop "even :: nat \ bool" "plus :: nat \ nat \ nat" "0 :: nat" Suc "1 :: nat" "2 :: nat" "3 :: nat" } \ ML_val %quote \ check_nat \<^context> \<^cprop>\even (Suc 0 + 1 + 2 + 3 + 4 + 5)\ \ paragraph \An example for \<^typ>\int\\ ML %quote \ val check_int = @{computation_check terms: Trueprop "even :: int \ bool" "plus :: int \ int \ int" "0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int" } \ ML_val %quote \ check_int \<^context> \<^cprop>\even ((0::int) + 1 + 2 + 3 + -1 + -2 + -3)\ \ paragraph \An example for \<^typ>\String.literal\\ definition %quote is_cap_letter :: "String.literal \ bool" where "is_cap_letter s \ (case String.asciis_of_literal s of [] \ False | k # _ \ 65 \ k \ k \ 90)" (*<*) (*>*) ML %quote \ val check_literal = @{computation_check terms: Trueprop is_cap_letter datatypes: bool String.literal } \ ML_val %quote \ check_literal \<^context> \<^cprop>\is_cap_letter (STR ''Q'')\ \ subsubsection \Preprocessing HOL terms into evaluable shape\ text \ When integrating decision procedures developed inside HOL into HOL itself, a common approach is to use a deep embedding where operators etc. are represented by datatypes in Isabelle/HOL. Then it is necessary to turn generic shallowly embedded statements into that particular deep embedding (\qt{reification}). One option is to hardcode using code antiquotations (see \secref{sec:code_antiq}). Another option is to use pre-existing infrastructure in HOL: \<^ML>\Reification.conv\ and \<^ML>\Reification.tac\. A simplistic example: \ datatype %quote form_ord = T | F | Less nat nat | And form_ord form_ord | Or form_ord form_ord | Neg form_ord primrec %quote interp :: "form_ord \ 'a::order list \ bool" where "interp T vs \ True" | "interp F vs \ False" | "interp (Less i j) vs \ vs ! i < vs ! j" | "interp (And f1 f2) vs \ interp f1 vs \ interp f2 vs" | "interp (Or f1 f2) vs \ interp f1 vs \ interp f2 vs" | "interp (Neg f) vs \ \ interp f vs" text \ \noindent The datatype \<^type>\form_ord\ represents formulae whose semantics is given by \<^const>\interp\. Note that values are represented by variable indices (\<^typ>\nat\) whose concrete values are given in list \<^term>\vs\. \ ML %quote (*<*) \\ lemma "thm": fixes x y z :: "'a::order" shows "x < y \ x < z \ interp (And (Less (Suc 0) (Suc (Suc 0))) (Less (Suc 0) 0)) [z, x, y]" ML_prf %quote (*>*) \val thm = Reification.conv \<^context> @{thms interp.simps} \<^cterm>\x < y \ x < z\\ (*<*) by (tactic \ALLGOALS (resolve_tac \<^context> [thm])\) (*>*) text \ \noindent By virtue of @{fact interp.simps}, \<^ML>\Reification.conv\ provides a conversion which, for this concrete example, yields @{thm thm [no_vars]}. Note that the argument to \<^const>\interp\ does not contain any free variables and can thus be evaluated using evaluation. A less meager example can be found in the AFP, session \Regular-Sets\, theory \Regexp_Method\. \ end diff --git a/src/Doc/Codegen/Evaluation.thy b/src/Doc/Codegen/Evaluation.thy --- a/src/Doc/Codegen/Evaluation.thy +++ b/src/Doc/Codegen/Evaluation.thy @@ -1,204 +1,204 @@ theory Evaluation imports Setup begin (*<*) ML \ Isabelle_System.make_directory (File.tmp_path (Path.basic "examples")) \ (*>*) section \Evaluation \label{sec:evaluation}\ text \ Recalling \secref{sec:principle}, code generation turns a system of equations into a program with the \emph{same} equational semantics. As a consequence, this program can be used as a \emph{rewrite engine} for terms: rewriting a term \<^term>\t\ using a program to a term \<^term>\t'\ yields the theorems \<^prop>\t \ t'\. This application of code generation in the following is referred to as \emph{evaluation}. \ subsection \Evaluation techniques\ text \ There is a rich palette of evaluation techniques, each comprising different aspects: \begin{description} \item[Expressiveness.] Depending on the extent to which symbolic computation is possible, the class of terms which can be evaluated can be bigger or smaller. \item[Efficiency.] The more machine-near the technique, the faster it is. \item[Trustability.] Techniques which a huge (and also probably more configurable infrastructure) are more fragile and less trustable. \end{description} \ subsubsection \The simplifier (\simp\)\ text \ The simplest way for evaluation is just using the simplifier with the original code equations of the underlying program. This gives fully symbolic evaluation and highest trustablity, with the usual performance of the simplifier. Note that for operations on abstract datatypes (cf.~\secref{sec:invariant}), the original theorems as given by the users are used, not the modified ones. \ subsubsection \Normalization by evaluation (\nbe\)\ text \ - Normalization by evaluation @{cite "Aehlig-Haftmann-Nipkow:2008:nbe"} + Normalization by evaluation \<^cite>\"Aehlig-Haftmann-Nipkow:2008:nbe"\ provides a comparably fast partially symbolic evaluation which permits also normalization of functions and uninterpreted symbols; the stack of code to be trusted is considerable. \ subsubsection \Evaluation in ML (\code\)\ text \ Considerable performance can be achieved using evaluation in ML, at the cost of being restricted to ground results and a layered stack of code to be trusted, including a user's specific code generator setup. Evaluation is carried out in a target language \emph{Eval} which inherits from \emph{SML} but for convenience uses parts of the Isabelle runtime environment. Hence soundness depends crucially on the correctness of the code generator setup; this is one of the reasons why you should not use adaptation (see \secref{sec:adaptation}) frivolously. \ subsection \Dynamic evaluation\ text \ Dynamic evaluation takes the code generator configuration \qt{as it is} at the point where evaluation is issued and computes a corresponding result. Best example is the @{command_def value} command for ad-hoc evaluation of terms: \ value %quote "42 / (12 :: rat)" text \ \noindent @{command value} tries first to evaluate using ML, falling back to normalization by evaluation if this fails. A particular technique may be specified in square brackets, e.g. \ value %quote [nbe] "42 / (12 :: rat)" text \ To employ dynamic evaluation in documents, there is also a \value\ antiquotation with the same evaluation techniques as @{command value}. \ subsubsection \Term reconstruction in ML\ text \ Results from evaluation in ML must be turned into Isabelle's internal term representation again. Since that setup is highly configurable, it is never assumed to be trustable. Hence evaluation in ML provides no full term reconstruction but distinguishes the following kinds: \begin{description} \item[Plain evaluation.] A term is normalized using the vanilla term reconstruction from ML to Isabelle; this is a pragmatic approach for applications which do not need trustability. \item[Property conversion.] Evaluates propositions; since these are monomorphic, the term reconstruction is fixed once and for all and therefore trustable -- in the sense that only the regular code generator setup has to be trusted, without relying on term reconstruction from ML to Isabelle. \end{description} \noindent The different degree of trustability is also manifest in the types of the corresponding ML functions: plain evaluation operates on uncertified terms, whereas property conversion operates on certified terms. \ subsubsection \The partiality principle \label{sec:partiality_principle}\ text \ During evaluation exceptions indicating a pattern match failure or a non-implemented function are treated specially: as sketched in \secref{sec:partiality}, such exceptions can be interpreted as partiality. For plain evaluation, the result hence is optional; property conversion falls back to reflexivity in such cases. \ subsubsection \Schematic overview\ text \ \newcommand{\ttsize}{\fontsize{5.8pt}{8pt}\selectfont} \fontsize{9pt}{12pt}\selectfont \begin{tabular}{l||c|c|c} & \simp\ & \nbe\ & \code\ \tabularnewline \hline \hline interactive evaluation & @{command value} \[simp]\ & @{command value} \[nbe]\ & @{command value} \[code]\ \tabularnewline plain evaluation & & & \ttsize\<^ML>\Code_Evaluation.dynamic_value\ \tabularnewline \hline evaluation method & @{method code_simp} & @{method normalization} & @{method eval} \tabularnewline property conversion & & & \ttsize\<^ML>\Code_Runtime.dynamic_holds_conv\ \tabularnewline \hline conversion & \ttsize\<^ML>\Code_Simp.dynamic_conv\ & \ttsize\<^ML>\Nbe.dynamic_conv\ \end{tabular} \ subsection \Static evaluation\ text \ When implementing proof procedures using evaluation, in most cases the code generator setup is appropriate \qt{as it was} when the proof procedure was written in ML, not an arbitrary later potentially modified setup. This is called static evaluation. \ subsubsection \Static evaluation using \simp\ and \nbe\\ text \ For \simp\ and \nbe\ static evaluation can be achieved using \<^ML>\Code_Simp.static_conv\ and \<^ML>\Nbe.static_conv\. Note that \<^ML>\Nbe.static_conv\ by its very nature requires an invocation of the ML compiler for every call, which can produce significant overhead. \ subsubsection \Intimate connection between logic and system runtime\ text \ Static evaluation for \eval\ operates differently -- the required generated code is inserted directly into an ML block using antiquotations. The idea is that generated code performing static evaluation (called a \<^emph>\computation\) is compiled once and for all such that later calls do not require any invocation of the code generator or the ML compiler at all. This topic deserves a dedicated chapter \secref{sec:computations}. \ end diff --git a/src/Doc/Codegen/Further.thy b/src/Doc/Codegen/Further.thy --- a/src/Doc/Codegen/Further.thy +++ b/src/Doc/Codegen/Further.thy @@ -1,220 +1,220 @@ theory Further imports Setup begin section \Further issues \label{sec:further}\ subsection \Incorporating generated code directly into the system runtime -- \code_reflect\\ subsubsection \Static embedding of generated code into the system runtime\ text \ The @{ML_antiquotation code} antiquotation is lightweight, but the generated code is only accessible while the ML section is processed. Sometimes this is not appropriate, especially if the generated code contains datatype declarations which are shared with other parts of the system. In these cases, @{command_def code_reflect} can be used: \ code_reflect %quote Sum_Type datatypes sum = Inl | Inr functions "Sum_Type.sum.projl" "Sum_Type.sum.projr" text \ \noindent @{command code_reflect} takes a structure name and references to datatypes and functions; for these code is compiled into the named ML structure and the \emph{Eval} target is modified in a way that future code generation will reference these precompiled versions of the given datatypes and functions. This also allows to refer to the referenced datatypes and functions from arbitrary ML code as well. A typical example for @{command code_reflect} can be found in the \<^theory>\HOL.Predicate\ theory. \ subsubsection \Separate compilation\ text \ For technical reasons it is sometimes necessary to separate generation and compilation of code which is supposed to be used in the system runtime. For this @{command code_reflect} with an optional \<^theory_text>\file_prefix\ argument can be used: \ code_reflect %quote Rat datatypes rat functions Fract "(plus :: rat \ rat \ rat)" "(minus :: rat \ rat \ rat)" "(times :: rat \ rat \ rat)" "(divide :: rat \ rat \ rat)" file_prefix rat text \ \noindent This generates the referenced code as logical files within the theory context, similar to @{command export_code}. \ subsection \Specialities of the \Scala\ target language \label{sec:scala}\ text \ \Scala\ deviates from languages of the ML family in a couple of aspects; those which affect code generation mainly have to do with \Scala\'s type system: \begin{itemize} \item \Scala\ prefers tupled syntax over curried syntax. \item \Scala\ sacrifices Hindely-Milner type inference for a much more rich type system with subtyping etc. For this reason type arguments sometimes have to be given explicitly in square brackets (mimicking System F syntax). \item In contrast to \Haskell\ where most specialities of the type system are implemented using \emph{type classes}, \Scala\ provides a sophisticated system of \emph{implicit arguments}. \end{itemize} \noindent Concerning currying, the \Scala\ serializer counts arguments in code equations to determine how many arguments shall be tupled; remaining arguments and abstractions in terms rather than function definitions are always curried. The second aspect affects user-defined adaptations with @{command code_printing}. For regular terms, the \Scala\ serializer prints all type arguments explicitly. For user-defined term adaptations this is only possible for adaptations which take no arguments: here the type arguments are just appended. Otherwise they are ignored; hence user-defined adaptations for polymorphic constants have to be designed very carefully to avoid ambiguity. Note also that name clashes can occur when generating Scala code to case-insensitive file systems; these can be avoid using the ``\(case_insensitive)\'' option for @{command export_code}. \ subsection \Modules namespace\ text \ When invoking the @{command export_code} command it is possible to leave out the @{keyword "module_name"} part; then code is distributed over different modules, where the module name space roughly is induced by the Isabelle theory name space. Then sometimes the awkward situation occurs that dependencies between definitions introduce cyclic dependencies between modules, which in the \Haskell\ world leaves you to the mercy of the \Haskell\ implementation you are using, while for \SML\/\OCaml\ code generation is not possible. A solution is to declare module names explicitly. Let use assume the three cyclically dependent modules are named \emph{A}, \emph{B} and \emph{C}. Then, by stating \ code_identifier %quote code_module A \ (SML) ABC | code_module B \ (SML) ABC | code_module C \ (SML) ABC text \ \noindent we explicitly map all those modules on \emph{ABC}, resulting in an ad-hoc merge of this three modules at serialisation time. \ subsection \Locales and interpretation\ text \ A technical issue comes to surface when generating code from specifications stemming from locale interpretation into global theories. Let us assume a locale specifying a power operation on arbitrary types: \ locale %quote power = fixes power :: "'a \ 'b \ 'b" assumes power_commute: "power x \ power y = power y \ power x" begin text \ \noindent Inside that locale we can lift \power\ to exponent lists by means of a specification relative to that locale: \ primrec %quote powers :: "'a list \ 'b \ 'b" where "powers [] = id" | "powers (x # xs) = power x \ powers xs" lemma %quote powers_append: "powers (xs @ ys) = powers xs \ powers ys" by (induct xs) simp_all lemma %quote powers_power: "powers xs \ power x = power x \ powers xs" by (induct xs) (simp_all del: o_apply id_apply add: comp_assoc, simp del: o_apply add: o_assoc power_commute) lemma %quote powers_rev: "powers (rev xs) = powers xs" by (induct xs) (simp_all add: powers_append powers_power) end %quote text \ After an interpretation of this locale (say, @{command_def global_interpretation} \fun_power:\ @{term [source] "power (\n (f :: 'a \ 'a). f ^^ n)"}), one could naively expect to have a constant \fun_power.powers :: nat list \ ('a \ 'a) \ 'a \ 'a\ for which code can be generated. But this not the case: internally, the term \fun_power.powers\ is an abbreviation for the foundational term @{term [source] "power.powers (\n (f :: 'a \ 'a). f ^^ n)"} - (see @{cite "isabelle-locale"} for the details behind). + (see \<^cite>\"isabelle-locale"\ for the details behind). Fortunately, a succint solution is available: a dedicated rewrite definition: \ global_interpretation %quote fun_power: power "(\n (f :: 'a \ 'a). f ^^ n)" defines funpows = fun_power.powers by unfold_locales (simp_all add: fun_eq_iff funpow_mult mult.commute) text \ \noindent This amends the interpretation morphisms such that occurrences of the foundational term @{term [source] "power.powers (\n (f :: 'a \ 'a). f ^^ n)"} are folded to a newly defined constant \<^const>\funpows\. After this setup procedure, code generation can continue as usual: \ text %quote \ @{code_stmts funpows constant: Nat.funpow funpows (Haskell)} \ subsection \Parallel computation\ text \ Theory \Parallel\ in \<^dir>\~~/src/HOL/Library\ contains operations to exploit parallelism inside the Isabelle/ML runtime engine. \ subsection \Imperative data structures\ text \ If you consider imperative data structures as inevitable for a specific application, you should consider \emph{Imperative Functional Programming with Isabelle/HOL} - @{cite "bulwahn-et-al:2008:imperative"}; the framework described there + \<^cite>\"bulwahn-et-al:2008:imperative"\; the framework described there is available in session \Imperative_HOL\, together with a short primer document. \ end diff --git a/src/Doc/Codegen/Inductive_Predicate.thy b/src/Doc/Codegen/Inductive_Predicate.thy --- a/src/Doc/Codegen/Inductive_Predicate.thy +++ b/src/Doc/Codegen/Inductive_Predicate.thy @@ -1,268 +1,268 @@ theory Inductive_Predicate imports Setup begin (*<*) hide_const %invisible append inductive %invisible append where "append [] ys ys" | "append xs ys zs \ append (x # xs) ys (x # zs)" lemma %invisible append: "append xs ys zs = (xs @ ys = zs)" by (induct xs arbitrary: ys zs) (auto elim: append.cases intro: append.intros) lemmas lexordp_def = lexordp_def [unfolded lexord_def mem_Collect_eq split] (*>*) section \Inductive Predicates \label{sec:inductive}\ text \ The \predicate compiler\ is an extension of the code generator which turns inductive specifications into equational ones, from which in turn executable code can be generated. The mechanisms of this compiler are described in detail in - @{cite "Berghofer-Bulwahn-Haftmann:2009:TPHOL"}. + \<^cite>\"Berghofer-Bulwahn-Haftmann:2009:TPHOL"\. Consider the simple predicate \<^const>\append\ given by these two introduction rules: \ text %quote \ @{thm append.intros(1)[of ys]} \\ @{thm append.intros(2)[of xs ys zs x]} \ text \ \noindent To invoke the compiler, simply use @{command_def "code_pred"}: \ code_pred %quote append . text \ \noindent The @{command "code_pred"} command takes the name of the inductive predicate and then you put a period to discharge a trivial correctness proof. The compiler infers possible modes for the predicate and produces the derived code equations. Modes annotate which (parts of the) arguments are to be taken as input, and which output. Modes are similar to types, but use the notation \i\ for input and \o\ for output. For \<^term>\append\, the compiler can infer the following modes: \begin{itemize} \item \i \ i \ i \ bool\ \item \i \ i \ o \ bool\ \item \o \ o \ i \ bool\ \end{itemize} You can compute sets of predicates using @{command_def "values"}: \ values %quote "{zs. append [(1::nat),2,3] [4,5] zs}" text \\noindent outputs \{[1, 2, 3, 4, 5]}\, and\ values %quote "{(xs, ys). append xs ys [(2::nat),3]}" text \\noindent outputs \{([], [2, 3]), ([2], [3]), ([2, 3], [])}\.\ text \ \noindent If you are only interested in the first elements of the set comprehension (with respect to a depth-first search on the introduction rules), you can pass an argument to @{command "values"} to specify the number of elements you want: \ values %quote 1 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}" values %quote 3 "{(xs, ys). append xs ys [(1::nat), 2, 3, 4]}" text \ \noindent The @{command "values"} command can only compute set comprehensions for which a mode has been inferred. The code equations for a predicate are made available as theorems with the suffix \equation\, and can be inspected with: \ thm %quote append.equation text \ \noindent More advanced options are described in the following subsections. \ subsection \Alternative names for functions\ text \ By default, the functions generated from a predicate are named after the predicate with the mode mangled into the name (e.g., \append_i_i_o\). You can specify your own names as follows: \ code_pred %quote (modes: i \ i \ o \ bool as concat, o \ o \ i \ bool as split, i \ o \ i \ bool as suffix) append . subsection \Alternative introduction rules\ text \ Sometimes the introduction rules of an predicate are not executable because they contain non-executable constants or specific modes could not be inferred. It is also possible that the introduction rules yield a function that loops forever due to the execution in a depth-first search manner. Therefore, you can declare alternative introduction rules for predicates with the attribute @{attribute "code_pred_intro"}. For example, the transitive closure is defined by: \ text %quote \ @{lemma [source] "r a b \ tranclp r a b" by (fact tranclp.intros(1))}\\ @{lemma [source] "tranclp r a b \ r b c \ tranclp r a c" by (fact tranclp.intros(2))} \ text \ \noindent These rules do not suit well for executing the transitive closure with the mode \(i \ o \ bool) \ i \ o \ bool\, as the second rule will cause an infinite loop in the recursive call. This can be avoided using the following alternative rules which are declared to the predicate compiler by the attribute @{attribute "code_pred_intro"}: \ lemma %quote [code_pred_intro]: "r a b \ tranclp r a b" "r a b \ tranclp r b c \ tranclp r a c" by auto text \ \noindent After declaring all alternative rules for the transitive closure, you invoke @{command "code_pred"} as usual. As you have declared alternative rules for the predicate, you are urged to prove that these introduction rules are complete, i.e., that you can derive an elimination rule for the alternative rules: \ code_pred %quote tranclp proof - case tranclp from this converse_tranclpE [OF tranclp.prems] show thesis by metis qed text \ \noindent Alternative rules can also be used for constants that have not been defined inductively. For example, the lexicographic order which is defined as: \ text %quote \ @{thm [display] lexordp_def [of r]} \ text \ \noindent To make it executable, you can derive the following two rules and prove the elimination rule: \ lemma %quote [code_pred_intro]: "append xs (a # v) ys \ lexordp r xs ys" (*<*)unfolding lexordp_def by (auto simp add: append)(*>*) lemma %quote [code_pred_intro]: "append u (a # v) xs \ append u (b # w) ys \ r a b \ lexordp r xs ys" (*<*)unfolding lexordp_def append apply simp apply (rule disjI2) by auto(*>*) code_pred %quote lexordp (*<*)proof - fix r xs ys assume lexord: "lexordp r xs ys" assume 1: "\r' xs' ys' a v. r = r' \ xs = xs' \ ys = ys' \ append xs' (a # v) ys' \ thesis" assume 2: "\r' xs' ys' u a v b w. r = r' \ xs = xs' \ ys = ys' \ append u (a # v) xs' \ append u (b # w) ys' \ r' a b \ thesis" { assume "\a v. ys = xs @ a # v" from this 1 have thesis by (fastforce simp add: append) } moreover { assume "\u a b v w. r a b \ xs = u @ a # v \ ys = u @ b # w" from this 2 have thesis by (fastforce simp add: append) } moreover note lexord ultimately show thesis unfolding lexordp_def by fastforce qed(*>*) subsection \Options for values\ text \ In the presence of higher-order predicates, multiple modes for some predicate could be inferred that are not disambiguated by the pattern of the set comprehension. To disambiguate the modes for the arguments of a predicate, you can state the modes explicitly in the @{command "values"} command. Consider the simple predicate \<^term>\succ\: \ inductive %quote succ :: "nat \ nat \ bool" where "succ 0 (Suc 0)" | "succ x y \ succ (Suc x) (Suc y)" code_pred %quote succ . text \ \noindent For this, the predicate compiler can infer modes \o \ o \ bool\, \i \ o \ bool\, \o \ i \ bool\ and \i \ i \ bool\. The invocation of @{command "values"} \{n. tranclp succ 10 n}\ loops, as multiple modes for the predicate \succ\ are possible and here the first mode \o \ o \ bool\ is chosen. To choose another mode for the argument, you can declare the mode for the argument between the @{command "values"} and the number of elements. \ values %quote [mode: i \ o \ bool] 1 "{n. tranclp succ 10 n}" (*FIXME does not terminate for n\1*) values %quote [mode: o \ i \ bool] 1 "{n. tranclp succ n 10}" subsection \Embedding into functional code within Isabelle/HOL\ text \ To embed the computation of an inductive predicate into functions that are defined in Isabelle/HOL, you have a number of options: \begin{itemize} \item You want to use the first-order predicate with the mode where all arguments are input. Then you can use the predicate directly, e.g. @{text [display] \valid_suffix ys zs = (if append [Suc 0, 2] ys zs then Some ys else None)\} \item If you know that the execution returns only one value (it is deterministic), then you can use the combinator \<^term>\Predicate.the\, e.g., a functional concatenation of lists is defined with @{term [display] "functional_concat xs ys = Predicate.the (append_i_i_o xs ys)"} Note that if the evaluation does not return a unique value, it raises a run-time error \<^term>\not_unique\. \end{itemize} \ subsection \Further Examples\ text \ Further examples for compiling inductive predicates can be found in \<^file>\~~/src/HOL/Predicate_Compile_Examples/Examples.thy\. There are also some examples in the Archive of Formal Proofs, notably in the \POPLmark-deBruijn\ and the \FeatherweightJava\ sessions. \ end diff --git a/src/Doc/Codegen/Introduction.thy b/src/Doc/Codegen/Introduction.thy --- a/src/Doc/Codegen/Introduction.thy +++ b/src/Doc/Codegen/Introduction.thy @@ -1,256 +1,256 @@ theory Introduction imports Setup begin (*<*) ML \ Isabelle_System.make_directory (File.tmp_path (Path.basic "examples")) \ (*>*) section \Introduction\ text \ This tutorial introduces the code generator facilities of \Isabelle/HOL\. It allows to turn (a certain class of) HOL specifications into corresponding executable code in the programming - languages \SML\ @{cite SML}, \OCaml\ @{cite OCaml}, - \Haskell\ @{cite "haskell-revised-report"} and \Scala\ - @{cite "scala-overview-tech-report"}. + languages \SML\ \<^cite>\SML\, \OCaml\ \<^cite>\OCaml\, + \Haskell\ \<^cite>\"haskell-revised-report"\ and \Scala\ + \<^cite>\"scala-overview-tech-report"\. To profit from this tutorial, some familiarity and experience with - Isabelle/HOL @{cite "isa-tutorial"} and its basic theories is assumed. + Isabelle/HOL \<^cite>\"isa-tutorial"\ and its basic theories is assumed. \ subsection \Code generation principle: shallow embedding \label{sec:principle}\ text \ The key concept for understanding Isabelle's code generation is \emph{shallow embedding}: logical entities like constants, types and classes are identified with corresponding entities in the target language. In particular, the carrier of a generated program's semantics are \emph{equational theorems} from the logic. If we view a generated program as an implementation of a higher-order rewrite system, then every rewrite step performed by the program can be simulated in the logic, which guarantees partial correctness - @{cite "Haftmann-Nipkow:2010:code"}. + \<^cite>\"Haftmann-Nipkow:2010:code"\. \ subsection \A quick start with the Isabelle/HOL toolbox \label{sec:queue_example}\ text \ In a HOL theory, the @{command_def datatype} and @{command_def definition}/@{command_def primrec}/@{command_def fun} declarations form the core of a functional programming language. By default equational theorems stemming from those are used for generated code, therefore \qt{naive} code generation can proceed without further ado. For example, here a simple \qt{implementation} of amortised queues: \ datatype %quote 'a queue = AQueue "'a list" "'a list" definition %quote empty :: "'a queue" where "empty = AQueue [] []" primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where "dequeue (AQueue [] []) = (None, AQueue [] [])" | "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" | "dequeue (AQueue xs []) = (case rev xs of y # ys \ (Some y, AQueue [] ys))" (*<*) lemma %invisible dequeue_nonempty_Nil [simp]: "xs \ [] \ dequeue (AQueue xs []) = (case rev xs of y # ys \ (Some y, AQueue [] ys))" by (cases xs) (simp_all split: list.splits) (*>*) text \\noindent Then we can generate code e.g.~for \SML\ as follows:\ export_code %quote empty dequeue enqueue in SML module_name Example text \\noindent resulting in the following code:\ text %quote \ @{code_stmts empty enqueue dequeue (SML)} \ text \ \noindent The @{command_def export_code} command takes multiple constants for which code shall be generated; anything else needed for those is added implicitly. Then follows a target language identifier and a freely chosen \<^theory_text>\module_name\. Output is written to a logical file-system within the theory context, with the theory name and ``\<^verbatim>\code\'' as overall prefix. There is also a formal session export using the same name: the result may be explored in the Isabelle/jEdit Prover IDE using the file-browser on the URL ``\<^verbatim>\isabelle-export:\''. The file name is determined by the target language together with an optional \<^theory_text>\file_prefix\ (the default is ``\<^verbatim>\export\'' with a consecutive number within the current theory). For \SML\, \OCaml\ and \Scala\, the file prefix becomes a plain file with extension (e.g.\ ``\<^verbatim>\.ML\'' for SML). For \Haskell\ the file prefix becomes a directory that is populated with a separate file for each module (with extension ``\<^verbatim>\.hs\''). Consider the following example: \ export_code %quote empty dequeue enqueue in Haskell module_name Example file_prefix example text \ \noindent This is the corresponding code: \ text %quote \ @{code_stmts empty enqueue dequeue (Haskell)} \ text \ \noindent For more details about @{command export_code} see \secref{sec:further}. \ subsection \Type classes\ text \ Code can also be generated from type classes in a Haskell-like manner. For illustration here an example from abstract algebra: \ class %quote semigroup = fixes mult :: "'a \ 'a \ 'a" (infixl "\" 70) assumes assoc: "(x \ y) \ z = x \ (y \ z)" class %quote monoid = semigroup + fixes neutral :: 'a ("\") assumes neutl: "\ \ x = x" and neutr: "x \ \ = x" instantiation %quote nat :: monoid begin primrec %quote mult_nat where "0 \ n = (0::nat)" | "Suc m \ n = n + m \ n" definition %quote neutral_nat where "\ = Suc 0" lemma %quote add_mult_distrib: fixes n m q :: nat shows "(n + m) \ q = n \ q + m \ q" by (induct n) simp_all instance %quote proof fix m n q :: nat show "m \ n \ q = m \ (n \ q)" by (induct m) (simp_all add: add_mult_distrib) show "\ \ n = n" by (simp add: neutral_nat_def) show "m \ \ = m" by (induct m) (simp_all add: neutral_nat_def) qed end %quote text \ \noindent We define the natural operation of the natural numbers on monoids: \ primrec %quote (in monoid) pow :: "nat \ 'a \ 'a" where "pow 0 a = \" | "pow (Suc n) a = a \ pow n a" text \ \noindent This we use to define the discrete exponentiation function: \ definition %quote bexp :: "nat \ nat" where "bexp n = pow n (Suc (Suc 0))" text \ \noindent The corresponding code in Haskell uses that language's native classes: \ text %quote \ @{code_stmts bexp (Haskell)} \ text \ \noindent This is a convenient place to show how explicit dictionary construction manifests in generated code -- the same example in \SML\: \ text %quote \ @{code_stmts bexp (SML)} \ text \ \noindent Note the parameters with trailing underscore (\<^verbatim>\A_\), which are the dictionary parameters. \ subsection \How to continue from here\ text \ What you have seen so far should be already enough in a lot of cases. If you are content with this, you can quit reading here. Anyway, to understand situations where problems occur or to increase the scope of code generation beyond default, it is necessary to gain some understanding how the code generator actually works: \begin{itemize} \item The foundations of the code generator are described in \secref{sec:foundations}. \item In particular \secref{sec:utterly_wrong} gives hints how to debug situations where code generation does not succeed as expected. \item The scope and quality of generated code can be increased dramatically by applying refinement techniques, which are introduced in \secref{sec:refinement}. \item How to define partial functions such that code can be generated is explained in \secref{sec:partial}. \item Inductive predicates can be turned executable using an extension of the code generator \secref{sec:inductive}. \item If you want to utilize code generation to obtain fast evaluators e.g.~for decision procedures, have a look at \secref{sec:evaluation}. \item You may want to skim over the more technical sections \secref{sec:adaptation} and \secref{sec:further}. - \item The target language Scala @{cite "scala-overview-tech-report"} + \item The target language Scala \<^cite>\"scala-overview-tech-report"\ comes with some specialities discussed in \secref{sec:scala}. \item For exhaustive syntax diagrams etc. you should visit the - Isabelle/Isar Reference Manual @{cite "isabelle-isar-ref"}. + Isabelle/Isar Reference Manual \<^cite>\"isabelle-isar-ref"\. \end{itemize} \bigskip \begin{center}\fbox{\fbox{\begin{minipage}{8cm} \begin{center}\textit{Happy proving, happy hacking!}\end{center} \end{minipage}}}\end{center} \ end diff --git a/src/Doc/Codegen/Partial_Functions.thy b/src/Doc/Codegen/Partial_Functions.thy --- a/src/Doc/Codegen/Partial_Functions.thy +++ b/src/Doc/Codegen/Partial_Functions.thy @@ -1,261 +1,261 @@ theory Partial_Functions imports Setup "HOL-Library.Monad_Syntax" begin section \Partial Functions \label{sec:partial}\ text \We demonstrate three approaches to defining executable partial recursive functions, i.e.\ functions that do not terminate for all inputs. The main points are the definitions of the functions and the inductive proofs about them. Our concrete example represents a typical termination problem: following a data structure that may contain cycles. We want to follow a mapping from \nat\ to \nat\ to the end (until we leave its domain). The mapping is represented by a list \ns :: nat list\ that maps \n\ to \ns ! n\. The details of the example are in some sense irrelevant but make the exposition more realistic. However, we hide most proofs or show only the characteristic opening.\ text \The list representation of the mapping can be abstracted to a relation. The order @{term "(ns!n, n)"} is the order that @{const wf} expects.\ definition rel :: "nat list \ (nat * nat) set" where "rel ns = set(zip ns [0..*) text \\noindent This relation should be acyclic to guarantee termination of the partial functions defined below.\ subsection \Tail recursion\ text \If a function is tail-recursive, an executable definition is easy:\ partial_function (tailrec) follow :: "nat list \ nat \ nat" where "follow ns n = (if n < length ns then follow ns (ns!n) else n)" text \Informing the code generator:\ declare follow.simps[code] text \Now @{const follow} is executable:\ value "follow [1,2,3] 0" text \For proofs about @{const follow} we need a @{const wf} relation on @{term "(ns,n)"} pairs that decreases with each recursive call. The first component stays the same but must be acyclic. The second component must decrease w.r.t @{const rel}:\ definition "rel_follow = same_fst (acyclic o rel) rel" lemma wf_follow: "wf rel_follow" (*<*) by (auto intro: wf_same_fst simp: rel_follow_def finite_acyclic_wf) text \A more explicit formulation of \rel_follow\: The first component stays the same but must be acyclic. The second component decreases w.r.t \rel\:\ lemma rel_follow_step: "\ m < length ms; acyclic (rel ms) \ \ ((ms, ms ! m), (ms, m)) \ rel_follow" by(force simp: rel_follow_def rel_def in_set_zip) (*>*) text \This is how you start an inductive proof about @{const follow} along @{const rel_follow}:\ lemma "acyclic(rel ms) \ follow ms m = n \ length ms \ n" proof (induction "(ms,m)" arbitrary: m n rule: wf_induct_rule[OF wf_follow]) (*<*) case 1 thus ?case using follow.simps rel_follow_step by fastforce qed (*>*) subsection \Option\ text \If the function is not tail-recursive, not all is lost: if we rewrite it to return an \option\ type, it can still be defined. In this case @{term "Some x"} represents the result \x\ and @{term None} represents represents nontermination. For example, counting the length of the chain represented by \ns\ can be defined like this:\ partial_function (option) count :: "nat list \ nat \ nat option" where "count ns n = (if n < length ns then do {k \ count ns (ns!n); Some (k+1)} else Some 0)" text \\noindent We use a Haskell-like \do\ notation (import @{theory "HOL-Library.Monad_Syntax"}) to abbreviate the clumsy explicit \noindent @{term "case count ns (ns!n) of Some k \ Some(k+1) | None \ None"}. \noindent The branch \None \ None\ represents the requirement that nontermination of a recursive call must lead to nontermination of the function. Now we can prove that @{const count} terminates for all acyclic maps:\ lemma "acyclic(rel ms) \ \k. count ms m = Some k" proof (induction "(ms,m)" arbitrary: ms m rule: wf_induct_rule[OF wf_follow]) (*<*) case 1 thus ?case by (metis bind.bind_lunit count.simps rel_follow_step) qed (*>*) subsection \Subtype\ text \In this approach we define a new type that contains only elements on which the function in question terminates. In our example that is the subtype of all \ns :: nat list\ such that @{term "rel ns"} is acyclic. Then @{const follow} can be defined as a total function on that subtype.\ text \The subtype is not empty:\ lemma acyclic_rel_Nil: "acyclic(rel [])" (*<*)by (simp add: rel_def acyclic_def)(*>*) text \Definition of subtype \acyc\:\ typedef acyc = "{ns. acyclic(rel ns)}" morphisms rep_acyc abs_acyc using acyclic_rel_Nil by auto text \\noindent This defines two functions @{term [source] "rep_acyc :: acyc \ nat list"} and @{const abs_acyc} \::\ \mbox{@{typ "nat list \ acyc"}}. Function @{const abs_acyc} is only defined on acyclic lists and not executable for that reason. Type \dlist\ in Section~\ref{sec:partiality} is defined in the same manner. The following command sets up infrastructure for lifting functions on @{typ "nat list"} -to @{typ acyc} (used by @{command_def lift_definition} below) \cite{isabelle-isar-ref}.\ +to @{typ acyc} (used by @{command_def lift_definition} below) \<^cite>\"isabelle-isar-ref"\.\ setup_lifting type_definition_acyc text \This is how @{const follow} can be defined on @{typ acyc}:\ function follow2 :: "acyc \ nat \ nat" where "follow2 ac n = (let ns = rep_acyc ac in if n < length ns then follow2 ac (ns!n) else n)" by pat_completeness auto text \Now we prepare for the termination proof. Relation \rel_follow2\ is almost identical to @{const rel_follow}.\ definition "rel_follow2 = same_fst (acyclic o rel o rep_acyc) (rel o rep_acyc)" lemma wf_follow2: "wf rel_follow2" (*<*) by (auto intro: wf_same_fst simp add: rel_follow2_def finite_acyclic_wf) text \A more explicit formulation of \rel_follow2\:\ lemma rel_follow2_step: "\ ns = rep_acyc ac; m < length ns; acyclic (rel ns) \ \ ((ac, ns ! m), (ac, m)) \ rel_follow2" by(force simp add: rel_follow2_def rel_def in_set_zip) (*>*) text \Here comes the actual termination proof:\ termination follow2 proof show "wf rel_follow2" (*<*) by (auto intro: wf_same_fst simp add: rel_follow2_def finite_acyclic_wf)(*>*) next show "\ac n ns. \ ns = rep_acyc ac; n < length ns \ \ ((ac, ns ! n), (ac, n)) \ rel_follow2" (*<*) using rel_follow2_step rep_acyc by simp(*>*) qed text \Inductive proofs about @{const follow2} can now simply use computation induction:\ lemma "follow2 ac m = n \ length (rep_acyc ac) \ n" proof (induction ac m arbitrary: n rule: follow2.induct) (*<*) case 1 thus ?case by (metis (full_types) follow2.simps linorder_not_le) qed (*>*) text \A complication with the subtype approach is that injection into the subtype (function \abs_acyc\ in our example) is not executable. But to call @{const follow2}, we need an argument of type \acyc\ and we need to obtain it in an executable manner. There are two approaches.\ text \In the first approach we check wellformedness (i.e. acyclicity) explicitly. This check needs to be executable (which @{const acyclic} and @{const rel} are). If the check fails, @{term "[]"} is returned (which is acyclic).\ lift_definition is_acyc :: "nat list \ acyc" is "\ns. if acyclic(rel ns) then ns else []" (*<*)by (auto simp: acyclic_rel_Nil)(*>*) text \\noindent This works because we can easily prove that for any \ns\, the \\\-term produces an acyclic list. But it requires the possibly expensive check @{prop "acyclic(rel ns)"}.\ definition "follow_test ns n = follow2 (is_acyc ns) n" text \The relation is acyclic (a chain):\ value "follow_test [1,2,3] 1" text \In the second approach, wellformedness of the argument is guaranteed by construction. In our example \[1.. represents an acyclic chain \i \ i+1\\ lemma acyclic_chain: "acyclic (rel [1..*) text \\ lift_definition acyc_chain :: "nat \ acyc" is "\n. [1..*) text \\ definition "follow_chain m n = follow2 (acyc_chain m) n" value "follow_chain 5 1" text \The subtype approach requires neither tail-recursion nor \option\ but you cannot easily modify arguments of the subtype except via existing functions on the subtype. Otherwise you need to inject some value into the subtype and that injection is not computable. \ (*<*) text \The problem with subtypes: the Abs function must not be used explicitly. This can be avoided by using \lift_definition\. Example:\ typedef nat1 = "{n::nat. n > 0}" by auto print_theorems setup_lifting type_definition_nat1 (* just boiler plate *) lift_definition mk1 :: "nat \ nat1" is "\n. if n>0 then n else 1" by auto lift_definition g :: "nat1 \ nat1" is "\n. if n \ 2 then n-1 else n" by auto print_theorems text \This generates \g.rep_eq: Rep_nat1 (g x) = (if 2 \ Rep_nat1 x then Rep_nat1 x - 1 else Rep_nat1 x)\ which is acceptable as an abstract code eqn. The manual definition of \g n1 = (let n = Rep_nat1 n1 in if 2 \ n then Abs_nat1 (n - 1) else Abs_nat1 n)\ looks non-executable but \g.rep_eq\ can be derived from it.\ value "g (mk1 3)" text \However, this trick does not work when passing an Abs-term as an argument, eg in a recursive call, because the Abs-term is `hidden' by the function call.\ (*>*) end \ No newline at end of file diff --git a/src/Doc/Codegen/Refinement.thy b/src/Doc/Codegen/Refinement.thy --- a/src/Doc/Codegen/Refinement.thy +++ b/src/Doc/Codegen/Refinement.thy @@ -1,277 +1,277 @@ theory Refinement imports Setup begin section \Program and datatype refinement \label{sec:refinement}\ text \ Code generation by shallow embedding (cf.~\secref{sec:principle}) allows to choose code equations and datatype constructors freely, given that some very basic syntactic properties are met; this flexibility opens up mechanisms for refinement which allow to extend the scope and quality of generated code dramatically. \ subsection \Program refinement\ text \ Program refinement works by choosing appropriate code equations explicitly (cf.~\secref{sec:equations}); as example, we use Fibonacci numbers: \ fun %quote fib :: "nat \ nat" where "fib 0 = 0" | "fib (Suc 0) = Suc 0" | "fib (Suc (Suc n)) = fib n + fib (Suc n)" text \ \noindent The runtime of the corresponding code grows exponential due to two recursive calls: \ text %quote \ @{code_stmts fib constant: fib (Haskell)} \ text \ \noindent A more efficient implementation would use dynamic programming, e.g.~sharing of common intermediate results between recursive calls. This idea is expressed by an auxiliary operation which computes a Fibonacci number and its successor simultaneously: \ definition %quote fib_step :: "nat \ nat \ nat" where "fib_step n = (fib (Suc n), fib n)" text \ \noindent This operation can be implemented by recursion using dynamic programming: \ lemma %quote [code]: "fib_step 0 = (Suc 0, 0)" "fib_step (Suc n) = (let (m, q) = fib_step n in (m + q, m))" by (simp_all add: fib_step_def) text \ \noindent What remains is to implement \<^const>\fib\ by \<^const>\fib_step\ as follows: \ lemma %quote [code]: "fib 0 = 0" "fib (Suc n) = fst (fib_step n)" by (simp_all add: fib_step_def) text \ \noindent The resulting code shows only linear growth of runtime: \ text %quote \ @{code_stmts fib constant: fib fib_step (Haskell)} \ subsection \Datatype refinement\ text \ Selecting specific code equations \emph{and} datatype constructors leads to datatype refinement. As an example, we will develop an alternative representation of the queue example given in \secref{sec:queue_example}. The amortised representation is convenient for generating code but exposes its \qt{implementation} details, which may be cumbersome when proving theorems about it. Therefore, here is a simple, straightforward representation of queues: \ datatype %quote 'a queue = Queue "'a list" definition %quote empty :: "'a queue" where "empty = Queue []" primrec %quote enqueue :: "'a \ 'a queue \ 'a queue" where "enqueue x (Queue xs) = Queue (xs @ [x])" fun %quote dequeue :: "'a queue \ 'a option \ 'a queue" where "dequeue (Queue []) = (None, Queue [])" | "dequeue (Queue (x # xs)) = (Some x, Queue xs)" text \ \noindent This we can use directly for proving; for executing, we provide an alternative characterisation: \ definition %quote AQueue :: "'a list \ 'a list \ 'a queue" where "AQueue xs ys = Queue (ys @ rev xs)" code_datatype %quote AQueue text \ \noindent Here we define a \qt{constructor} \<^const>\AQueue\ which is defined in terms of \Queue\ and interprets its arguments according to what the \emph{content} of an amortised queue is supposed to be. The prerequisite for datatype constructors is only syntactical: a constructor must be of type \\ = \ \ \ \\<^sub>1 \ \\<^sub>n\ where \{\\<^sub>1, \, \\<^sub>n}\ is exactly the set of \emph{all} type variables in \\\; then \\\ is its corresponding datatype. The HOL datatype package by default registers any new datatype with its constructors, but this may be changed using @{command_def code_datatype}; the currently chosen constructors can be inspected using the @{command print_codesetup} command. Equipped with this, we are able to prove the following equations for our primitive queue operations which \qt{implement} the simple queues in an amortised fashion: \ lemma %quote empty_AQueue [code]: "empty = AQueue [] []" by (simp add: AQueue_def empty_def) lemma %quote enqueue_AQueue [code]: "enqueue x (AQueue xs ys) = AQueue (x # xs) ys" by (simp add: AQueue_def) lemma %quote dequeue_AQueue [code]: "dequeue (AQueue xs []) = (if xs = [] then (None, AQueue [] []) else dequeue (AQueue [] (rev xs)))" "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)" by (simp_all add: AQueue_def) text \ \noindent It is good style, although no absolute requirement, to provide code equations for the original artefacts of the implemented type, if possible; in our case, these are the datatype constructor \<^const>\Queue\ and the case combinator \<^const>\case_queue\: \ lemma %quote Queue_AQueue [code]: "Queue = AQueue []" by (simp add: AQueue_def fun_eq_iff) lemma %quote case_queue_AQueue [code]: "case_queue f (AQueue xs ys) = f (ys @ rev xs)" by (simp add: AQueue_def) text \ \noindent The resulting code looks as expected: \ text %quote \ @{code_stmts empty enqueue dequeue Queue case_queue (SML)} \ text \ The same techniques can also be applied to types which are not specified as datatypes, e.g.~type \<^typ>\int\ is originally specified as quotient type by means of @{command_def typedef}, but for code generation constants allowing construction of binary numeral values are used as constructors for \<^typ>\int\. This approach however fails if the representation of a type demands invariants; this issue is discussed in the next section. \ subsection \Datatype refinement involving invariants \label{sec:invariant}\ text \ Datatype representation involving invariants require a dedicated setup for the type and its primitive operations. As a running example, we implement a type \<^typ>\'a dlist\ of lists consisting of distinct elements. The specification of \<^typ>\'a dlist\ itself can be found in theory \<^theory>\HOL-Library.Dlist\. The first step is to decide on which representation the abstract type (in our example \<^typ>\'a dlist\) should be implemented. Here we choose \<^typ>\'a list\. Then a conversion from the concrete type to the abstract type must be specified, here: \ text %quote \ \<^term_type>\Dlist\ \ text \ \noindent Next follows the specification of a suitable \emph{projection}, i.e.~a conversion from abstract to concrete type: \ text %quote \ \<^term_type>\list_of_dlist\ \ text \ \noindent This projection must be specified such that the following \emph{abstract datatype certificate} can be proven: \ lemma %quote [code abstype]: "Dlist (list_of_dlist dxs) = dxs" by (fact Dlist_list_of_dlist) text \ \noindent Note that so far the invariant on representations (\<^term_type>\distinct\) has never been mentioned explicitly: the invariant is only referred to implicitly: all values in set \<^term>\{xs. list_of_dlist (Dlist xs) = xs}\ are invariant, and in our example this is exactly \<^term>\{xs. distinct xs}\. The primitive operations on \<^typ>\'a dlist\ are specified indirectly using the projection \<^const>\list_of_dlist\. For the empty \dlist\, \<^const>\Dlist.empty\, we finally want the code equation \ text %quote \ \<^term>\Dlist.empty = Dlist []\ \ text \ \noindent This we have to prove indirectly as follows: \ lemma %quote [code]: "list_of_dlist Dlist.empty = []" by (fact list_of_dlist_empty) text \ \noindent This equation logically encodes both the desired code equation and that the expression \<^const>\Dlist\ is applied to obeys the implicit invariant. Equations for insertion and removal are similar: \ lemma %quote [code]: "list_of_dlist (Dlist.insert x dxs) = List.insert x (list_of_dlist dxs)" by (fact list_of_dlist_insert) lemma %quote [code]: "list_of_dlist (Dlist.remove x dxs) = remove1 x (list_of_dlist dxs)" by (fact list_of_dlist_remove) text \ \noindent Then the corresponding code is as follows: \ text %quote \ @{code_stmts Dlist.empty Dlist.insert Dlist.remove list_of_dlist (SML)} \ text \ - See further @{cite "Haftmann-Kraus-Kuncar-Nipkow:2013:data_refinement"} + See further \<^cite>\"Haftmann-Kraus-Kuncar-Nipkow:2013:data_refinement"\ for the meta theory of datatype refinement involving invariants. Typical data structures implemented by representations involving invariants are available in the library, theory \<^theory>\HOL-Library.Mapping\ specifies key-value-mappings (type \<^typ>\('a, 'b) mapping\); these can be implemented by red-black-trees (theory \<^theory>\HOL-Library.RBT\). \ end diff --git a/src/Doc/Corec/Corec.thy b/src/Doc/Corec/Corec.thy --- a/src/Doc/Corec/Corec.thy +++ b/src/Doc/Corec/Corec.thy @@ -1,971 +1,971 @@ (* Title: Doc/Corec/Corec.thy Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Aymeric Bouzy, Ecole polytechnique Author: Andreas Lochbihler, ETH Zuerich Author: Andrei Popescu, Middlesex University Author: Dmitriy Traytel, ETH Zuerich Tutorial for nonprimitively corecursive definitions. *) theory Corec imports Main Datatypes.Setup "HOL-Library.BNF_Corec" "HOL-Library.FSet" begin section \Introduction \label{sec:introduction}\ text \ -Isabelle's (co)datatype package @{cite "isabelle-datatypes"} offers a convenient +Isabelle's (co)datatype package \<^cite>\"isabelle-datatypes"\ offers a convenient syntax for introducing codatatypes. For example, the type of (infinite) streams can be defined as follows (cf. \<^file>\~~/src/HOL/Library/Stream.thy\): \ codatatype 'a stream = SCons (shd: 'a) (stl: "'a stream") text \ \noindent The (co)datatype package also provides two commands, \keyw{primcorec} and \keyw{prim\-corec\-ur\-sive}, for defining primitively corecursive functions. This tutorial presents a definitional package for functions beyond primitive corecursion. It describes @{command corec} and related commands:\ @{command corecursive}, @{command friend_of_corec}, and @{command coinduction_upto}. It also covers the @{method corec_unique} proof method. The package is not part of \<^theory>\Main\; it is located in \<^file>\~~/src/HOL/Library/BNF_Corec.thy\. The @{command corec} command generalizes \keyw{primcorec} in three main respects. First, it allows multiple constructors around corecursive calls, where \keyw{primcorec} expects exactly one. For example: \ corec oneTwos :: "nat stream" where "oneTwos = SCons 1 (SCons 2 oneTwos)" text \ Second, @{command corec} allows other functions than constructors to appear in the corecursive call context (i.e., around any self-calls on the right-hand side of the equation). The requirement on these functions is that they must be \emph{friendly}. Intuitively, a function is friendly if it needs to destruct at most one constructor of input to produce one constructor of output. We can register functions as friendly using the @{command friend_of_corec} command, or by passing the \friend\ option to @{command corec}. The friendliness check relies on an internal syntactic check in combination with a parametricity subgoal, which must be discharged manually (typically using @{method transfer_prover} or @{method transfer_prover_eq}). Third, @{command corec} allows self-calls that are not guarded by a constructor, as long as these calls occur in a friendly context (a context consisting exclusively of friendly functions) and can be shown to be terminating (well founded). The mixture of recursive and corecursive calls in a single function can be quite useful in practice. Internally, the package synthesizes corecursors that take into account the possible call contexts. The corecursor is accompanined by a corresponding, equally general coinduction principle. The corecursor and the coinduction principle grow in expressiveness as we interact with it. In process algebra terminology, corecursion and coinduction take place \emph{up to} friendly contexts. -The package fully adheres to the LCF philosophy @{cite mgordon79}: The +The package fully adheres to the LCF philosophy \<^cite>\mgordon79\: The characteristic theorems associated with the specified corecursive functions are derived rather than introduced axiomatically. (Exceptionally, most of the internal proof obligations are omitted if the \quick_and_dirty\ option is enabled.) The package is described in a pair of scientific papers -@{cite "blanchette-et-al-2015-fouco" and "blanchette-et-al-201x-amico"}. Some +\<^cite>\"blanchette-et-al-2015-fouco" and "blanchette-et-al-201x-amico"\. Some of the text and examples below originate from there. This tutorial is organized as follows: \begin{itemize} \setlength{\itemsep}{0pt} \item Section \ref{sec:introductory-examples}, ``Introductory Examples,'' describes how to specify corecursive functions and to reason about them. \item Section \ref{sec:command-syntax}, ``Command Syntax,'' describes the syntax of the commands offered by the package. \item Section \ref{sec:generated-theorems}, ``Generated Theorems,'' lists the theorems produced by the package's commands. \item Section \ref{sec:proof-methods}, ``Proof Methods,'' briefly describes the @{method corec_unique} and @{method transfer_prover_eq} proof methods. \item Section \ref{sec:attribute}, ``Attribute,'' briefly describes the @{attribute friend_of_corec_simps} attribute, which can be used to strengthen the tactics underlying the @{command friend_of_corec} and @{command corec} \(friend)\ commands. \item Section \ref{sec:known-bugs-and-limitations}, ``Known Bugs and Limitations,'' concludes with known open issues. \end{itemize} Although it is more powerful than \keyw{primcorec} in many respects, @{command corec} suffers from a number of limitations. Most notably, it does not support mutually corecursive codatatypes, and it is less efficient than \keyw{primcorec} because it needs to dynamically synthesize corecursors and corresponding coinduction principles to accommodate the friends. \newbox\boxA \setbox\boxA=\hbox{\texttt{NOSPAM}} \newcommand\authoremaili{\texttt{jasmin.blan{\color{white}NOSPAM}\kern-\wd\boxA{}chette@\allowbreak gmail.\allowbreak com}} Comments and bug reports concerning either the package or this tutorial should be directed to the first author at \authoremaili{} or to the \texttt{cl-isabelle-users} mailing list. \ section \Introductory Examples \label{sec:introductory-examples}\ text \ The package is illustrated through concrete examples featuring different flavors of corecursion. More examples can be found in the directory \<^dir>\~~/src/HOL/Corec_Examples\. \ subsection \Simple Corecursion \label{ssec:simple-corecursion}\ text \ -The case studies by Rutten~@{cite rutten05} and Hinze~@{cite hinze10} on stream +The case studies by Rutten~\<^cite>\rutten05\ and Hinze~\<^cite>\hinze10\ on stream calculi serve as our starting point. The following definition of pointwise sum can be performed with either \keyw{primcorec} or @{command corec}: \ primcorec ssum :: "('a :: plus) stream \ 'a stream \ 'a stream" where "ssum xs ys = SCons (shd xs + shd ys) (ssum (stl xs) (stl ys))" text \ \noindent Pointwise sum meets the friendliness criterion. We register it as a friend using the @{command friend_of_corec} command. The command requires us to give a specification of \<^const>\ssum\ where a constructor (\<^const>\SCons\) occurs at the outermost position on the right-hand side. Here, we can simply reuse the \keyw{primcorec} specification above: \ friend_of_corec ssum :: "('a :: plus) stream \ 'a stream \ 'a stream" where "ssum xs ys = SCons (shd xs + shd ys) (ssum (stl xs) (stl ys))" apply (rule ssum.code) by transfer_prover text \ \noindent The command emits two subgoals. The first subgoal corresponds to the equation we specified and is trivial to discharge. The second subgoal is a parametricity property that captures the the requirement that the function may destruct at most one constructor of input to produce one constructor of output. This subgoal can usually be discharged using the \transfer_prover\ or @{method transfer_prover_eq} proof method (Section~\ref{ssec:transfer-prover-eq}). The latter replaces equality relations by their relator terms according to the @{thm [source] relator_eq} theorem collection before it invokes @{method transfer_prover}. After registering \<^const>\ssum\ as a friend, we can use it in the corecursive call context, either inside or outside the constructor guard: \ corec fibA :: "nat stream" where "fibA = SCons 0 (ssum (SCons 1 fibA) fibA)" text \\blankline\ corec fibB :: "nat stream" where "fibB = ssum (SCons 0 (SCons 1 fibB)) (SCons 0 fibB)" text \ Using the \friend\ option, we can simultaneously define a function and register it as a friend: \ corec (friend) sprod :: "('a :: {plus,times}) stream \ 'a stream \ 'a stream" where "sprod xs ys = SCons (shd xs * shd ys) (ssum (sprod xs (stl ys)) (sprod (stl xs) ys))" text \\blankline\ corec (friend) sexp :: "nat stream \ nat stream" where "sexp xs = SCons (2 ^^ shd xs) (sprod (stl xs) (sexp xs))" text \ \noindent The parametricity subgoal is given to \transfer_prover_eq\ (Section~\ref{ssec:transfer-prover-eq}). The \<^const>\sprod\ and \<^const>\sexp\ functions provide shuffle product and exponentiation on streams. We can use them to define the stream of factorial numbers in two different ways: \ corec factA :: "nat stream" where "factA = (let zs = SCons 1 factA in sprod zs zs)" text \\blankline\ corec factB :: "nat stream" where "factB = sexp (SCons 0 factB)" text \ The arguments of friendly functions can be of complex types involving the target codatatype. The following example defines the supremum of a finite set of streams by primitive corecursion and registers it as friendly: \ corec (friend) sfsup :: "nat stream fset \ nat stream" where "sfsup X = SCons (Sup (fset (fimage shd X))) (sfsup (fimage stl X))" text \ \noindent In general, the arguments may be any bounded natural functor (BNF) -@{cite "isabelle-datatypes"}, with the restriction that the target codatatype +\<^cite>\"isabelle-datatypes"\, with the restriction that the target codatatype (\<^typ>\nat stream\) may occur only in a \emph{live} position of the BNF. For this reason, the following function, on unbounded sets, cannot be registered as a friend: \ primcorec ssup :: "nat stream set \ nat stream" where "ssup X = SCons (Sup (image shd X)) (ssup (image stl X))" subsection \Nested Corecursion \label{ssec:nested-corecursion}\ text \ The package generally supports arbitrary codatatypes with multiple constructors and nesting through other type constructors (BNFs). Consider the following type of finitely branching Rose trees of potentially infinite depth: \ codatatype 'a tree = Node (lab: 'a) (sub: "'a tree list") text \ We first define the pointwise sum of two trees analogously to \<^const>\ssum\: \ corec (friend) tsum :: "('a :: plus) tree \ 'a tree \ 'a tree" where "tsum t u = Node (lab t + lab u) (map (\(t', u'). tsum t' u') (zip (sub t) (sub u)))" text \ \noindent Here, \<^const>\map\ is the standard map function on lists, and \<^const>\zip\ converts two parallel lists into a list of pairs. The \<^const>\tsum\ function is primitively corecursive. Instead of @{command corec} \(friend)\, we could also have used \keyw{primcorec} and @{command friend_of_corec}, as we did for \<^const>\ssum\. Once \<^const>\tsum\ is registered as friendly, we can use it in the corecursive call context of another function: \ corec (friend) ttimes :: "('a :: {plus,times}) tree \ 'a tree \ 'a tree" where "ttimes t u = Node (lab t * lab u) (map (\(t', u'). tsum (ttimes t u') (ttimes t' u)) (zip (sub t) (sub u)))" text \ All the syntactic convenience provided by \keyw{primcorec} is also supported by @{command corec}, @{command corecursive}, and @{command friend_of_corec}. In particular, nesting through the function type can be expressed using \\\-abstractions and function applications rather than through composition (\<^term>\(\)\, the map function for \\\). For example: \ codatatype 'a language = Lang (\: bool) (\
: "'a \ 'a language") text \\blankline\ corec (friend) Plus :: "'a language \ 'a language \ 'a language" where "Plus r s = Lang (\ r \ \ s) (\a. Plus (\
r a) (\
s a))" text \\blankline\ corec (friend) Times :: "'a language \ 'a language \ 'a language" where "Times r s = Lang (\ r \ \ s) (\a. if \ r then Plus (Times (\
r a) s) (\
s a) else Times (\
r a) s)" text \\blankline\ corec (friend) Star :: "'a language \ 'a language" where "Star r = Lang True (\a. Times (\
r a) (Star r))" text \\blankline\ corec (friend) Inter :: "'a language \ 'a language \ 'a language" where "Inter r s = Lang (\ r \ \ s) (\a. Inter (\
r a) (\
s a))" text \\blankline\ corec (friend) PLUS :: "'a language list \ 'a language" where "PLUS xs = Lang (\x \ set xs. \ x) (\a. PLUS (map (\r. \
r a) xs))" subsection \Mixed Recursion--Corecursion \label{ssec:mixed-recursion-corecursion}\ text \ It is often convenient to let a corecursive function perform some finite computation before producing a constructor. With mixed recursion--corecursion, a finite number of unguarded recursive calls perform this calculation before reaching a guarded corecursive call. Intuitively, the unguarded recursive call can be unfolded to arbitrary finite depth, ultimately yielding a purely corecursive definition. An example is the \<^term>\primes\ function from Di -Gianantonio and Miculan @{cite "di-gianantonio-miculan-2003"}: +Gianantonio and Miculan \<^cite>\"di-gianantonio-miculan-2003"\: \ corecursive primes :: "nat \ nat \ nat stream" where "primes m n = (if (m = 0 \ n > 1) \ coprime m n then SCons n (primes (m * n) (n + 1)) else primes m (n + 1))" apply (relation "measure (\(m, n). if n = 0 then 1 else if coprime m n then 0 else m - n mod m)") apply (auto simp: mod_Suc diff_less_mono2 intro: Suc_lessI elim!: not_coprimeE) apply (metis dvd_1_iff_1 dvd_eq_mod_eq_0 mod_0 mod_Suc mod_Suc_eq mod_mod_cancel) done text \ \noindent The @{command corecursive} command is a variant of @{command corec} that allows us to specify a termination argument for any unguarded self-call. When called with \m = 1\ and \n = 2\, the \<^const>\primes\ function computes the stream of prime numbers. The unguarded call in the \else\ branch increments \<^term>\n\ until it is coprime to the first argument \<^term>\m\ (i.e., the greatest common divisor of \<^term>\m\ and \<^term>\n\ is \1\). For any positive integers \<^term>\m\ and \<^term>\n\, the numbers \<^term>\m\ and \m * n + 1\ are coprime, yielding an upper bound on the number of times \<^term>\n\ is increased. Hence, the function will take the \else\ branch at most finitely often before taking the then branch and producing one constructor. There is a slight complication when \m = 0 \ n > 1\: Without the first disjunct in the \if\ condition, the function could stall. (This corner case was overlooked in the original example -@{cite "di-gianantonio-miculan-2003"}.) +\<^cite>\"di-gianantonio-miculan-2003"\.) In the following examples, termination is discharged automatically by @{command corec} by invoking @{method lexicographic_order}: \ corec catalan :: "nat \ nat stream" where "catalan n = (if n > 0 then ssum (catalan (n - 1)) (SCons 0 (catalan (n + 1))) else SCons 1 (catalan 1))" text \\blankline\ corec collatz :: "nat \ nat stream" where "collatz n = (if even n \ n > 0 then collatz (n div 2) else SCons n (collatz (3 * n + 1)))" text \ A more elaborate case study, revolving around the filter function on lazy lists, is presented in \<^file>\~~/src/HOL/Corec_Examples/LFilter.thy\. \ subsection \Self-Friendship \label{ssec:self-friendship}\ text \ The package allows us to simultaneously define a function and use it as its own friend, as in the following definition of a ``skewed product'': \ corec (friend) sskew :: "('a :: {plus,times}) stream \ 'a stream \ 'a stream" where "sskew xs ys = SCons (shd xs * shd ys) (sskew (sskew xs (stl ys)) (sskew (stl xs) ys))" text \ \noindent Such definitions, with nested self-calls on the right-hand side, cannot be separated into a @{command corec} part and a @{command friend_of_corec} part. \ subsection \Coinduction \label{ssec:coinduction}\ text \ Once a corecursive specification has been accepted, we normally want to reason about it. The \codatatype\ command generates a structural coinduction principle that matches primitively corecursive functions. For nonprimitive specifications, our package provides the more advanced proof principle of \emph{coinduction up to congruence}---or simply \emph{coinduction up-to}. The structural coinduction principle for \<^typ>\'a stream\, called @{thm [source] stream.coinduct}, is as follows: % \begin{indentblock} @{thm stream.coinduct[no_vars]} \end{indentblock} % Coinduction allows us to prove an equality \l = r\ on streams by providing a relation \R\ that relates \l\ and \r\ (first premise) and that constitutes a bisimulation (second premise). Streams that are related by a bisimulation cannot be distinguished by taking observations (via the selectors \<^const>\shd\ and \<^const>\stl\); hence they must be equal. The coinduction up-to principle after registering \<^const>\sskew\ as friendly is available as @{thm [source] sskew.coinduct} and as one of the components of the theorem collection @{thm [source] stream.coinduct_upto}: % \begin{indentblock} @{thm sskew.coinduct[no_vars]} \end{indentblock} % This rule is almost identical to structural coinduction, except that the corecursive application of \<^term>\R\ is generalized to \<^term>\stream.v5.congclp R\. The \<^const>\stream.v5.congclp\ predicate is equipped with the following introduction rules: \begin{indentblock} \begin{description} \item[@{thm [source] sskew.cong_base}\rm:] ~ \\ @{thm sskew.cong_base[no_vars]} \item[@{thm [source] sskew.cong_refl}\rm:] ~ \\ @{thm sskew.cong_refl[no_vars]} \item[@{thm [source] sskew.cong_sym}\rm:] ~ \\ @{thm sskew.cong_sym[no_vars]} \item[@{thm [source] sskew.cong_trans}\rm:] ~ \\ @{thm sskew.cong_trans[no_vars]} \item[@{thm [source] sskew.cong_SCons}\rm:] ~ \\ @{thm sskew.cong_SCons[no_vars]} \item[@{thm [source] sskew.cong_ssum}\rm:] ~ \\ @{thm sskew.cong_ssum[no_vars]} \item[@{thm [source] sskew.cong_sprod}\rm:] ~ \\ @{thm sskew.cong_sprod[no_vars]} \item[@{thm [source] sskew.cong_sskew}\rm:] ~ \\ @{thm sskew.cong_sskew[no_vars]} \end{description} \end{indentblock} % The introduction rules are also available as @{thm [source] sskew.cong_intros}. Notice that there is no introduction rule corresponding to \<^const>\sexp\, because \<^const>\sexp\ has a more restrictive result type than \<^const>\sskew\ (\<^typ>\nat stream\ vs. \<^typ>\('a :: {plus,times}) stream\. The version numbers, here \v5\, distinguish the different congruence closures generated for a given codatatype as more friends are registered. As much as possible, it is recommended to avoid referring to them in proof documents. Since the package maintains a set of incomparable corecursors, there is also a set of associated coinduction principles and a set of sets of introduction rules. A technically subtle point is to make Isabelle choose the right rules in most situations. For this purpose, the package maintains the collection @{thm [source] stream.coinduct_upto} of coinduction principles ordered by increasing generality, which works well with Isabelle's philosophy of applying the first rule that matches. For example, after registering \<^const>\ssum\ as a friend, proving the equality \<^term>\l = r\ on \<^typ>\nat stream\ might require coinduction principle for \<^term>\nat stream\, which is up to \<^const>\ssum\. The collection @{thm [source] stream.coinduct_upto} is guaranteed to be complete and up to date with respect to the type instances of definitions considered so far, but occasionally it may be necessary to take the union of two incomparable coinduction principles. This can be done using the @{command coinduction_upto} command. Consider the following definitions: \ codatatype ('a, 'b) tllist = TNil (terminal: 'b) | TCons (thd: 'a) (ttl: "('a, 'b) tllist") text \\blankline\ corec (friend) square_elems :: "(nat, 'b) tllist \ (nat, 'b) tllist" where "square_elems xs = (case xs of TNil z \ TNil z | TCons y ys \ TCons (y ^^ 2) (square_elems ys))" text \\blankline\ corec (friend) square_terminal :: "('a, int) tllist \ ('a, int) tllist" where "square_terminal xs = (case xs of TNil z \ TNil (z ^^ 2) | TCons y ys \ TCons y (square_terminal ys))" text \ At this point, @{thm [source] tllist.coinduct_upto} contains three variants of the coinduction principles: % \begin{itemize} \item \<^typ>\('a, int) tllist\ up to \<^const>\TNil\, \<^const>\TCons\, and \<^const>\square_terminal\; \item \<^typ>\(nat, 'b) tllist\ up to \<^const>\TNil\, \<^const>\TCons\, and \<^const>\square_elems\; \item \<^typ>\('a, 'b) tllist\ up to \<^const>\TNil\ and \<^const>\TCons\. \end{itemize} % The following variant is missing: % \begin{itemize} \item \<^typ>\(nat, int) tllist\ up to \<^const>\TNil\, \<^const>\TCons\, \<^const>\square_elems\, and \<^const>\square_terminal\. \end{itemize} % To generate it without having to define a new function with @{command corec}, we can use the following command: \ coinduction_upto nat_int_tllist: "(nat, int) tllist" text \ \noindent This produces the theorems % \begin{indentblock} @{thm [source] nat_int_tllist.coinduct_upto} \\ @{thm [source] nat_int_tllist.cong_intros} \end{indentblock} % (as well as the individually named introduction rules) and extends the dynamic collections @{thm [source] tllist.coinduct_upto} and @{thm [source] tllist.cong_intros}. \ subsection \Uniqueness Reasoning \label{ssec:uniqueness-reasoning}\ text \ It is sometimes possible to achieve better automation by using a more specialized proof method than coinduction. Uniqueness principles maintain a good balance between expressiveness and automation. They exploit the property that a corecursive definition is the unique solution to a fixpoint equation. The @{command corec}, @{command corecursive}, and @{command friend_of_corec} commands generate a property \f.unique\ about the function of interest \<^term>\f\ that can be used to prove that any function that satisfies \<^term>\f\'s corecursive specification must be equal to~\<^term>\f\. For example: \[@{thm ssum.unique[no_vars]}\] The uniqueness principles are not restricted to functions defined using @{command corec} or @{command corecursive} or registered with @{command friend_of_corec}. Suppose \<^term>\t x\ is an arbitrary term depending on \<^term>\x\. The @{method corec_unique} proof method, provided by our tool, transforms subgoals of the form \[\<^term>\(\x. f x = H x f) \ f x = t x\\] into \[\<^term>\\x. t x = H x t\\] The higher-order functional \<^term>\H\ must be such that \<^term>\f x = H x f\ would be a valid @{command corec} specification, but without nested self-calls or unguarded (recursive) calls. Thus, @{method corec_unique} proves uniqueness of \<^term>\t\ with respect to the given corecursive equation regardless of how \<^term>\t\ was defined. For example: \ lemma fixes f :: "nat stream \ nat stream \ nat stream" assumes "\xs ys. f xs ys = SCons (shd ys * shd xs) (ssum (f xs (stl ys)) (f (stl xs) ys))" shows "f = sprod" using assms proof corec_unique show "sprod = (\xs ys :: nat stream. SCons (shd ys * shd xs) (ssum (sprod xs (stl ys)) (sprod (stl xs) ys)))" apply (rule ext)+ apply (subst sprod.code) by simp qed text \ The proof method relies on some theorems generated by the package. If no function over a given codatatype has been defined using @{command corec} or @{command corecursive} or registered as friendly using @{command friend_of_corec}, the theorems will not be available yet. In such cases, the theorems can be explicitly generated using the command \ coinduction_upto stream: "'a stream" section \Command Syntax \label{sec:command-syntax}\ subsection \\keyw{corec} and \keyw{corecursive} \label{ssec:corec-and-corecursive-syntax}\ text \ \begin{matharray}{rcl} @{command_def "corec"} & : & \local_theory \ local_theory\ \\ @{command_def "corecursive"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ (@@{command corec} | @@{command corecursive}) target? \ @{syntax cr_options}? fix @'where' prop ; @{syntax_def cr_options}: '(' ((@{syntax plugins} | 'friend' | 'transfer') + ',') ')' \ \medskip \noindent The @{command corec} and @{command corecursive} commands introduce a corecursive function over a codatatype. The syntactic entity \synt{target} can be used to specify a local context, \synt{fix} denotes name with an optional type signature, and \synt{prop} denotes -a HOL proposition @{cite "isabelle-isar-ref"}. +a HOL proposition \<^cite>\"isabelle-isar-ref"\. The optional target is optionally followed by a combination of the following options: \begin{itemize} \setlength{\itemsep}{0pt} \item The \plugins\ option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. \item The \friend\ option indicates that the defined function should be registered as a friend. This gives rise to additional proof obligations. \item The \transfer\ option indicates that an unconditional transfer rule should be generated and proved \by transfer_prover\. The \[transfer_rule]\ attribute is set on the generated theorem. \end{itemize} The @{command corec} command is an abbreviation for @{command corecursive} with appropriate applications of @{method transfer_prover_eq} (Section~\ref{ssec:transfer-prover-eq}) and @{method lexicographic_order} to discharge any emerging proof obligations. \ subsection \\keyw{friend_of_corec} \label{ssec:friend-of-corec-syntax}\ text \ \begin{matharray}{rcl} @{command_def "friend_of_corec"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ @@{command friend_of_corec} target? \ @{syntax foc_options}? fix @'where' prop ; @{syntax_def foc_options}: '(' ((@{syntax plugins} | 'transfer') + ',') ')' \ \medskip \noindent The @{command friend_of_corec} command registers a corecursive function as friendly. The syntactic entity \synt{target} can be used to specify a local context, \synt{fix} denotes name with an optional type signature, and \synt{prop} denotes -a HOL proposition @{cite "isabelle-isar-ref"}. +a HOL proposition \<^cite>\"isabelle-isar-ref"\. The optional target is optionally followed by a combination of the following options: \begin{itemize} \setlength{\itemsep}{0pt} \item The \plugins\ option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. \item The \transfer\ option indicates that an unconditional transfer rule should be generated and proved \by transfer_prover\. The \[transfer_rule]\ attribute is set on the generated theorem. \end{itemize} \ subsection \\keyw{coinduction_upto} \label{ssec:coinduction-upto-syntax}\ text \ \begin{matharray}{rcl} @{command_def "coinduction_upto"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command coinduction_upto} target? name ':' type \ \medskip \noindent The @{command coinduction_upto} generates a coinduction up-to rule for a given instance of a (possibly polymorphic) codatatype and notes the result with the specified prefix. The syntactic entity \synt{name} denotes an identifier and \synt{type} denotes a -type @{cite "isabelle-isar-ref"}. +type \<^cite>\"isabelle-isar-ref"\. \ section \Generated Theorems \label{sec:generated-theorems}\ text \ The full list of named theorems generated by the package can be obtained by issuing the command \keyw{print_theorems} immediately after the datatype definition. This list excludes low-level theorems that reveal internal constructions. To make these accessible, add the line \ declare [[bnf_internals]] (*<*) declare [[bnf_internals = false]] (*>*) text \ In addition to the theorem listed below for each command provided by the package, all commands update the dynamic theorem collections \begin{indentblock} \begin{description} \item[\t.\\hthm{coinduct_upto}] \item[\t.\\hthm{cong_intros}] \end{description} \end{indentblock} % for the corresponding codatatype \t\ so that they always contain the most powerful coinduction up-to principles derived so far. \ subsection \\keyw{corec} and \keyw{corecursive} \label{ssec:corec-and-corecursive-theorems}\ text \ For a function \<^term>\f\ over codatatype \t\, the @{command corec} and @{command corecursive} commands generate the following properties (listed for \<^const>\sexp\, cf. Section~\ref{ssec:simple-corecursion}): \begin{indentblock} \begin{description} \item[\f.\\hthm{code} \[code]\\rm:] ~ \\ @{thm sexp.code[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin -@{cite "isabelle-datatypes"}. +\<^cite>\"isabelle-datatypes"\. \item[\f.\\hthm{coinduct} \[consumes 1, case_names t, case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm:] ~ \\ @{thm sexp.coinduct[no_vars]} \item[\f.\\hthm{cong_intros}\rm:] ~ \\ @{thm sexp.cong_intros[no_vars]} \item[\f.\\hthm{unique}\rm:] ~ \\ @{thm sexp.unique[no_vars]} \\ This property is not generated for mixed recursive--corecursive definitions. \item[\f.\\hthm{inner_induct}\rm:] ~ \\ This property is only generated for mixed recursive--corecursive definitions. For \<^const>\primes\ (Section~\ref{ssec:mixed-recursion-corecursion}, it reads as follows: \\[\jot] @{thm primes.inner_induct[no_vars]} \end{description} \end{indentblock} \noindent The individual rules making up \f.cong_intros\ are available as \begin{indentblock} \begin{description} \item[\f.\\hthm{cong_base}] \item[\f.\\hthm{cong_refl}] \item[\f.\\hthm{cong_sym}] \item[\f.\\hthm{cong_trans}] \item[\f.\\hthm{cong_C}\\<^sub>1\, \ldots, \f.\\hthm{cong_C}\\<^sub>n\] ~ \\ where \C\<^sub>1\, \\\, \C\<^sub>n\ are \t\'s constructors \item[\f.\\hthm{cong_f}\\<^sub>1\, \ldots, \f.\\hthm{cong_f}\\<^sub>m\] ~ \\ where \f\<^sub>1\, \\\, \f\<^sub>m\ are the available friends for \t\ \end{description} \end{indentblock} \ subsection \\keyw{friend_of_corec} \label{ssec:friend-of-corec-theorems}\ text \ The @{command friend_of_corec} command generates the same theorems as @{command corec} and @{command corecursive}, except that it adds an optional \friend.\ component to the names to prevent potential clashes (e.g., \f.friend.code\). \ subsection \\keyw{coinduction_upto} \label{ssec:coinduction-upto-theorems}\ text \ The @{command coinduction_upto} command generates the following properties (listed for \nat_int_tllist\): \begin{indentblock} \begin{description} \item[\begin{tabular}{@ {}l@ {}} \t.\\hthm{coinduct_upto} \[consumes 1, case_names t,\ \\ \phantom{\t.\\hthm{coinduct_upto} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \end{tabular}] ~ \\ @{thm nat_int_tllist.coinduct_upto[no_vars]} \item[\t.\\hthm{cong_intros}\rm:] ~ \\ @{thm nat_int_tllist.cong_intros[no_vars]} \end{description} \end{indentblock} \noindent The individual rules making up \t.cong_intros\ are available separately as \t.cong_base\, \t.cong_refl\, etc.\ (Section~\ref{ssec:corec-and-corecursive-theorems}). \ section \Proof Methods \label{sec:proof-methods}\ subsection \\textit{corec_unique} \label{ssec:corec-unique}\ text \ The @{method corec_unique} proof method can be used to prove the uniqueness of a corecursive specification. See Section~\ref{ssec:uniqueness-reasoning} for details. \ subsection \\textit{transfer_prover_eq} \label{ssec:transfer-prover-eq}\ text \ The @{method transfer_prover_eq} proof method replaces the equality relation \<^term>\(=)\ with compound relator expressions according to @{thm [source] relator_eq} before calling @{method transfer_prover} on the current subgoal. It tends to work better than plain @{method transfer_prover} on the parametricity proof obligations of @{command corecursive} and @{command friend_of_corec}, because they often contain equality relations on complex types, which @{method transfer_prover} cannot reason about. \ section \Attribute \label{sec:attribute}\ subsection \\textit{friend_of_corec_simps} \label{ssec:friend-of-corec-simps}\ text \ The @{attribute friend_of_corec_simps} attribute declares naturality theorems to be used by @{command friend_of_corec} and @{command corec} \(friend)\ in deriving the user specification from reduction to primitive corecursion. Internally, these commands derive naturality theorems from the parametricity proof obligations dischared by the user or the @{method transfer_prover_eq} method, but this derivation fails if in the arguments of a higher-order constant a type variable occurs on both sides of the function type constructor. The required naturality theorem can then be declared with @{attribute friend_of_corec_simps}. See \<^file>\~~/src/HOL/Corec_Examples/Tests/Iterate_GPV.thy\ for an example. \ section \Known Bugs and Limitations \label{sec:known-bugs-and-limitations}\ text \ This section lists the known bugs and limitations of the corecursion package at the time of this writing. \begin{enumerate} \setlength{\itemsep}{0pt} \item \emph{Mutually corecursive codatatypes are not supported.} \item \emph{The signature of friend functions may not depend on type variables beyond those that appear in the codatatype.} \item \emph{The internal tactics may fail on legal inputs.} In some cases, this limitation can be circumvented using the @{attribute friend_of_corec_simps} attribute (Section~\ref{ssec:friend-of-corec-simps}). \item \emph{The \transfer\ option is not implemented yet.} \item \emph{The constructor and destructor views offered by {\upshape\keyw{primcorec}} are not supported by @{command corec} and @{command corecursive}.} \item \emph{There is no mechanism for registering custom plugins.} \item \emph{The package does not interact well with locales.} \item \emph{The undocumented \corecUU_transfer\ theorem is not as polymorphic as it could be.} \item \emph{All type variables occurring in the arguments of a friendly function must occur as direct arguments of the type constructor of the resulting type.} \end{enumerate} \ end diff --git a/src/Doc/Datatypes/Datatypes.thy b/src/Doc/Datatypes/Datatypes.thy --- a/src/Doc/Datatypes/Datatypes.thy +++ b/src/Doc/Datatypes/Datatypes.thy @@ -1,3688 +1,3687 @@ (* Title: Doc/Datatypes/Datatypes.thy Author: Julian Biendarra, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, Ecole de technologie superieure Author: Lorenz Panny, TU Muenchen Author: Andrei Popescu, TU Muenchen Author: Dmitriy Traytel, TU Muenchen Tutorial for (co)datatype definitions. *) theory Datatypes imports Setup "HOL-Library.BNF_Axiomatization" "HOL-Library.Countable" "HOL-Library.FSet" "HOL-Library.Simps_Case_Conv" begin (*<*) unbundle cardinal_syntax (*>*) section \Introduction \label{sec:introduction}\ text \ The 2013 edition of Isabelle introduced a definitional package for freely generated datatypes and codatatypes. This package replaces the earlier -implementation due to Berghofer and Wenzel @{cite "Berghofer-Wenzel:1999:TPHOL"}. +implementation due to Berghofer and Wenzel \<^cite>\"Berghofer-Wenzel:1999:TPHOL"\. Perhaps the main advantage of the new package is that it supports recursion through a large class of non-datatypes, such as finite sets: \ datatype 'a tree\<^sub>f\<^sub>s = Node\<^sub>f\<^sub>s (lbl\<^sub>f\<^sub>s: 'a) (sub\<^sub>f\<^sub>s: "'a tree\<^sub>f\<^sub>s fset") text \ \noindent Another strong point is the support for local definitions: \ context linorder begin datatype flag = Less | Eq | Greater end text \ \noindent Furthermore, the package provides a lot of convenience, including automatically generated discriminators, selectors, and relators as well as a wealth of properties about them. In addition to inductive datatypes, the package supports coinductive datatypes, or \emph{codatatypes}, which allow infinite values. For example, the following command introduces the type of lazy lists, which comprises both finite and infinite values: \ (*<*) locale early locale late (*>*) codatatype (*<*)(in early) (*>*)'a llist = LNil | LCons 'a "'a llist" text \ \noindent Mixed inductive--coinductive recursion is possible via nesting. Compare the following four Rose tree examples: \ datatype (*<*)(in early) (*>*)'a tree\<^sub>f\<^sub>f = Node\<^sub>f\<^sub>f 'a "'a tree\<^sub>f\<^sub>f list" datatype (*<*)(in early) (*>*)'a tree\<^sub>f\<^sub>i = Node\<^sub>f\<^sub>i 'a "'a tree\<^sub>f\<^sub>i llist" codatatype (*<*)(in early) (*>*)'a tree\<^sub>i\<^sub>f = Node\<^sub>i\<^sub>f 'a "'a tree\<^sub>i\<^sub>f list" codatatype (*<*)(in early) (*>*)'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i 'a "'a tree\<^sub>i\<^sub>i llist" text \ The first two tree types allow only paths of finite length, whereas the last two allow infinite paths. Orthogonally, the nodes in the first and third types have finitely many direct subtrees, whereas those of the second and fourth may have infinite branching. The package is part of \<^theory>\Main\. Additional functionality is provided by the theory \<^file>\~~/src/HOL/Library/BNF_Axiomatization.thy\. The package, like its predecessor, fully adheres to the LCF philosophy -@{cite mgordon79}: The characteristic theorems associated with the specified +\<^cite>\mgordon79\: The characteristic theorems associated with the specified (co)datatypes are derived rather than introduced axiomatically.% \footnote{However, some of the internal constructions and most of the internal proof obligations are omitted if the \quick_and_dirty\ option is enabled.} The package is described in a number of scientific papers -@{cite "traytel-et-al-2012" and "blanchette-et-al-2014-impl" and -"panny-et-al-2014" and "blanchette-et-al-2015-wit"}. +\<^cite>\"traytel-et-al-2012" and "blanchette-et-al-2014-impl" and +"panny-et-al-2014" and "blanchette-et-al-2015-wit"\. The central notion is that of a \emph{bounded natural functor} (BNF)---a well-behaved type constructor for which nested (co)recursion is supported. This tutorial is organized as follows: \begin{itemize} \setlength{\itemsep}{0pt} \item Section \ref{sec:defining-datatypes}, ``Defining Datatypes,'' describes how to specify datatypes using the @{command datatype} command. \item Section \ref{sec:defining-primitively-recursive-functions}, ``Defining Primitively Recursive Functions,'' describes how to specify functions -using @{command primrec}. (A separate tutorial @{cite "isabelle-function"} +using @{command primrec}. (A separate tutorial \<^cite>\"isabelle-function"\ describes the more powerful \keyw{fun} and \keyw{function} commands.) \item Section \ref{sec:defining-codatatypes}, ``Defining Codatatypes,'' describes how to specify codatatypes using the @{command codatatype} command. \item Section \ref{sec:defining-primitively-corecursive-functions}, ``Defining Primitively Corecursive Functions,'' describes how to specify functions using the @{command primcorec} and @{command primcorecursive} commands. (A separate tutorial -@{cite "isabelle-corec"} describes the more powerful \keyw{corec} and +\<^cite>\"isabelle-corec"\ describes the more powerful \keyw{corec} and \keyw{corecursive} commands.) \item Section \ref{sec:registering-bounded-natural-functors}, ``Registering Bounded Natural Functors,'' explains how to use the @{command bnf} command to register arbitrary type constructors as BNFs. \item Section \ref{sec:deriving-destructors-and-constructor-theorems}, ``Deriving Destructors and Constructor Theorems,'' explains how to use the command @{command free_constructors} to derive destructor constants and theorems for freely generated types, as performed internally by @{command datatype} and @{command codatatype}. %\item Section \ref{sec:using-the-standard-ml-interface}, ``Using the Standard %ML Interface,'' describes the package's programmatic interface. \item Section \ref{sec:selecting-plugins}, ``Selecting Plugins,'' is concerned with the package's interoperability with other Isabelle packages and tools, such as the code generator, Transfer, Lifting, and Quickcheck. \item Section \ref{sec:known-bugs-and-limitations}, ``Known Bugs and Limitations,'' concludes with known open issues. \end{itemize} \newbox\boxA \setbox\boxA=\hbox{\texttt{NOSPAM}} \newcommand\authoremaili{\texttt{jasmin.blan{\color{white}NOSPAM}\kern-\wd\boxA{}chette@\allowbreak gmail.\allowbreak com}} Comments and bug reports concerning either the package or this tutorial should be directed to the second author at \authoremaili{} or to the \texttt{cl-isabelle-users} mailing list. \ section \Defining Datatypes \label{sec:defining-datatypes}\ text \ Datatypes can be specified using the @{command datatype} command. \ subsection \Introductory Examples \label{ssec:datatype-introductory-examples}\ text \ Datatypes are illustrated through concrete examples featuring different flavors of recursion. More examples can be found in the directory \<^dir>\~~/src/HOL/Datatype_Examples\. \ subsubsection \Nonrecursive Types \label{sssec:datatype-nonrecursive-types}\ text \ Datatypes are introduced by specifying the desired names and argument types for their constructors. \emph{Enumeration} types are the simplest form of datatype. All their constructors are nullary: \ datatype trool = Truue | Faalse | Perhaaps text \ \noindent \<^const>\Truue\, \<^const>\Faalse\, and \<^const>\Perhaaps\ have the type \<^typ>\trool\. Polymorphic types are possible, such as the following option type, modeled after its homologue from the \<^theory>\HOL.Option\ theory: \ (*<*) hide_const None Some map_option hide_type option (*>*) datatype 'a option = None | Some 'a text \ \noindent The constructors are \None :: 'a option\ and \Some :: 'a \ 'a option\. The next example has three type parameters: \ datatype ('a, 'b, 'c) triple = Triple 'a 'b 'c text \ \noindent The constructor is \Triple :: 'a \ 'b \ 'c \ ('a, 'b, 'c) triple\. Unlike in Standard ML, curried constructors are supported. The uncurried variant is also possible: \ datatype ('a, 'b, 'c) triple\<^sub>u = Triple\<^sub>u "'a * 'b * 'c" text \ \noindent Occurrences of nonatomic types on the right-hand side of the equal sign must be enclosed in double quotes, as is customary in Isabelle. \ subsubsection \Simple Recursion \label{sssec:datatype-simple-recursion}\ text \ Natural numbers are the simplest example of a recursive type: \ datatype nat = Zero | Succ nat text \ \noindent Lists were shown in the introduction. Terminated lists are a variant that stores a value of type \<^typ>\'b\ at the very end: \ datatype (*<*)(in early) (*>*)('a, 'b) tlist = TNil 'b | TCons 'a "('a, 'b) tlist" subsubsection \Mutual Recursion \label{sssec:datatype-mutual-recursion}\ text \ \emph{Mutually recursive} types are introduced simultaneously and may refer to each other. The example below introduces a pair of types for even and odd natural numbers: \ datatype even_nat = Even_Zero | Even_Succ odd_nat and odd_nat = Odd_Succ even_nat text \ \noindent Arithmetic expressions are defined via terms, terms via factors, and factors via expressions: \ datatype ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp" and ('a, 'b) trm = Factor "('a, 'b) fct" | Prod "('a, 'b) fct" "('a, 'b) trm" and ('a, 'b) fct = Const 'a | Var 'b | Expr "('a, 'b) exp" subsubsection \Nested Recursion \label{sssec:datatype-nested-recursion}\ text \ \emph{Nested recursion} occurs when recursive occurrences of a type appear under a type constructor. The introduction showed some examples of trees with nesting through lists. A more complex example, that reuses our \<^type>\option\ type, follows: \ datatype 'a btree = BNode 'a "'a btree option" "'a btree option" text \ \noindent Not all nestings are admissible. For example, this command will fail: \ datatype 'a wrong = W1 | W2 (*<*)'a typ (*>*)"'a wrong \ 'a" text \ \noindent The issue is that the function arrow \\\ allows recursion only through its right-hand side. This issue is inherited by polymorphic datatypes defined in terms of~\\\: \ datatype ('a, 'b) fun_copy = Fun "'a \ 'b" datatype 'a also_wrong = W1 | W2 (*<*)'a typ (*>*)"('a also_wrong, 'a) fun_copy" text \ \noindent The following definition of \<^typ>\'a\-branching trees is legal: \ datatype 'a ftree = FTLeaf 'a | FTNode "'a \ 'a ftree" text \ \noindent And so is the definition of hereditarily finite sets: \ datatype hfset = HFSet "hfset fset" text \ \noindent In general, type constructors \('a\<^sub>1, \, 'a\<^sub>m) t\ allow recursion on a subset of their type arguments \'a\<^sub>1\, \ldots, \'a\<^sub>m\. These type arguments are called \emph{live}; the remaining type arguments are called \emph{dead}. In \<^typ>\'a \ 'b\ and \<^typ>\('a, 'b) fun_copy\, the type variable \<^typ>\'a\ is dead and \<^typ>\'b\ is live. Type constructors must be registered as BNFs to have live arguments. This is done automatically for datatypes and codatatypes introduced by the @{command datatype} and @{command codatatype} commands. Section~\ref{sec:registering-bounded-natural-functors} explains how to register arbitrary type constructors as BNFs. Here is another example that fails: \ datatype 'a pow_list = PNil 'a (*<*)'a datatype 'a pow_list' = PNil' 'a (*>*)| PCons "('a * 'a) pow_list" text \ \noindent This attempted definition features a different flavor of nesting, where the recursive call in the type specification occurs around (rather than inside) another type constructor. \ subsubsection \Auxiliary Constants \label{sssec:datatype-auxiliary-constants}\ text \ The @{command datatype} command introduces various constants in addition to the constructors. With each datatype are associated set functions, a map function, a predicator, a relator, discriminators, and selectors, all of which can be given custom names. In the example below, the familiar names \null\, \hd\, \tl\, \set\, \map\, and \list_all2\ override the default names \is_Nil\, \un_Cons1\, \un_Cons2\, \set_list\, \map_list\, and \rel_list\: \ (*<*) no_translations "[x, xs]" == "x # [xs]" "[x]" == "x # []" no_notation Nil ("[]") and Cons (infixr "#" 65) hide_type list hide_const Nil Cons case_list hd tl set map list_all2 rec_list size_list list_all context early begin (*>*) datatype (set: 'a) list = null: Nil | Cons (hd: 'a) (tl: "'a list") for map: map rel: list_all2 pred: list_all where "tl Nil = Nil" text \ \noindent The types of the constants that appear in the specification are listed below. \medskip \begin{tabular}{@ {}ll@ {}} Constructors: & \Nil :: 'a list\ \\ & \Cons :: 'a \ 'a list \ 'a list\ \\ Discriminator: & \null :: 'a list \ bool\ \\ Selectors: & \hd :: 'a list \ 'a\ \\ & \tl :: 'a list \ 'a list\ \\ Set function: & \set :: 'a list \ 'a set\ \\ Map function: & \map :: ('a \ 'b) \ 'a list \ 'b list\ \\ Relator: & \list_all2 :: ('a \ 'b \ bool) \ 'a list \ 'b list \ bool\ \end{tabular} \medskip The discriminator \<^const>\null\ and the selectors \<^const>\hd\ and \<^const>\tl\ are characterized by the following conditional equations: % \[@{thm list.collapse(1)[of xs, no_vars]} \qquad @{thm list.collapse(2)[of xs, no_vars]}\] % For two-constructor datatypes, a single discriminator constant is sufficient. The discriminator associated with \<^const>\Cons\ is simply \<^term>\\xs. \ null xs\. The \keyw{where} clause at the end of the command specifies a default value for selectors applied to constructors on which they are not a priori specified. In the example, it is used to ensure that the tail of the empty list is itself (instead of being left unspecified). Because \<^const>\Nil\ is nullary, it is also possible to use \<^term>\\xs. xs = Nil\ as a discriminator. This is the default behavior if we omit the identifier \<^const>\null\ and the associated colon. Some users argue against this, because the mixture of constructors and selectors in the characteristic theorems can lead Isabelle's automation to switch between the constructor and the destructor view in surprising ways. The usual mixfix syntax annotations are available for both types and constructors. For example: \ (*<*) end (*>*) datatype ('a, 'b) prod (infixr "*" 20) = Pair 'a 'b text \\blankline\ datatype (set: 'a) list = null: Nil ("[]") | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65) for map: map rel: list_all2 pred: list_all text \ \noindent Incidentally, this is how the traditional syntax can be set up: \ syntax "_list" :: "args \ 'a list" ("[(_)]") text \\blankline\ translations "[x, xs]" == "x # [xs]" "[x]" == "x # []" subsection \Command Syntax \label{ssec:datatype-command-syntax}\ subsubsection \\keyw{datatype} \label{sssec:datatype}\ text \ \begin{matharray}{rcl} @{command_def "datatype"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command datatype} target? @{syntax dt_options}? @{syntax dt_spec} ; @{syntax_def dt_options}: '(' ((@{syntax plugins} | 'discs_sels') + ',') ')' ; @{syntax_def plugins}: 'plugins' ('only' | 'del') ':' (name +) ; @{syntax_def dt_spec}: (@{syntax dt_name} '=' (@{syntax dt_ctor} + '|') \ @{syntax map_rel_pred}? (@'where' (prop + '|'))? + @'and') ; @{syntax_def map_rel_pred}: @'for' ((('map' | 'rel' | 'pred') ':' name) +) \ \medskip \noindent The @{command datatype} command introduces a set of mutually recursive datatypes specified by their constructors. The syntactic entity \synt{target} can be used to specify a local -context (e.g., \(in linorder)\ @{cite "isabelle-isar-ref"}), +context (e.g., \(in linorder)\ \<^cite>\"isabelle-isar-ref"\), and \synt{prop} denotes a HOL proposition. The optional target is optionally followed by a combination of the following options: \begin{itemize} \setlength{\itemsep}{0pt} \item The \plugins\ option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. \item The \discs_sels\ option indicates that discriminators and selectors should be generated. The option is implicitly enabled if names are specified for discriminators or selectors. \end{itemize} The optional \keyw{where} clause specifies default values for selectors. Each proposition must be an equation of the form \un_D (C \) = \\, where \C\ is a constructor and \un_D\ is a selector. The left-hand sides of the datatype equations specify the name of the type to define, its type parameters, and additional information: \<^rail>\ @{syntax_def dt_name}: @{syntax tyargs}? name mixfix? ; @{syntax_def tyargs}: typefree | '(' (('dead' | name ':')? typefree + ',') ')' \ \medskip \noindent The syntactic entity \synt{name} denotes an identifier, \synt{mixfix} denotes the usual parenthesized mixfix notation, and \synt{typefree} denotes fixed type -variable (\<^typ>\'a\, \<^typ>\'b\, \ldots) @{cite "isabelle-isar-ref"}. +variable (\<^typ>\'a\, \<^typ>\'b\, \ldots) \<^cite>\"isabelle-isar-ref"\. The optional names preceding the type variables allow to override the default names of the set functions (\set\<^sub>1_t\, \ldots, \set\<^sub>m_t\). Type arguments can be marked as dead by entering \dead\ in front of the type variable (e.g., \(dead 'a)\); otherwise, they are live or dead (and a set function is generated or not) depending on where they occur in the right-hand sides of the definition. Declaring a type argument as dead can speed up the type definition but will prevent any later (co)recursion through that type argument. Inside a mutually recursive specification, all defined datatypes must mention exactly the same type variables in the same order. \<^rail>\ @{syntax_def dt_ctor}: (name ':')? name (@{syntax dt_ctor_arg} * ) mixfix? \ \medskip \noindent The main constituents of a constructor specification are the name of the constructor and the list of its argument types. An optional discriminator name can be supplied at the front. If discriminators are enabled (cf.\ the \discs_sels\ option) but no name is supplied, the default is \\x. x = C\<^sub>j\ for nullary constructors and \t.is_C\<^sub>j\ otherwise. \<^rail>\ @{syntax_def dt_ctor_arg}: type | '(' name ':' type ')' \ \medskip \noindent -The syntactic entity \synt{type} denotes a HOL type @{cite "isabelle-isar-ref"}. +The syntactic entity \synt{type} denotes a HOL type \<^cite>\"isabelle-isar-ref"\. In addition to the type of a constructor argument, it is possible to specify a name for the corresponding selector. The same selector name can be reused for arguments to several constructors as long as the arguments share the same type. If selectors are enabled (cf.\ the \discs_sels\ option) but no name is supplied, the default name is \un_C\<^sub>ji\. \ subsubsection \\keyw{datatype_compat} \label{sssec:datatype-compat}\ text \ \begin{matharray}{rcl} @{command_def "datatype_compat"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command datatype_compat} (name +) \ \medskip \noindent The @{command datatype_compat} command registers new-style datatypes as old-style datatypes and invokes the old-style plugins. For example: \ datatype_compat even_nat odd_nat text \\blankline\ ML \Old_Datatype_Data.get_info \<^theory> \<^type_name>\even_nat\\ text \ -The syntactic entity \synt{name} denotes an identifier @{cite "isabelle-isar-ref"}. +The syntactic entity \synt{name} denotes an identifier \<^cite>\"isabelle-isar-ref"\. The command is sometimes useful when migrating from the old datatype package to the new one. A few remarks concern nested recursive datatypes: \begin{itemize} \setlength{\itemsep}{0pt} \item The old-style, nested-as-mutual induction rule and recursor theorems are generated under their usual names but with ``\compat_\'' prefixed (e.g., \compat_tree.induct\, \compat_tree.inducts\, and \compat_tree.rec\). These theorems should be identical to the ones generated by the old datatype package, \emph{up to the order of the premises}---meaning that the subgoals generated by the \induct\ or \induction\ method may be in a different order than before. \item All types through which recursion takes place must be new-style datatypes or the function type. \end{itemize} \ subsection \Generated Constants \label{ssec:datatype-generated-constants}\ text \ Given a datatype \('a\<^sub>1, \, 'a\<^sub>m) t\ with $m$ live type variables and $n$ constructors \t.C\<^sub>1\, \ldots, \t.C\<^sub>n\, the following auxiliary constants are introduced: \medskip \begin{tabular}{@ {}ll@ {}} Case combinator: & \t.case_t\ (rendered using the familiar \case\--\of\ syntax) \\ Discriminators: & \t.is_C\<^sub>1\$, \ldots, $\t.is_C\<^sub>n\ \\ Selectors: & \t.un_C\<^sub>11\$, \ldots, $\t.un_C\<^sub>1k\<^sub>1\ \\ & \quad\vdots \\ & \t.un_C\<^sub>n1\$, \ldots, $\t.un_C\<^sub>nk\<^sub>n\ \\ Set functions: & \t.set\<^sub>1_t\, \ldots, \t.set\<^sub>m_t\ \\ Map function: & \t.map_t\ \\ Relator: & \t.rel_t\ \\ Recursor: & \t.rec_t\ \end{tabular} \medskip \noindent The discriminators and selectors are generated only if the \discs_sels\ option is enabled or if names are specified for discriminators or selectors. The set functions, map function, predicator, and relator are generated only if $m > 0$. In addition, some of the plugins introduce their own constants (Section~\ref{sec:selecting-plugins}). The case combinator, discriminators, and selectors are collectively called \emph{destructors}. The prefix ``\t.\'' is an optional component of the names and is normally hidden. \ subsection \Generated Theorems \label{ssec:datatype-generated-theorems}\ text \ The characteristic theorems generated by @{command datatype} are grouped in three broad categories: \begin{itemize} \setlength{\itemsep}{0pt} \item The \emph{free constructor theorems} (Section~\ref{sssec:free-constructor-theorems}) are properties of the constructors and destructors that can be derived for any freely generated type. Internally, the derivation is performed by @{command free_constructors}. \item The \emph{functorial theorems} (Section~\ref{sssec:functorial-theorems}) are properties of datatypes related to their BNF nature. \item The \emph{inductive theorems} (Section~\ref{sssec:inductive-theorems}) are properties of datatypes related to their inductive nature. \end{itemize} \noindent The full list of named theorems can be obtained by issuing the command \keyw{print_theorems} immediately after the datatype definition. This list includes theorems produced by plugins (Section~\ref{sec:selecting-plugins}), but normally excludes low-level theorems that reveal internal constructions. To make these accessible, add the line \ declare [[bnf_internals]] (*<*) declare [[bnf_internals = false]] (*>*) subsubsection \Free Constructor Theorems \label{sssec:free-constructor-theorems}\ (*<*) consts nonnull :: 'a (*>*) text \ The free constructor theorems are partitioned in three subgroups. The first subgroup of properties is concerned with the constructors. They are listed below for \<^typ>\'a list\: \begin{indentblock} \begin{description} \item[\t.\\hthm{inject} \[iff, induct_simp]\\rm:] ~ \\ @{thm list.inject[no_vars]} \item[\t.\\hthm{distinct} \[simp, induct_simp]\\rm:] ~ \\ @{thm list.distinct(1)[no_vars]} \\ @{thm list.distinct(2)[no_vars]} \item[\t.\\hthm{exhaust} \[cases t, case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust[no_vars]} \item[\t.\\hthm{nchotomy}\rm:] ~ \\ @{thm list.nchotomy[no_vars]} \end{description} \end{indentblock} \noindent In addition, these nameless theorems are registered as safe elimination rules: \begin{indentblock} \begin{description} \item[\t.\\hthm{distinct {\upshape[}THEN notE}\, elim!\\hthm{\upshape]}\rm:] ~ \\ @{thm list.distinct(1)[THEN notE, elim!, no_vars]} \\ @{thm list.distinct(2)[THEN notE, elim!, no_vars]} \end{description} \end{indentblock} \noindent The next subgroup is concerned with the case combinator: \begin{indentblock} \begin{description} \item[\t.\\hthm{case} \[simp, code]\\rm:] ~ \\ @{thm list.case(1)[no_vars]} \\ @{thm list.case(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{case_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.case_cong[no_vars]} \item[\t.\\hthm{case_cong_weak} \[cong]\\rm:] ~ \\ @{thm list.case_cong_weak[no_vars]} \item[\t.\\hthm{case_distrib}\rm:] ~ \\ @{thm list.case_distrib[no_vars]} \item[\t.\\hthm{split}\rm:] ~ \\ @{thm list.split[no_vars]} \item[\t.\\hthm{split_asm}\rm:] ~ \\ @{thm list.split_asm[no_vars]} \item[\t.\\hthm{splits} = \split split_asm\] \end{description} \end{indentblock} \noindent The third subgroup revolves around discriminators and selectors: \begin{indentblock} \begin{description} \item[\t.\\hthm{disc} \[simp]\\rm:] ~ \\ @{thm list.disc(1)[no_vars]} \\ @{thm list.disc(2)[no_vars]} \item[\t.\\hthm{discI}\rm:] ~ \\ @{thm list.discI(1)[no_vars]} \\ @{thm list.discI(2)[no_vars]} \item[\t.\\hthm{sel} \[simp, code]\\rm:] ~ \\ @{thm list.sel(1)[no_vars]} \\ @{thm list.sel(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{collapse} \[simp]\\rm:] ~ \\ @{thm list.collapse(1)[no_vars]} \\ @{thm list.collapse(2)[no_vars]} \\ The \[simp]\ attribute is exceptionally omitted for datatypes equipped with a single nullary constructor, because a property of the form \<^prop>\x = C\ is not suitable as a simplification rule. \item[\t.\\hthm{distinct_disc} \[dest]\\rm:] ~ \\ These properties are missing for \<^typ>\'a list\ because there is only one proper discriminator. If the datatype had been introduced with a second discriminator called \<^const>\nonnull\, they would have read as follows: \\[\jot] \<^prop>\null list \ \ nonnull list\ \\ \<^prop>\nonnull list \ \ null list\ \item[\t.\\hthm{exhaust_disc} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust_disc[no_vars]} \item[\t.\\hthm{exhaust_sel} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ @{thm list.exhaust_sel[no_vars]} \item[\t.\\hthm{expand}\rm:] ~ \\ @{thm list.expand[no_vars]} \item[\t.\\hthm{split_sel}\rm:] ~ \\ @{thm list.split_sel[no_vars]} \item[\t.\\hthm{split_sel_asm}\rm:] ~ \\ @{thm list.split_sel_asm[no_vars]} \item[\t.\\hthm{split_sels} = \split_sel split_sel_asm\] \item[\t.\\hthm{case_eq_if}\rm:] ~ \\ @{thm list.case_eq_if[no_vars]} \item[\t.\\hthm{disc_eq_case}\rm:] ~ \\ @{thm list.disc_eq_case(1)[no_vars]} \\ @{thm list.disc_eq_case(2)[no_vars]} \end{description} \end{indentblock} \noindent In addition, equational versions of \t.disc\ are registered with the \[code]\ attribute. The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \ subsubsection \Functorial Theorems \label{sssec:functorial-theorems}\ text \ The functorial theorems are generated for type constructors with at least one live type argument (e.g., \<^typ>\'a list\). They are partitioned in two subgroups. The first subgroup consists of properties involving the constructors or the destructors and either a set function, the map function, the predicator, or the relator: \begin{indentblock} \begin{description} \item[\t.\\hthm{case_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.case_transfer[no_vars]} \\ This property is generated by the \transfer\ plugin (Section~\ref{ssec:transfer}). %The \[transfer_rule]\ attribute is set by the \transfer\ plugin %(Section~\ref{ssec:transfer}). \item[\t.\\hthm{sel_transfer} \[transfer_rule]\\rm:] ~ \\ This property is missing for \<^typ>\'a list\ because there is no common selector to all constructors. \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}). \item[\t.\\hthm{ctr_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.ctr_transfer(1)[no_vars]} \\ @{thm list.ctr_transfer(2)[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) . \item[\t.\\hthm{disc_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.disc_transfer(1)[no_vars]} \\ @{thm list.disc_transfer(2)[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}). \item[\t.\\hthm{set} \[simp, code]\\rm:] ~ \\ @{thm list.set(1)[no_vars]} \\ @{thm list.set(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{set_cases} \[consumes 1, cases set: set\<^sub>i_t]\\rm:] ~ \\ @{thm list.set_cases[no_vars]} \item[\t.\\hthm{set_intros}\rm:] ~ \\ @{thm list.set_intros(1)[no_vars]} \\ @{thm list.set_intros(2)[no_vars]} \item[\t.\\hthm{set_sel}\rm:] ~ \\ @{thm list.set_sel[no_vars]} \item[\t.\\hthm{map} \[simp, code]\\rm:] ~ \\ @{thm list.map(1)[no_vars]} \\ @{thm list.map(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{map_disc_iff} \[simp]\\rm:] ~ \\ @{thm list.map_disc_iff[no_vars]} \item[\t.\\hthm{map_sel}\rm:] ~ \\ @{thm list.map_sel[no_vars]} \item[\t.\\hthm{pred_inject} \[simp]\\rm:] ~ \\ @{thm list.pred_inject(1)[no_vars]} \\ @{thm list.pred_inject(2)[no_vars]} \item[\t.\\hthm{rel_inject} \[simp]\\rm:] ~ \\ @{thm list.rel_inject(1)[no_vars]} \\ @{thm list.rel_inject(2)[no_vars]} \item[\t.\\hthm{rel_distinct} \[simp]\\rm:] ~ \\ @{thm list.rel_distinct(1)[no_vars]} \\ @{thm list.rel_distinct(2)[no_vars]} \item[\t.\\hthm{rel_intros}\rm:] ~ \\ @{thm list.rel_intros(1)[no_vars]} \\ @{thm list.rel_intros(2)[no_vars]} \item[\t.\\hthm{rel_cases} \[consumes 1, case_names t\<^sub>1 \ t\<^sub>m, cases pred]\\rm:] ~ \\ @{thm list.rel_cases[no_vars]} \item[\t.\\hthm{rel_sel}\rm:] ~ \\ @{thm list.rel_sel[no_vars]} \end{description} \end{indentblock} \noindent In addition, equational versions of \t.rel_inject\ and \rel_distinct\ are registered with the \[code]\ attribute. The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). The second subgroup consists of more abstract properties of the set functions, the map function, the predicator, and the relator: \begin{indentblock} \begin{description} \item[\t.\\hthm{inj_map}\rm:] ~ \\ @{thm list.inj_map[no_vars]} \item[\t.\\hthm{inj_map_strong}\rm:] ~ \\ @{thm list.inj_map_strong[no_vars]} \item[\t.\\hthm{map_comp}\rm:] ~ \\ @{thm list.map_comp[no_vars]} \item[\t.\\hthm{map_cong0}\rm:] ~ \\ @{thm list.map_cong0[no_vars]} \item[\t.\\hthm{map_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.map_cong[no_vars]} \item[\t.\\hthm{map_cong_pred}\rm:] ~ \\ @{thm list.map_cong_pred[no_vars]} \item[\t.\\hthm{map_cong_simp}\rm:] ~ \\ @{thm list.map_cong_simp[no_vars]} \item[\t.\\hthm{map_id0}\rm:] ~ \\ @{thm list.map_id0[no_vars]} \item[\t.\\hthm{map_id}\rm:] ~ \\ @{thm list.map_id[no_vars]} \item[\t.\\hthm{map_ident}\rm:] ~ \\ @{thm list.map_ident[no_vars]} \item[\t.\\hthm{map_ident_strong}\rm:] ~ \\ @{thm list.map_ident_strong[no_vars]} \item[\t.\\hthm{map_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.map_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \item[\t.\\hthm{pred_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.pred_cong[no_vars]} \item[\t.\\hthm{pred_cong_simp}\rm:] ~ \\ @{thm list.pred_cong_simp[no_vars]} \item[\t.\\hthm{pred_map}\rm:] ~ \\ @{thm list.pred_map[no_vars]} \item[\t.\\hthm{pred_mono} \[mono]\\rm:] ~ \\ @{thm list.pred_mono[no_vars]} \item[\t.\\hthm{pred_mono_strong}\rm:] ~ \\ @{thm list.pred_mono_strong[no_vars]} \item[\t.\\hthm{pred_rel}\rm:] ~ \\ @{thm list.pred_rel[no_vars]} \item[\t.\\hthm{pred_set}\rm:] ~ \\ @{thm list.pred_set[no_vars]} \item[\t.\\hthm{pred_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.pred_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \item[\t.\\hthm{pred_True}\rm:] ~ \\ @{thm list.pred_True[no_vars]} \item[\t.\\hthm{set_map}\rm:] ~ \\ @{thm list.set_map[no_vars]} \item[\t.\\hthm{set_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.set_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \item[\t.\\hthm{rel_compp} \[relator_distr]\\rm:] ~ \\ @{thm list.rel_compp[no_vars]} \\ The \[relator_distr]\ attribute is set by the \lifting\ plugin (Section~\ref{ssec:lifting}). \item[\t.\\hthm{rel_conversep}\rm:] ~ \\ @{thm list.rel_conversep[no_vars]} \item[\t.\\hthm{rel_eq}\rm:] ~ \\ @{thm list.rel_eq[no_vars]} \item[\t.\\hthm{rel_eq_onp}\rm:] ~ \\ @{thm list.rel_eq_onp[no_vars]} \item[\t.\\hthm{rel_flip}\rm:] ~ \\ @{thm list.rel_flip[no_vars]} \item[\t.\\hthm{rel_map}\rm:] ~ \\ @{thm list.rel_map(1)[no_vars]} \\ @{thm list.rel_map(2)[no_vars]} \item[\t.\\hthm{rel_mono} \[mono, relator_mono]\\rm:] ~ \\ @{thm list.rel_mono[no_vars]} \\ The \[relator_mono]\ attribute is set by the \lifting\ plugin (Section~\ref{ssec:lifting}). \item[\t.\\hthm{rel_mono_strong}\rm:] ~ \\ @{thm list.rel_mono_strong[no_vars]} \item[\t.\\hthm{rel_cong} \[fundef_cong]\\rm:] ~ \\ @{thm list.rel_cong[no_vars]} \item[\t.\\hthm{rel_cong_simp}\rm:] ~ \\ @{thm list.rel_cong_simp[no_vars]} \item[\t.\\hthm{rel_refl}\rm:] ~ \\ @{thm list.rel_refl[no_vars]} \item[\t.\\hthm{rel_refl_strong}\rm:] ~ \\ @{thm list.rel_refl_strong[no_vars]} \item[\t.\\hthm{rel_reflp}\rm:] ~ \\ @{thm list.rel_reflp[no_vars]} \item[\t.\\hthm{rel_symp}\rm:] ~ \\ @{thm list.rel_symp[no_vars]} \item[\t.\\hthm{rel_transp}\rm:] ~ \\ @{thm list.rel_transp[no_vars]} \item[\t.\\hthm{rel_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.rel_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} \end{indentblock} \ subsubsection \Inductive Theorems \label{sssec:inductive-theorems}\ text \ The inductive theorems are as follows: \begin{indentblock} \begin{description} \item[\t.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n, induct t]\\rm:] ~ \\ @{thm list.induct[no_vars]} \item[\t.\\hthm{rel_induct} \[case_names C\<^sub>1 \ C\<^sub>n, induct pred]\\rm:] ~ \\ @{thm list.rel_induct[no_vars]} \item[\begin{tabular}{@ {}l@ {}} \t\<^sub>1_\_t\<^sub>m.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm: \\ \t\<^sub>1_\_t\<^sub>m.\\hthm{rel_induct} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm: \\ \end{tabular}] ~ \\ Given $m > 1$ mutually recursive datatypes, this induction rule can be used to prove $m$ properties simultaneously. \item[\t.\\hthm{rec} \[simp, code]\\rm:] ~ \\ @{thm list.rec(1)[no_vars]} \\ @{thm list.rec(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{rec_o_map}\rm:] ~ \\ @{thm list.rec_o_map[no_vars]} \item[\t.\\hthm{rec_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm list.rec_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} \end{indentblock} \noindent For convenience, @{command datatype} also provides the following collection: \begin{indentblock} \begin{description} \item[\t.\\hthm{simps}] = \t.inject\ \t.distinct\ \t.case\ \t.rec\ \t.map\ \t.rel_inject\ \\ \t.rel_distinct\ \t.set\ \end{description} \end{indentblock} \ subsection \Proof Method \label{ssec:datatype-proof-method}\ subsubsection \\textit{countable_datatype} \label{sssec:datatype-compat}\ text \ The theory \<^file>\~~/src/HOL/Library/Countable.thy\ provides a proof method called @{method countable_datatype} that can be used to prove the countability of many datatypes, building on the countability of the types appearing in their definitions and of any type arguments. For example: \ instance list :: (countable) countable by countable_datatype subsection \Antiquotation \label{ssec:datatype-antiquotation}\ subsubsection \\textit{datatype} \label{sssec:datatype-datatype}\ text \ The \textit{datatype} antiquotation, written \texttt{\char`\\\char`<\char`^}\verb|datatype>|\guilsinglleft\textit{t}\guilsinglright{} or \texttt{@}\verb|{datatype| \textit{t}\verb|}|, where \textit{t} is a type name, expands to \LaTeX{} code for the definition of the datatype, with each constructor listed with its argument types. For example, if \textit{t} is @{type option}: \begin{quote} \<^datatype>\option\ \end{quote} \ subsection \Compatibility Issues \label{ssec:datatype-compatibility-issues}\ text \ The command @{command datatype} has been designed to be highly compatible with the old, pre-Isabelle2015 command, to ease migration. There are nonetheless a few incompatibilities that may arise when porting: \begin{itemize} \setlength{\itemsep}{0pt} \item \emph{The Standard ML interfaces are different.} Tools and extensions written to call the old ML interfaces will need to be adapted to the new interfaces. The \BNF_LFP_Compat\ structure provides convenience functions that simulate the old interfaces in terms of the new ones. \item \emph{The recursor \rec_t\ has a different signature for nested recursive datatypes.} In the old package, nested recursion through non-functions was internally reduced to mutual recursion. This reduction was visible in the type of the recursor, used by \keyw{primrec}. Recursion through functions was handled specially. In the new package, nested recursion (for functions and non-functions) is handled in a more modular fashion. The old-style recursor can be generated on demand using @{command primrec} if the recursion is via new-style datatypes, as explained in Section~\ref{sssec:primrec-nested-as-mutual-recursion}, or using @{command datatype_compat}. \item \emph{Accordingly, the induction rule is different for nested recursive datatypes.} Again, the old-style induction rule can be generated on demand using @{command primrec} if the recursion is via new-style datatypes, as explained in Section~\ref{sssec:primrec-nested-as-mutual-recursion}, or using @{command datatype_compat}. For recursion through functions, the old-style induction rule can be obtained by applying the \[unfolded all_mem_range]\ attribute on \t.induct\. \item \emph{The \<^const>\size\ function has a slightly different definition.} The new function returns \1\ instead of \0\ for some nonrecursive constructors. This departure from the old behavior made it possible to implement \<^const>\size\ in terms of the generic function \t.size_t\. Moreover, the new function considers nested occurrences of a value, in the nested recursive case. The old behavior can be obtained by disabling the \size\ plugin (Section~\ref{sec:selecting-plugins}) and instantiating the \size\ type class manually. \item \emph{The internal constructions are completely different.} Proof texts that unfold the definition of constants introduced by the old command will be difficult to port. \item \emph{Some constants and theorems have different names.} For non-mutually recursive datatypes, the alias \t.inducts\ for \t.induct\ is no longer generated. For $m > 1$ mutually recursive datatypes, \rec_t\<^sub>1_\_t\<^sub>m_i\ has been renamed \rec_t\<^sub>i\ for each \i \ {1, \, m}\, \t\<^sub>1_\_t\<^sub>m.inducts(i)\ has been renamed \t\<^sub>i.induct\ for each \i \ {1, \, m}\, and the collection \t\<^sub>1_\_t\<^sub>m.size\ (generated by the \size\ plugin, Section~\ref{ssec:size}) has been divided into \t\<^sub>1.size\, \ldots, \t\<^sub>m.size\. \item \emph{The \t.simps\ collection has been extended.} Previously available theorems are available at the same index as before. \item \emph{Variables in generated properties have different names.} This is rarely an issue, except in proof texts that refer to variable names in the \[where \]\ attribute. The solution is to use the more robust \[of \]\ syntax. \end{itemize} \ section \Defining Primitively Recursive Functions \label{sec:defining-primitively-recursive-functions}\ text \ Recursive functions over datatypes can be specified using the @{command primrec} command, which supports primitive recursion, or using the \keyw{fun}, \keyw{function}, and \keyw{partial_function} commands. In this tutorial, the focus is on @{command primrec}; \keyw{fun} and \keyw{function} are described in -a separate tutorial @{cite "isabelle-function"}. +a separate tutorial \<^cite>\"isabelle-function"\. Because it is restricted to primitive recursion, @{command primrec} is less powerful than \keyw{fun} and \keyw{function}. However, there are primitively recursive specifications (e.g., based on infinitely branching or mutually recursive datatypes) for which \keyw{fun}'s termination check fails. It is also good style to use the simpler @{command primrec} mechanism when it works, both as an optimization and as documentation. \ subsection \Introductory Examples \label{ssec:primrec-introductory-examples}\ text \ Primitive recursion is illustrated through concrete examples based on the datatypes defined in Section~\ref{ssec:datatype-introductory-examples}. More examples can be found in the directory \<^dir>\~~/src/HOL/Datatype_Examples\. \ subsubsection \Nonrecursive Types \label{sssec:primrec-nonrecursive-types}\ text \ Primitive recursion removes one layer of constructors on the left-hand side in each equation. For example: \ primrec (nonexhaustive) bool_of_trool :: "trool \ bool" where "bool_of_trool Faalse \ False" | "bool_of_trool Truue \ True" text \\blankline\ primrec the_list :: "'a option \ 'a list" where "the_list None = []" | "the_list (Some a) = [a]" text \\blankline\ primrec the_default :: "'a \ 'a option \ 'a" where "the_default d None = d" | "the_default _ (Some a) = a" text \\blankline\ primrec mirrror :: "('a, 'b, 'c) triple \ ('c, 'b, 'a) triple" where "mirrror (Triple a b c) = Triple c b a" text \ \noindent The equations can be specified in any order, and it is acceptable to leave out some cases, which are then unspecified. Pattern matching on the left-hand side is restricted to a single datatype, which must correspond to the same argument in all equations. \ subsubsection \Simple Recursion \label{sssec:primrec-simple-recursion}\ text \ For simple recursive types, recursive calls on a constructor argument are allowed on the right-hand side: \ primrec replicate :: "nat \ 'a \ 'a list" where "replicate Zero _ = []" | "replicate (Succ n) x = x # replicate n x" text \\blankline\ primrec (nonexhaustive) at :: "'a list \ nat \ 'a" where "at (x # xs) j = (case j of Zero \ x | Succ j' \ at xs j')" text \\blankline\ primrec (*<*)(in early) (transfer) (*>*)tfold :: "('a \ 'b \ 'b) \ ('a, 'b) tlist \ 'b" where "tfold _ (TNil y) = y" | "tfold f (TCons x xs) = f x (tfold f xs)" text \ \noindent Pattern matching is only available for the argument on which the recursion takes place. Fortunately, it is easy to generate pattern-maching equations using the @{command simps_of_case} command provided by the theory \<^file>\~~/src/HOL/Library/Simps_Case_Conv.thy\. \ simps_of_case at_simps_alt: at.simps text \ This generates the lemma collection @{thm [source] at_simps_alt}: % \[@{thm at_simps_alt(1)[no_vars]} \qquad @{thm at_simps_alt(2)[no_vars]}\] % The next example is defined using \keyw{fun} to escape the syntactic restrictions imposed on primitively recursive functions: \ fun at_least_two :: "nat \ bool" where "at_least_two (Succ (Succ _)) \ True" | "at_least_two _ \ False" subsubsection \Mutual Recursion \label{sssec:primrec-mutual-recursion}\ text \ The syntax for mutually recursive functions over mutually recursive datatypes is straightforward: \ primrec nat_of_even_nat :: "even_nat \ nat" and nat_of_odd_nat :: "odd_nat \ nat" where "nat_of_even_nat Even_Zero = Zero" | "nat_of_even_nat (Even_Succ n) = Succ (nat_of_odd_nat n)" | "nat_of_odd_nat (Odd_Succ n) = Succ (nat_of_even_nat n)" text \\blankline\ primrec eval\<^sub>e :: "('a \ int) \ ('b \ int) \ ('a, 'b) exp \ int" and eval\<^sub>t :: "('a \ int) \ ('b \ int) \ ('a, 'b) trm \ int" and eval\<^sub>f :: "('a \ int) \ ('b \ int) \ ('a, 'b) fct \ int" where "eval\<^sub>e \ \ (Term t) = eval\<^sub>t \ \ t" | "eval\<^sub>e \ \ (Sum t e) = eval\<^sub>t \ \ t + eval\<^sub>e \ \ e" | "eval\<^sub>t \ \ (Factor f) = eval\<^sub>f \ \ f" | "eval\<^sub>t \ \ (Prod f t) = eval\<^sub>f \ \ f + eval\<^sub>t \ \ t" | "eval\<^sub>f \ _ (Const a) = \ a" | "eval\<^sub>f _ \ (Var b) = \ b" | "eval\<^sub>f \ \ (Expr e) = eval\<^sub>e \ \ e" text \ \noindent Mutual recursion is possible within a single type, using \keyw{fun}: \ fun even :: "nat \ bool" and odd :: "nat \ bool" where "even Zero = True" | "even (Succ n) = odd n" | "odd Zero = False" | "odd (Succ n) = even n" subsubsection \Nested Recursion \label{sssec:primrec-nested-recursion}\ text \ In a departure from the old datatype package, nested recursion is normally handled via the map functions of the nesting type constructors. For example, recursive calls are lifted to lists using \<^const>\map\: \ (*<*) datatype 'a tree\<^sub>f\<^sub>f = Node\<^sub>f\<^sub>f (lbl\<^sub>f\<^sub>f: 'a) (sub\<^sub>f\<^sub>f: "'a tree\<^sub>f\<^sub>f list") (*>*) primrec at\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \ nat list \ 'a" where "at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js = (case js of [] \ a | j # js' \ at (map (\t. at\<^sub>f\<^sub>f t js') ts) j)" text \ \noindent The next example features recursion through the \option\ type. Although \option\ is not a new-style datatype, it is registered as a BNF with the map function \<^const>\map_option\: \ primrec (*<*)(in early) (*>*)sum_btree :: "('a::{zero,plus}) btree \ 'a" where "sum_btree (BNode a lt rt) = a + the_default 0 (map_option sum_btree lt) + the_default 0 (map_option sum_btree rt)" text \ \noindent The same principle applies for arbitrary type constructors through which recursion is possible. Notably, the map function for the function type (\\\) is simply composition (\(\)\): \ primrec (*<*)(in early) (*>*)relabel_ft :: "('a \ 'a) \ 'a ftree \ 'a ftree" where "relabel_ft f (FTLeaf x) = FTLeaf (f x)" | "relabel_ft f (FTNode g) = FTNode (relabel_ft f \ g)" text \ \noindent For convenience, recursion through functions can also be expressed using $\lambda$-abstractions and function application rather than through composition. For example: \ primrec relabel_ft :: "('a \ 'a) \ 'a ftree \ 'a ftree" where "relabel_ft f (FTLeaf x) = FTLeaf (f x)" | "relabel_ft f (FTNode g) = FTNode (\x. relabel_ft f (g x))" text \\blankline\ primrec (nonexhaustive) subtree_ft :: "'a \ 'a ftree \ 'a ftree" where "subtree_ft x (FTNode g) = g x" text \ \noindent For recursion through curried $n$-ary functions, $n$ applications of \<^term>\(\)\ are necessary. The examples below illustrate the case where $n = 2$: \ datatype 'a ftree2 = FTLeaf2 'a | FTNode2 "'a \ 'a \ 'a ftree2" text \\blankline\ primrec (*<*)(in early) (*>*)relabel_ft2 :: "('a \ 'a) \ 'a ftree2 \ 'a ftree2" where "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" | "relabel_ft2 f (FTNode2 g) = FTNode2 ((\) ((\) (relabel_ft2 f)) g)" text \\blankline\ primrec relabel_ft2 :: "('a \ 'a) \ 'a ftree2 \ 'a ftree2" where "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" | "relabel_ft2 f (FTNode2 g) = FTNode2 (\x y. relabel_ft2 f (g x y))" text \\blankline\ primrec (nonexhaustive) subtree_ft2 :: "'a \ 'a \ 'a ftree2 \ 'a ftree2" where "subtree_ft2 x y (FTNode2 g) = g x y" text \ For any datatype featuring nesting, the predicator can be used instead of the map function, typically when defining predicates. For example: \ primrec increasing_tree :: "int \ int tree\<^sub>f\<^sub>f \ bool" where "increasing_tree m (Node\<^sub>f\<^sub>f n ts) \ n \ m \ list_all (increasing_tree (n + 1)) ts" subsubsection \Nested-as-Mutual Recursion \label{sssec:primrec-nested-as-mutual-recursion}\ (*<*) locale n2m begin (*>*) text \ For compatibility with the old package, but also because it is sometimes convenient in its own right, it is possible to treat nested recursive datatypes as mutually recursive ones if the recursion takes place though new-style datatypes. For example: \ primrec (nonexhaustive) at\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f \ nat list \ 'a" and ats\<^sub>f\<^sub>f :: "'a tree\<^sub>f\<^sub>f list \ nat \ nat list \ 'a" where "at\<^sub>f\<^sub>f (Node\<^sub>f\<^sub>f a ts) js = (case js of [] \ a | j # js' \ ats\<^sub>f\<^sub>f ts j js')" | "ats\<^sub>f\<^sub>f (t # ts) j = (case j of Zero \ at\<^sub>f\<^sub>f t | Succ j' \ ats\<^sub>f\<^sub>f ts j')" text \ \noindent Appropriate induction rules are generated as @{thm [source] at\<^sub>f\<^sub>f.induct}, @{thm [source] ats\<^sub>f\<^sub>f.induct}, and @{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}. The induction rules and the underlying recursors are generated dynamically and are kept in a cache to speed up subsequent definitions. Here is a second example: \ primrec sum_btree :: "('a::{zero,plus}) btree \ 'a" and sum_btree_option :: "'a btree option \ 'a" where "sum_btree (BNode a lt rt) = a + sum_btree_option lt + sum_btree_option rt" | "sum_btree_option None = 0" | "sum_btree_option (Some t) = sum_btree t" text \ % * can pretend a nested type is mutually recursive (if purely inductive) % * avoids the higher-order map % * e.g. % * this can always be avoided; % * e.g. in our previous example, we first mapped the recursive % calls, then we used a generic at function to retrieve the result % % * there's no hard-and-fast rule of when to use one or the other, % just like there's no rule when to use fold and when to use % primrec % % * higher-order approach, considering nesting as nesting, is more % compositional -- e.g. we saw how we could reuse an existing polymorphic % at or the_default, whereas \<^const>\ats\<^sub>f\<^sub>f\ is much more specific % % * but: % * is perhaps less intuitive, because it requires higher-order thinking % * may seem inefficient, and indeed with the code generator the % mutually recursive version might be nicer % * is somewhat indirect -- must apply a map first, then compute a result % (cannot mix) % * the auxiliary functions like \<^const>\ats\<^sub>f\<^sub>f\ are sometimes useful in own right % % * impact on automation unclear % \ (*<*) end (*>*) subsection \Command Syntax \label{ssec:primrec-command-syntax}\ subsubsection \\keyw{primrec} \label{sssec:primrec}\ text \ \begin{matharray}{rcl} @{command_def "primrec"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command primrec} target? @{syntax pr_options}? fixes \ @'where' (@{syntax pr_equation} + '|') ; @{syntax_def pr_options}: '(' ((@{syntax plugins} | 'nonexhaustive' | 'transfer') + ',') ')' ; @{syntax_def pr_equation}: thmdecl? prop \ \medskip \noindent The @{command primrec} command introduces a set of mutually recursive functions over datatypes. The syntactic entity \synt{target} can be used to specify a local context, \synt{fixes} denotes a list of names with optional type signatures, \synt{thmdecl} denotes an optional name for the formula that follows, and -\synt{prop} denotes a HOL proposition @{cite "isabelle-isar-ref"}. +\synt{prop} denotes a HOL proposition \<^cite>\"isabelle-isar-ref"\. The optional target is optionally followed by a combination of the following options: \begin{itemize} \setlength{\itemsep}{0pt} \item The \plugins\ option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. \item The \nonexhaustive\ option indicates that the functions are not necessarily specified for all constructors. It can be used to suppress the warning that is normally emitted when some constructors are missing. \item The \transfer\ option indicates that an unconditional transfer rule should be generated and proved \by transfer_prover\. The \[transfer_rule]\ attribute is set on the generated theorem. \end{itemize} %%% TODO: elaborate on format of the equations %%% TODO: elaborate on mutual and nested-to-mutual \ subsection \Generated Theorems \label{ssec:primrec-generated-theorems}\ (*<*) context early begin (*>*) text \ The @{command primrec} command generates the following properties (listed for \<^const>\tfold\): \begin{indentblock} \begin{description} \item[\f.\\hthm{simps} \[simp, code]\\rm:] ~ \\ @{thm tfold.simps(1)[no_vars]} \\ @{thm tfold.simps(2)[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\f.\\hthm{transfer} \[transfer_rule]\\rm:] ~ \\ @{thm tfold.transfer[no_vars]} \\ This theorem is generated by the \transfer\ plugin (Section~\ref{ssec:transfer}) for functions declared with the \transfer\ option enabled. \item[\f.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ This induction rule is generated for nested-as-mutual recursive functions (Section~\ref{sssec:primrec-nested-as-mutual-recursion}). \item[\f\<^sub>1_\_f\<^sub>m.\\hthm{induct} \[case_names C\<^sub>1 \ C\<^sub>n]\\rm:] ~ \\ This induction rule is generated for nested-as-mutual recursive functions (Section~\ref{sssec:primrec-nested-as-mutual-recursion}). Given $m > 1$ mutually recursive functions, this rule can be used to prove $m$ properties simultaneously. \end{description} \end{indentblock} \ (*<*) end (*>*) subsection \Recursive Default Values for Selectors \label{ssec:primrec-recursive-default-values-for-selectors}\ text \ A datatype selector \un_D\ can have a default value for each constructor on which it is not otherwise specified. Occasionally, it is useful to have the default value be defined recursively. This leads to a chicken-and-egg situation, because the datatype is not introduced yet at the moment when the selectors are introduced. Of course, we can always define the selectors manually afterward, but we then have to state and prove all the characteristic theorems ourselves instead of letting the package do it. Fortunately, there is a workaround that relies on overloading to relieve us from the tedium of manual derivations: \begin{enumerate} \setlength{\itemsep}{0pt} \item Introduce a fully unspecified constant \un_D\<^sub>0 :: 'a\ using @{command consts}. \item Define the datatype, specifying \un_D\<^sub>0\ as the selector's default value. \item Define the behavior of \un_D\<^sub>0\ on values of the newly introduced datatype using the \keyw{overloading} command. \item Derive the desired equation on \un_D\ from the characteristic equations for \un_D\<^sub>0\. \end{enumerate} \noindent The following example illustrates this procedure: \ (*<*) hide_const termi (*>*) consts termi\<^sub>0 :: 'a text \\blankline\ datatype ('a, 'b) tlist = TNil (termi: 'b) | TCons (thd: 'a) (ttl: "('a, 'b) tlist") where "ttl (TNil y) = TNil y" | "termi (TCons _ xs) = termi\<^sub>0 xs" text \\blankline\ overloading termi\<^sub>0 \ "termi\<^sub>0 :: ('a, 'b) tlist \ 'b" begin primrec termi\<^sub>0 :: "('a, 'b) tlist \ 'b" where "termi\<^sub>0 (TNil y) = y" | "termi\<^sub>0 (TCons x xs) = termi\<^sub>0 xs" end text \\blankline\ lemma termi_TCons[simp]: "termi (TCons x xs) = termi xs" by (cases xs) auto subsection \Compatibility Issues \label{ssec:primrec-compatibility-issues}\ text \ The command @{command primrec}'s behavior on new-style datatypes has been designed to be highly compatible with that for old, pre-Isabelle2015 datatypes, to ease migration. There is nonetheless at least one incompatibility that may arise when porting to the new package: \begin{itemize} \setlength{\itemsep}{0pt} \item \emph{Some theorems have different names.} For $m > 1$ mutually recursive functions, \f\<^sub>1_\_f\<^sub>m.simps\ has been broken down into separate subcollections \f\<^sub>i.simps\. \end{itemize} \ section \Defining Codatatypes \label{sec:defining-codatatypes}\ text \ Codatatypes can be specified using the @{command codatatype} command. The command is first illustrated through concrete examples featuring different flavors of corecursion. More examples can be found in the directory \<^dir>\~~/src/HOL/Datatype_Examples\. The \emph{Archive of Formal Proofs} also -includes some useful codatatypes, notably for lazy lists @{cite -"lochbihler-2010"}. +includes some useful codatatypes, notably for lazy lists \<^cite>\"lochbihler-2010"\. \ subsection \Introductory Examples \label{ssec:codatatype-introductory-examples}\ subsubsection \Simple Corecursion \label{sssec:codatatype-simple-corecursion}\ text \ Non-corecursive codatatypes coincide with the corresponding datatypes, so they are rarely used in practice. \emph{Corecursive codatatypes} have the same syntax as recursive datatypes, except for the command name. For example, here is the definition of lazy lists: \ codatatype (lset: 'a) llist = lnull: LNil | LCons (lhd: 'a) (ltl: "'a llist") for map: lmap rel: llist_all2 pred: llist_all where "ltl LNil = LNil" text \ \noindent Lazy lists can be infinite, such as \LCons 0 (LCons 0 (\))\ and \LCons 0 (LCons 1 (LCons 2 (\)))\. Here is a related type, that of infinite streams: \ codatatype (sset: 'a) stream = SCons (shd: 'a) (stl: "'a stream") for map: smap rel: stream_all2 text \ \noindent Another interesting type that can be defined as a codatatype is that of the extended natural numbers: \ codatatype enat = EZero | ESucc enat text \ \noindent This type has exactly one infinite element, \ESucc (ESucc (ESucc (\)))\, that represents $\infty$. In addition, it has finite values of the form \ESucc (\ (ESucc EZero)\)\. Here is an example with many constructors: \ codatatype 'a process = Fail | Skip (cont: "'a process") | Action (prefix: 'a) (cont: "'a process") | Choice (left: "'a process") (right: "'a process") text \ \noindent Notice that the \<^const>\cont\ selector is associated with both \<^const>\Skip\ and \<^const>\Action\. \ subsubsection \Mutual Corecursion \label{sssec:codatatype-mutual-corecursion}\ text \ \noindent The example below introduces a pair of \emph{mutually corecursive} types: \ codatatype even_enat = Even_EZero | Even_ESucc odd_enat and odd_enat = Odd_ESucc even_enat subsubsection \Nested Corecursion \label{sssec:codatatype-nested-corecursion}\ text \ \noindent The next examples feature \emph{nested corecursion}: \ codatatype 'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i (lbl\<^sub>i\<^sub>i: 'a) (sub\<^sub>i\<^sub>i: "'a tree\<^sub>i\<^sub>i llist") text \\blankline\ codatatype 'a tree\<^sub>i\<^sub>s = Node\<^sub>i\<^sub>s (lbl\<^sub>i\<^sub>s: 'a) (sub\<^sub>i\<^sub>s: "'a tree\<^sub>i\<^sub>s fset") text \\blankline\ codatatype 'a sm = SM (accept: bool) (trans: "'a \ 'a sm") subsection \Command Syntax \label{ssec:codatatype-command-syntax}\ subsubsection \\keyw{codatatype} \label{sssec:codatatype}\ text \ \begin{matharray}{rcl} @{command_def "codatatype"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command codatatype} target? @{syntax dt_options}? @{syntax dt_spec} \ \medskip \noindent Definitions of codatatypes have almost exactly the same syntax as for datatypes (Section~\ref{ssec:datatype-command-syntax}). The \discs_sels\ option is superfluous because discriminators and selectors are always generated for codatatypes. \ subsection \Generated Constants \label{ssec:codatatype-generated-constants}\ text \ Given a codatatype \('a\<^sub>1, \, 'a\<^sub>m) t\ with $m > 0$ live type variables and $n$ constructors \t.C\<^sub>1\, \ldots, \t.C\<^sub>n\, the same auxiliary constants are generated as for datatypes (Section~\ref{ssec:datatype-generated-constants}), except that the recursor is replaced by a dual concept: \medskip \begin{tabular}{@ {}ll@ {}} Corecursor: & \t.corec_t\ \end{tabular} \ subsection \Generated Theorems \label{ssec:codatatype-generated-theorems}\ text \ The characteristic theorems generated by @{command codatatype} are grouped in three broad categories: \begin{itemize} \setlength{\itemsep}{0pt} \item The \emph{free constructor theorems} (Section~\ref{sssec:free-constructor-theorems}) are properties of the constructors and destructors that can be derived for any freely generated type. \item The \emph{functorial theorems} (Section~\ref{sssec:functorial-theorems}) are properties of datatypes related to their BNF nature. \item The \emph{coinductive theorems} (Section~\ref{sssec:coinductive-theorems}) are properties of datatypes related to their coinductive nature. \end{itemize} \noindent The first two categories are exactly as for datatypes. \ subsubsection \Coinductive Theorems \label{sssec:coinductive-theorems}\ text \ The coinductive theorems are listed below for \<^typ>\'a llist\: \begin{indentblock} \begin{description} \item[\begin{tabular}{@ {}l@ {}} \t.\\hthm{coinduct} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\t.\\hthm{coinduct} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n, coinduct t]\\rm: \end{tabular}] ~ \\ @{thm llist.coinduct[no_vars]} \item[\begin{tabular}{@ {}l@ {}} \t.\\hthm{coinduct_strong} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\t.\\hthm{coinduct_strong} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \end{tabular}] ~ \\ @{thm llist.coinduct_strong[no_vars]} \item[\begin{tabular}{@ {}l@ {}} \t.\\hthm{rel_coinduct} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\t.\\hthm{rel_coinduct} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n, coinduct pred]\\rm: \end{tabular}] ~ \\ @{thm llist.rel_coinduct[no_vars]} \item[\begin{tabular}{@ {}l@ {}} \t\<^sub>1_\_t\<^sub>m.\\hthm{coinduct} \[case_names t\<^sub>1 \ t\<^sub>m, case_conclusion D\<^sub>1 \ D\<^sub>n]\ \\ \t\<^sub>1_\_t\<^sub>m.\\hthm{coinduct_strong} \[case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\t\<^sub>1_\_t\<^sub>m.\\hthm{coinduct_strong} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \\ \t\<^sub>1_\_t\<^sub>m.\\hthm{rel_coinduct} \[case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\t\<^sub>1_\_t\<^sub>m.\\hthm{rel_coinduct} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \\ \end{tabular}] ~ \\ Given $m > 1$ mutually corecursive codatatypes, these coinduction rules can be used to prove $m$ properties simultaneously. \item[\begin{tabular}{@ {}l@ {}} \t\<^sub>1_\_t\<^sub>m.\\hthm{set_induct} \[case_names C\<^sub>1 \ C\<^sub>n,\ \\ \phantom{\t\<^sub>1_\_t\<^sub>m.\\hthm{set_induct} \[\}\induct set: set\<^sub>j_t\<^sub>1, \, induct set: set\<^sub>j_t\<^sub>m]\\rm: \\ \end{tabular}] ~ \\ @{thm llist.set_induct[no_vars]} \\ If $m = 1$, the attribute \[consumes 1]\ is generated as well. \item[\t.\\hthm{corec}\rm:] ~ \\ @{thm llist.corec(1)[no_vars]} \\ @{thm llist.corec(2)[no_vars]} \item[\t.\\hthm{corec_code} \[code]\\rm:] ~ \\ @{thm llist.corec_code[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\t.\\hthm{corec_disc}\rm:] ~ \\ @{thm llist.corec_disc(1)[no_vars]} \\ @{thm llist.corec_disc(2)[no_vars]} \item[\t.\\hthm{corec_disc_iff} \[simp]\\rm:] ~ \\ @{thm llist.corec_disc_iff(1)[no_vars]} \\ @{thm llist.corec_disc_iff(2)[no_vars]} \item[\t.\\hthm{corec_sel} \[simp]\\rm:] ~ \\ @{thm llist.corec_sel(1)[no_vars]} \\ @{thm llist.corec_sel(2)[no_vars]} \item[\t.\\hthm{map_o_corec}\rm:] ~ \\ @{thm llist.map_o_corec[no_vars]} \item[\t.\\hthm{corec_transfer} \[transfer_rule]\\rm:] ~ \\ @{thm llist.corec_transfer[no_vars]} \\ The \[transfer_rule]\ attribute is set by the \transfer\ plugin (Section~\ref{ssec:transfer}) for type constructors with no dead type arguments. \end{description} \end{indentblock} \noindent For convenience, @{command codatatype} also provides the following collection: \begin{indentblock} \begin{description} \item[\t.\\hthm{simps}] = \t.inject\ \t.distinct\ \t.case\ \t.corec_disc_iff\ \t.corec_sel\ \\ \t.map\ \t.rel_inject\ \t.rel_distinct\ \t.set\ \end{description} \end{indentblock} \ subsection \Antiquotation \label{ssec:codatatype-antiquotation}\ subsubsection \\textit{codatatype} \label{sssec:codatatype-codatatype}\ text \ The \textit{codatatype} antiquotation, written \texttt{\char`\\\char`<\char`^}\verb|codatatype>|\guilsinglleft\textit{t}\guilsinglright{} or \texttt{@}\verb|{codatatype| \textit{t}\verb|}|, where \textit{t} is a type name, expands to \LaTeX{} code for the definition of the codatatype, with each constructor listed with its argument types. For example, if \textit{t} is @{type llist}: \begin{quote} \<^codatatype>\llist\ \end{quote} \ section \Defining Primitively Corecursive Functions \label{sec:defining-primitively-corecursive-functions}\ text \ Corecursive functions can be specified using the @{command primcorec} and \keyw{prim\-corec\-ursive} commands, which support primitive corecursion. Other approaches include the more general \keyw{partial_function} command, the \keyw{corec} and \keyw{corecursive} commands, and techniques based on domains -and topologies @{cite "lochbihler-hoelzl-2014"}. In this tutorial, the focus is +and topologies \<^cite>\"lochbihler-hoelzl-2014"\. In this tutorial, the focus is on @{command primcorec} and @{command primcorecursive}; \keyw{corec} and \keyw{corecursive} are described in a separate tutorial -@{cite "isabelle-corec"}. More examples can be found in the directories +\<^cite>\"isabelle-corec"\. More examples can be found in the directories \<^dir>\~~/src/HOL/Datatype_Examples\ and \<^dir>\~~/src/HOL/Corec_Examples\. Whereas recursive functions consume datatypes one constructor at a time, corecursive functions construct codatatypes one constructor at a time. Partly reflecting a lack of agreement among proponents of coalgebraic methods, Isabelle supports three competing syntaxes for specifying a function $f$: \begin{itemize} \setlength{\itemsep}{0pt} \abovedisplayskip=.5\abovedisplayskip \belowdisplayskip=.5\belowdisplayskip \item The \emph{destructor view} specifies $f$ by implications of the form \[\\ \ is_C\<^sub>j (f x\<^sub>1 \ x\<^sub>n)\\] and equations of the form \[\un_C\<^sub>ji (f x\<^sub>1 \ x\<^sub>n) = \\\] This style is popular in the coalgebraic literature. \item The \emph{constructor view} specifies $f$ by equations of the form \[\\ \ f x\<^sub>1 \ x\<^sub>n = C\<^sub>j \\\] This style is often more concise than the previous one. \item The \emph{code view} specifies $f$ by a single equation of the form \[\f x\<^sub>1 \ x\<^sub>n = \\\] with restrictions on the format of the right-hand side. Lazy functional programming languages such as Haskell support a generalized version of this style. \end{itemize} All three styles are available as input syntax. Whichever syntax is chosen, characteristic theorems for all three styles are generated. %%% TODO: partial_function? E.g. for defining tail recursive function on lazy %%% lists (cf. terminal0 in TLList.thy) \ subsection \Introductory Examples \label{ssec:primcorec-introductory-examples}\ text \ Primitive corecursion is illustrated through concrete examples based on the codatatypes defined in Section~\ref{ssec:codatatype-introductory-examples}. More examples can be found in the directory \<^dir>\~~/src/HOL/Datatype_Examples\. The code view is favored in the examples below. Sections \ref{ssec:primrec-constructor-view} and \ref{ssec:primrec-destructor-view} present the same examples expressed using the constructor and destructor views. \ subsubsection \Simple Corecursion \label{sssec:primcorec-simple-corecursion}\ text \ Following the code view, corecursive calls are allowed on the right-hand side as long as they occur under a constructor, which itself appears either directly to the right of the equal sign or in a conditional expression: \ primcorec (*<*)(transfer) (*>*)literate :: "('a \ 'a) \ 'a \ 'a llist" where "literate g x = LCons x (literate g (g x))" text \\blankline\ primcorec siterate :: "('a \ 'a) \ 'a \ 'a stream" where "siterate g x = SCons x (siterate g (g x))" text \ \noindent The constructor ensures that progress is made---i.e., the function is \emph{productive}. The above functions compute the infinite lazy list or stream \[x, g x, g (g x), \]\. Productivity guarantees that prefixes \[x, g x, g (g x), \, (g ^^ k) x]\ of arbitrary finite length \k\ can be computed by unfolding the code equation a finite number of times. Corecursive functions construct codatatype values, but nothing prevents them from also consuming such values. The following function drops every second element in a stream: \ primcorec every_snd :: "'a stream \ 'a stream" where "every_snd s = SCons (shd s) (stl (stl s))" text \ \noindent Constructs such as \let\--\in\, \if\--\then\--\else\, and \case\--\of\ may appear around constructors that guard corecursive calls: \ primcorec lapp :: "'a llist \ 'a llist \ 'a llist" where "lapp xs ys = (case xs of LNil \ ys | LCons x xs' \ LCons x (lapp xs' ys))" text \ \noindent For technical reasons, \case\--\of\ is only supported for case distinctions on (co)datatypes that provide discriminators and selectors. Pattern matching is not supported by @{command primcorec}. Fortunately, it is easy to generate pattern-maching equations using the @{command simps_of_case} command provided by the theory \<^file>\~~/src/HOL/Library/Simps_Case_Conv.thy\. \ simps_of_case lapp_simps: lapp.code text \ This generates the lemma collection @{thm [source] lapp_simps}: % \begin{gather*} @{thm lapp_simps(1)[no_vars]} \\ @{thm lapp_simps(2)[no_vars]} \end{gather*} % Corecursion is useful to specify not only functions but also infinite objects: \ primcorec infty :: enat where "infty = ESucc infty" text \ \noindent The example below constructs a pseudorandom process value. It takes a stream of actions (\s\), a pseudorandom function generator (\f\), and a pseudorandom seed (\n\): \ (*<*) context early begin (*>*) primcorec random_process :: "'a stream \ (int \ int) \ int \ 'a process" where "random_process s f n = (if n mod 4 = 0 then Fail else if n mod 4 = 1 then Skip (random_process s f (f n)) else if n mod 4 = 2 then Action (shd s) (random_process (stl s) f (f n)) else Choice (random_process (every_snd s) (f \ f) (f n)) (random_process (every_snd (stl s)) (f \ f) (f (f n))))" (*<*) end (*>*) text \ \noindent The main disadvantage of the code view is that the conditions are tested sequentially. This is visible in the generated theorems. The constructor and destructor views offer nonsequential alternatives. \ subsubsection \Mutual Corecursion \label{sssec:primcorec-mutual-corecursion}\ text \ The syntax for mutually corecursive functions over mutually corecursive datatypes is unsurprising: \ primcorec even_infty :: even_enat and odd_infty :: odd_enat where "even_infty = Even_ESucc odd_infty" | "odd_infty = Odd_ESucc even_infty" subsubsection \Nested Corecursion \label{sssec:primcorec-nested-corecursion}\ text \ The next pair of examples generalize the \<^const>\literate\ and \<^const>\siterate\ functions (Section~\ref{sssec:primcorec-nested-corecursion}) to possibly infinite trees in which subnodes are organized either as a lazy list (\tree\<^sub>i\<^sub>i\) or as a finite set (\tree\<^sub>i\<^sub>s\). They rely on the map functions of the nesting type constructors to lift the corecursive calls: \ primcorec iterate\<^sub>i\<^sub>i :: "('a \ 'a llist) \ 'a \ 'a tree\<^sub>i\<^sub>i" where "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i g) (g x))" text \\blankline\ primcorec iterate\<^sub>i\<^sub>s :: "('a \ 'a fset) \ 'a \ 'a tree\<^sub>i\<^sub>s" where "iterate\<^sub>i\<^sub>s g x = Node\<^sub>i\<^sub>s x (fimage (iterate\<^sub>i\<^sub>s g) (g x))" text \ \noindent Both examples follow the usual format for constructor arguments associated with nested recursive occurrences of the datatype. Consider \<^const>\iterate\<^sub>i\<^sub>i\. The term \<^term>\g x\ constructs an \<^typ>\'a llist\ value, which is turned into an \<^typ>\'a tree\<^sub>i\<^sub>i llist\ value using \<^const>\lmap\. This format may sometimes feel artificial. The following function constructs a tree with a single, infinite branch from a stream: \ primcorec tree\<^sub>i\<^sub>i_of_stream :: "'a stream \ 'a tree\<^sub>i\<^sub>i" where "tree\<^sub>i\<^sub>i_of_stream s = Node\<^sub>i\<^sub>i (shd s) (lmap tree\<^sub>i\<^sub>i_of_stream (LCons (stl s) LNil))" text \ \noindent A more natural syntax, also supported by Isabelle, is to move corecursive calls under constructors: \ primcorec (*<*)(in late) (*>*)tree\<^sub>i\<^sub>i_of_stream :: "'a stream \ 'a tree\<^sub>i\<^sub>i" where "tree\<^sub>i\<^sub>i_of_stream s = Node\<^sub>i\<^sub>i (shd s) (LCons (tree\<^sub>i\<^sub>i_of_stream (stl s)) LNil)" text \ The next example illustrates corecursion through functions, which is a bit special. Deterministic finite automata (DFAs) are traditionally defined as 5-tuples \(Q, \, \, q\<^sub>0, F)\, where \Q\ is a finite set of states, \\\ is a finite alphabet, \\\ is a transition function, \q\<^sub>0\ is an initial state, and \F\ is a set of final states. The following function translates a DFA into a state machine: \ primcorec (*<*)(in early) (*>*)sm_of_dfa :: "('q \ 'a \ 'q) \ 'q set \ 'q \ 'a sm" where "sm_of_dfa \ F q = SM (q \ F) (sm_of_dfa \ F \ \ q)" text \ \noindent The map function for the function type (\\\) is composition (\(\)\). For convenience, corecursion through functions can also be expressed using $\lambda$-abstractions and function application rather than through composition. For example: \ primcorec sm_of_dfa :: "('q \ 'a \ 'q) \ 'q set \ 'q \ 'a sm" where "sm_of_dfa \ F q = SM (q \ F) (\a. sm_of_dfa \ F (\ q a))" text \\blankline\ primcorec empty_sm :: "'a sm" where "empty_sm = SM False (\_. empty_sm)" text \\blankline\ primcorec not_sm :: "'a sm \ 'a sm" where "not_sm M = SM (\ accept M) (\a. not_sm (trans M a))" text \\blankline\ primcorec or_sm :: "'a sm \ 'a sm \ 'a sm" where "or_sm M N = SM (accept M \ accept N) (\a. or_sm (trans M a) (trans N a))" text \ \noindent For recursion through curried $n$-ary functions, $n$ applications of \<^term>\(\)\ are necessary. The examples below illustrate the case where $n = 2$: \ codatatype ('a, 'b) sm2 = SM2 (accept2: bool) (trans2: "'a \ 'b \ ('a, 'b) sm2") text \\blankline\ primcorec (*<*)(in early) (*>*)sm2_of_dfa :: "('q \ 'a \ 'b \ 'q) \ 'q set \ 'q \ ('a, 'b) sm2" where "sm2_of_dfa \ F q = SM2 (q \ F) ((\) ((\) (sm2_of_dfa \ F)) (\ q))" text \\blankline\ primcorec sm2_of_dfa :: "('q \ 'a \ 'b \ 'q) \ 'q set \ 'q \ ('a, 'b) sm2" where "sm2_of_dfa \ F q = SM2 (q \ F) (\a b. sm2_of_dfa \ F (\ q a b))" subsubsection \Nested-as-Mutual Corecursion \label{sssec:primcorec-nested-as-mutual-corecursion}\ text \ Just as it is possible to recurse over nested recursive datatypes as if they were mutually recursive (Section~\ref{sssec:primrec-nested-as-mutual-recursion}), it is possible to pretend that nested codatatypes are mutually corecursive. For example: \ (*<*) context late begin (*>*) primcorec iterate\<^sub>i\<^sub>i :: "('a \ 'a llist) \ 'a \ 'a tree\<^sub>i\<^sub>i" and iterates\<^sub>i\<^sub>i :: "('a \ 'a llist) \ 'a llist \ 'a tree\<^sub>i\<^sub>i llist" where "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i g (g x))" | "iterates\<^sub>i\<^sub>i g xs = (case xs of LNil \ LNil | LCons x xs' \ LCons (iterate\<^sub>i\<^sub>i g x) (iterates\<^sub>i\<^sub>i g xs'))" text \ \noindent Coinduction rules are generated as @{thm [source] iterate\<^sub>i\<^sub>i.coinduct}, @{thm [source] iterates\<^sub>i\<^sub>i.coinduct}, and @{thm [source] iterate\<^sub>i\<^sub>i_iterates\<^sub>i\<^sub>i.coinduct} and analogously for \coinduct_strong\. These rules and the underlying corecursors are generated dynamically and are kept in a cache to speed up subsequent definitions. \ (*<*) end (*>*) subsubsection \Constructor View \label{ssec:primrec-constructor-view}\ (*<*) locale ctr_view begin (*>*) text \ The constructor view is similar to the code view, but there is one separate conditional equation per constructor rather than a single unconditional equation. Examples that rely on a single constructor, such as \<^const>\literate\ and \<^const>\siterate\, are identical in both styles. Here is an example where there is a difference: \ primcorec lapp :: "'a llist \ 'a llist \ 'a llist" where "lnull xs \ lnull ys \ lapp xs ys = LNil" | "_ \ lapp xs ys = LCons (lhd (if lnull xs then ys else xs)) (if xs = LNil then ltl ys else lapp (ltl xs) ys)" text \ \noindent With the constructor view, we must distinguish between the \<^const>\LNil\ and the \<^const>\LCons\ case. The condition for \<^const>\LCons\ is left implicit, as the negation of that for \<^const>\LNil\. For this example, the constructor view is slightly more involved than the code equation. Recall the code view version presented in Section~\ref{sssec:primcorec-simple-corecursion}. % TODO: \[{thm code_view.lapp.code}\] The constructor view requires us to analyze the second argument (\<^term>\ys\). The code equation generated from the constructor view also suffers from this. % TODO: \[{thm lapp.code}\] In contrast, the next example is arguably more naturally expressed in the constructor view: \ primcorec random_process :: "'a stream \ (int \ int) \ int \ 'a process" where "n mod 4 = 0 \ random_process s f n = Fail" | "n mod 4 = 1 \ random_process s f n = Skip (random_process s f (f n))" | "n mod 4 = 2 \ random_process s f n = Action (shd s) (random_process (stl s) f (f n))" | "n mod 4 = 3 \ random_process s f n = Choice (random_process (every_snd s) f (f n)) (random_process (every_snd (stl s)) f (f n))" (*<*) end (*>*) text \ \noindent Since there is no sequentiality, we can apply the equation for \<^const>\Choice\ without having first to discharge \<^term>\n mod (4::int) \ 0\, \<^term>\n mod (4::int) \ 1\, and \<^term>\n mod (4::int) \ 2\. The price to pay for this elegance is that we must discharge exclusiveness proof obligations, one for each pair of conditions \<^term>\(n mod (4::int) = i, n mod (4::int) = j)\ with \<^term>\i < j\. If we prefer not to discharge any obligations, we can enable the \sequential\ option. This pushes the problem to the users of the generated properties. %Here are more examples to conclude: \ subsubsection \Destructor View \label{ssec:primrec-destructor-view}\ (*<*) locale dtr_view begin (*>*) text \ The destructor view is in many respects dual to the constructor view. Conditions determine which constructor to choose, and these conditions are interpreted sequentially or not depending on the \sequential\ option. Consider the following examples: \ primcorec literate :: "('a \ 'a) \ 'a \ 'a llist" where "\ lnull (literate _ x)" | "lhd (literate _ x) = x" | "ltl (literate g x) = literate g (g x)" text \\blankline\ primcorec siterate :: "('a \ 'a) \ 'a \ 'a stream" where "shd (siterate _ x) = x" | "stl (siterate g x) = siterate g (g x)" text \\blankline\ primcorec every_snd :: "'a stream \ 'a stream" where "shd (every_snd s) = shd s" | "stl (every_snd s) = stl (stl s)" text \ \noindent The first formula in the \<^const>\literate\ specification indicates which constructor to choose. For \<^const>\siterate\ and \<^const>\every_snd\, no such formula is necessary, since the type has only one constructor. The last two formulas are equations specifying the value of the result for the relevant selectors. Corecursive calls appear directly to the right of the equal sign. Their arguments are unrestricted. The next example shows how to specify functions that rely on more than one constructor: \ primcorec lapp :: "'a llist \ 'a llist \ 'a llist" where "lnull xs \ lnull ys \ lnull (lapp xs ys)" | "lhd (lapp xs ys) = lhd (if lnull xs then ys else xs)" | "ltl (lapp xs ys) = (if xs = LNil then ltl ys else lapp (ltl xs) ys)" text \ \noindent For a codatatype with $n$ constructors, it is sufficient to specify $n - 1$ discriminator formulas. The command will then assume that the remaining constructor should be taken otherwise. This can be made explicit by adding \ (*<*) end locale dtr_view2 begin primcorec lapp :: "'a llist \ 'a llist \ 'a llist" where "lnull xs \ lnull ys \ lnull (lapp xs ys)" | "lhd (lapp xs ys) = lhd (if lnull xs then ys else xs)" | "ltl (lapp xs ys) = (if xs = LNil then ltl ys else lapp (ltl xs) ys)" | (*>*) "_ \ \ lnull (lapp xs ys)" text \ \noindent to the specification. The generated selector theorems are conditional. The next example illustrates how to cope with selectors defined for several constructors: \ primcorec random_process :: "'a stream \ (int \ int) \ int \ 'a process" where "n mod 4 = 0 \ random_process s f n = Fail" | "n mod 4 = 1 \ is_Skip (random_process s f n)" | "n mod 4 = 2 \ is_Action (random_process s f n)" | "n mod 4 = 3 \ is_Choice (random_process s f n)" | "cont (random_process s f n) = random_process s f (f n)" of Skip | "prefix (random_process s f n) = shd s" | "cont (random_process s f n) = random_process (stl s) f (f n)" of Action | "left (random_process s f n) = random_process (every_snd s) f (f n)" | "right (random_process s f n) = random_process (every_snd (stl s)) f (f n)" text \ \noindent Using the \of\ keyword, different equations are specified for \<^const>\cont\ depending on which constructor is selected. Here are more examples to conclude: \ primcorec even_infty :: even_enat and odd_infty :: odd_enat where "even_infty \ Even_EZero" | "un_Even_ESucc even_infty = odd_infty" | "un_Odd_ESucc odd_infty = even_infty" text \\blankline\ primcorec iterate\<^sub>i\<^sub>i :: "('a \ 'a llist) \ 'a \ 'a tree\<^sub>i\<^sub>i" where "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = x" | "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = lmap (iterate\<^sub>i\<^sub>i g) (g x)" (*<*) end (*>*) subsection \Command Syntax \label{ssec:primcorec-command-syntax}\ subsubsection \\keyw{primcorec} and \keyw{primcorecursive} \label{sssec:primcorecursive-and-primcorec}\ text \ \begin{matharray}{rcl} @{command_def "primcorec"} & : & \local_theory \ local_theory\ \\ @{command_def "primcorecursive"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ (@@{command primcorec} | @@{command primcorecursive}) target? \ @{syntax pcr_options}? fixes @'where' (@{syntax pcr_formula} + '|') ; @{syntax_def pcr_options}: '(' ((@{syntax plugins} | 'sequential' | 'exhaustive' | 'transfer') + ',') ')' ; @{syntax_def pcr_formula}: thmdecl? prop (@'of' (term * ))? \ \medskip \noindent The @{command primcorec} and @{command primcorecursive} commands introduce a set of mutually corecursive functions over codatatypes. The syntactic entity \synt{target} can be used to specify a local context, \synt{fixes} denotes a list of names with optional type signatures, \synt{thmdecl} denotes an optional name for the formula that follows, and -\synt{prop} denotes a HOL proposition @{cite "isabelle-isar-ref"}. +\synt{prop} denotes a HOL proposition \<^cite>\"isabelle-isar-ref"\. The optional target is optionally followed by a combination of the following options: \begin{itemize} \setlength{\itemsep}{0pt} \item The \plugins\ option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. \item The \sequential\ option indicates that the conditions in specifications expressed using the constructor or destructor view are to be interpreted sequentially. \item The \exhaustive\ option indicates that the conditions in specifications expressed using the constructor or destructor view cover all possible cases. This generally gives rise to an additional proof obligation. \item The \transfer\ option indicates that an unconditional transfer rule should be generated and proved \by transfer_prover\. The \[transfer_rule]\ attribute is set on the generated theorem. \end{itemize} The @{command primcorec} command is an abbreviation for @{command primcorecursive} with \by auto?\ to discharge any emerging proof obligations. %%% TODO: elaborate on format of the propositions %%% TODO: elaborate on mutual and nested-to-mutual \ subsection \Generated Theorems \label{ssec:primcorec-generated-theorems}\ text \ The @{command primcorec} and @{command primcorecursive} commands generate the following properties (listed for \<^const>\literate\): \begin{indentblock} \begin{description} \item[\f.\\hthm{code} \[code]\\rm:] ~ \\ @{thm literate.code[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\f.\\hthm{ctr}\rm:] ~ \\ @{thm literate.ctr[no_vars]} \item[\f.\\hthm{disc} \[simp, code]\\rm:] ~ \\ @{thm literate.disc[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). The \[simp]\ attribute is set only for functions for which \f.disc_iff\ is not available. \item[\f.\\hthm{disc_iff} \[simp]\\rm:] ~ \\ @{thm literate.disc_iff[no_vars]} \\ This property is generated only for functions declared with the \exhaustive\ option or whose conditions are trivially exhaustive. \item[\f.\\hthm{sel} \[simp, code]\\rm:] ~ \\ @{thm literate.disc[no_vars]} \\ The \[code]\ attribute is set by the \code\ plugin (Section~\ref{ssec:code-generator}). \item[\f.\\hthm{exclude}\rm:] ~ \\ These properties are missing for \<^const>\literate\ because no exclusiveness proof obligations arose. In general, the properties correspond to the discharged proof obligations. \item[\f.\\hthm{exhaust}\rm:] ~ \\ This property is missing for \<^const>\literate\ because no exhaustiveness proof obligation arose. In general, the property correspond to the discharged proof obligation. \item[\begin{tabular}{@ {}l@ {}} \f.\\hthm{coinduct} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\f.\\hthm{coinduct} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \end{tabular}] ~ \\ This coinduction rule is generated for nested-as-mutual corecursive functions (Section~\ref{sssec:primcorec-nested-as-mutual-corecursion}). \item[\begin{tabular}{@ {}l@ {}} \f.\\hthm{coinduct_strong} \[consumes m, case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\f.\\hthm{coinduct_strong} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \end{tabular}] ~ \\ This coinduction rule is generated for nested-as-mutual corecursive functions (Section~\ref{sssec:primcorec-nested-as-mutual-corecursion}). \item[\begin{tabular}{@ {}l@ {}} \f\<^sub>1_\_f\<^sub>m.\\hthm{coinduct} \[case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\f.\\hthm{coinduct} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \end{tabular}] ~ \\ This coinduction rule is generated for nested-as-mutual corecursive functions (Section~\ref{sssec:primcorec-nested-as-mutual-corecursion}). Given $m > 1$ mutually corecursive functions, this rule can be used to prove $m$ properties simultaneously. \item[\begin{tabular}{@ {}l@ {}} \f\<^sub>1_\_f\<^sub>m.\\hthm{coinduct_strong} \[case_names t\<^sub>1 \ t\<^sub>m,\ \\ \phantom{\f.\\hthm{coinduct_strong} \[\}\case_conclusion D\<^sub>1 \ D\<^sub>n]\\rm: \end{tabular}] ~ \\ This coinduction rule is generated for nested-as-mutual corecursive functions (Section~\ref{sssec:primcorec-nested-as-mutual-corecursion}). Given $m > 1$ mutually corecursive functions, this rule can be used to prove $m$ properties simultaneously. \end{description} \end{indentblock} \noindent For convenience, @{command primcorec} and @{command primcorecursive} also provide the following collection: \begin{indentblock} \begin{description} \item[\f.\\hthm{simps}] = \f.disc_iff\ (or \f.disc\) \t.sel\ \end{description} \end{indentblock} \ (* subsection \Recursive Default Values for Selectors \label{ssec:primcorec-recursive-default-values-for-selectors}\ text \ partial_function to the rescue \ *) section \Registering Bounded Natural Functors \label{sec:registering-bounded-natural-functors}\ text \ The (co)datatype package can be set up to allow nested recursion through arbitrary type constructors, as long as they adhere to the BNF requirements and are registered as BNFs. It is also possible to declare a BNF abstractly without specifying its internal structure. \ subsection \Bounded Natural Functors \label{ssec:bounded-natural-functors}\ text \ Bounded natural functors (BNFs) are a semantic criterion for where (co)re\-cur\-sion may appear on the right-hand side of an equation -@{cite "traytel-et-al-2012" and "blanchette-et-al-2015-wit"}. +\<^cite>\"traytel-et-al-2012" and "blanchette-et-al-2015-wit"\. An $n$-ary BNF is a type constructor equipped with a map function (functorial action), $n$ set functions (natural transformations), and an infinite cardinal bound that satisfy certain properties. For example, \<^typ>\'a llist\ is a unary BNF. Its predicator \llist_all :: ('a \ bool) \ 'a llist \ bool\ extends unary predicates over elements to unary predicates over lazy lists. Similarly, its relator \llist_all2 :: ('a \ 'b \ bool) \ 'a llist \ 'b llist \ bool\ extends binary predicates over elements to binary predicates over parallel lazy lists. The cardinal bound limits the number of elements returned by the set function; it may not depend on the cardinality of \<^typ>\'a\. The type constructors introduced by @{command datatype} and @{command codatatype} are automatically registered as BNFs. In addition, a number of old-style datatypes and non-free types are preregistered. Given an $n$-ary BNF, the $n$ type variables associated with set functions, and on which the map function acts, are \emph{live}; any other variables are \emph{dead}. Nested (co)recursion can only take place through live variables. \ subsection \Introductory Examples \label{ssec:bnf-introductory-examples}\ text \ The example below shows how to register a type as a BNF using the @{command bnf} command. Some of the proof obligations are best viewed with the bundle "cardinal_syntax" included. The type is simply a copy of the function space \<^typ>\'d \ 'a\, where \<^typ>\'a\ is live and \<^typ>\'d\ is dead. We introduce it together with its map function, set function, predicator, and relator. \ typedef ('d, 'a) fn = "UNIV :: ('d \ 'a) set" by simp text \\blankline\ setup_lifting type_definition_fn text \\blankline\ lift_definition map_fn :: "('a \ 'b) \ ('d, 'a) fn \ ('d, 'b) fn" is "(\)" . lift_definition set_fn :: "('d, 'a) fn \ 'a set" is range . text \\blankline\ lift_definition pred_fn :: "('a \ bool) \ ('d, 'a) fn \ bool" is "pred_fun (\_. True)" . lift_definition rel_fn :: "('a \ 'b \ bool) \ ('d, 'a) fn \ ('d, 'b) fn \ bool" is "rel_fun (=)" . text \\blankline\ bnf "('d, 'a) fn" map: map_fn sets: set_fn bd: "natLeq +c card_suc |UNIV :: 'd set|" rel: rel_fn pred: pred_fn proof - show "map_fn id = id" by transfer auto next fix f :: "'a \ 'b" and g :: "'b \ 'c" show "map_fn (g \ f) = map_fn g \ map_fn f" by transfer (auto simp add: comp_def) next fix F :: "('d, 'a) fn" and f g :: "'a \ 'b" assume "\x. x \ set_fn F \ f x = g x" then show "map_fn f F = map_fn g F" by transfer auto next fix f :: "'a \ 'b" show "set_fn \ map_fn f = (`) f \ set_fn" by transfer (auto simp add: comp_def) next show "card_order (natLeq +c card_suc |UNIV :: 'd set| )" by (rule card_order_bd_fun) next show "cinfinite (natLeq +c card_suc |UNIV :: 'd set| )" by (rule Cinfinite_bd_fun[THEN conjunct1]) next show "regularCard (natLeq +c card_suc |UNIV :: 'd set| )" by (rule regularCard_bd_fun) next fix F :: "('d, 'a) fn" have "|set_fn F| \o |UNIV :: 'd set|" (is "_ \o ?U") by transfer (rule card_of_image) also have "?U o natLeq +c card_suc ?U" using Card_order_card_suc card_of_card_order_on ordLeq_csum2 by blast finally show "|set_fn F| 'b \ bool" and S :: "'b \ 'c \ bool" show "rel_fn R OO rel_fn S \ rel_fn (R OO S)" by (rule, transfer) (auto simp add: rel_fun_def) next fix R :: "'a \ 'b \ bool" show "rel_fn R = (\x y. \z. set_fn z \ {(x, y). R x y} \ map_fn fst z = x \ map_fn snd z = y)" unfolding fun_eq_iff relcompp.simps conversep.simps by transfer (force simp: rel_fun_def subset_iff) next fix P :: "'a \ bool" show "pred_fn P = (\x. Ball (set_fn x) P)" unfolding fun_eq_iff by transfer simp qed text \\blankline\ print_theorems print_bnfs text \ \noindent Using \keyw{print_theorems} and @{command print_bnfs}, we can contemplate and show the world what we have achieved. This particular example does not need any nonemptiness witness, because the one generated by default is good enough, but in general this would be necessary. See \<^file>\~~/src/HOL/Basic_BNFs.thy\, \<^file>\~~/src/HOL/Library/Countable_Set_Type.thy\, \<^file>\~~/src/HOL/Library/FSet.thy\, and \<^file>\~~/src/HOL/Library/Multiset.thy\ for further examples of BNF registration, some of which feature custom witnesses. For many typedefs and quotient types, lifting the BNF structure from the raw typ to the abstract type can be done uniformly. This is the task of the @{command lift_bnf} command. Using @{command lift_bnf}, the above registration of \<^typ>\('d, 'a) fn\ as a BNF becomes much shorter: \ (*<*) context early begin (*>*) lift_bnf (*<*)(no_warn_wits) (*>*)('d, 'a) fn by force+ (*<*) end (*>*) text \ For type copies (@{command typedef}s with \<^term>\UNIV\ as the representing set), the proof obligations are so simple that they can be discharged automatically, yielding another command, @{command copy_bnf}, which does not emit any proof obligations: \ (*<*) context late begin (*>*) copy_bnf ('d, 'a) fn (*<*) end (*>*) text \ Since record schemas are type copies, @{command copy_bnf} can be used to register them as BNFs: \ record 'a point = xval :: 'a yval :: 'a text \\blankline\ copy_bnf (*<*)(no_warn_transfer) (*>*)('a, 'z) point_ext text \ In the general case, the proof obligations generated by @{command lift_bnf} are simpler than the acual BNF properties. In particular, no cardinality reasoning is required. Consider the following type of nonempty lists: \ typedef 'a nonempty_list = "{xs :: 'a list. xs \ []}" by auto text \ The @{command lift_bnf} command requires us to prove that the set of nonempty lists is closed under the map function and the zip function. The latter only occurs implicitly in the goal, in form of the variable \<^term>\zs :: ('a \ 'b) list\. \ lift_bnf (*<*)(no_warn_wits,no_warn_transfer) (*>*)'a nonempty_list proof - fix f (*<*):: "'a \ 'c"(*>*)and xs :: "'a list" assume "xs \ {xs. xs \ []}" then show "map f xs \ {xs. xs \ []}" by (cases xs(*<*) rule: list.exhaust(*>*)) auto next fix zs :: "('a \ 'b) list" assume "map fst zs \ {xs. xs \ []}" "map snd zs \ {xs. xs \ []}" then show "\zs'\{xs. xs \ []}. set zs' \ set zs \ map fst zs' = map fst zs \ map snd zs' = map snd zs" by (cases zs (*<*)rule: list.exhaust(*>*)) (auto intro!: exI[of _ zs]) qed text \The @{command lift_bnf} command also supports quotient types. Here is an example that defines the option type as a quotient of the sum type. The proof obligations generated by @{command lift_bnf} for quotients are different from the ones for typedefs. You can find additional examples of usages of @{command lift_bnf} for both quotients and subtypes in the session \emph{HOL-Datatype_Examples}.\ inductive ignore_Inl :: "'a + 'a \ 'a + 'a \ bool" where "ignore_Inl (Inl x) (Inl y)" | "ignore_Inl (Inr x) (Inr x)" (*<*) inductive_simps ignore_Inl_simps[simp]: "ignore_Inl (Inl x) y" "ignore_Inl (Inr x') y" "ignore_Inl y (Inl x)" "ignore_Inl y (Inr x')" (*>*) lemma ignore_Inl_equivp: "ignore_Inl x x" "ignore_Inl x y \ ignore_Inl y x" "ignore_Inl x y \ ignore_Inl y z \ ignore_Inl x z" by (cases x; cases y; cases z; auto)+ quotient_type 'a myoption = "'a + 'a" / ignore_Inl unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def by (blast intro: ignore_Inl_equivp) lift_bnf 'a myoption (*<*)[wits: "Inl undefined :: 'a + 'a"](*>*) proof - fix P :: "'a \ 'b \ bool" and Q :: "'b \ 'c \ bool" assume "P OO Q \ bot" then show "rel_sum P P OO ignore_Inl OO rel_sum Q Q \ ignore_Inl OO rel_sum (P OO Q) (P OO Q) OO ignore_Inl" by (fastforce(*<*) elim!: ignore_Inl.cases simp add: relcompp_apply fun_eq_iff rel_sum.simps(*>*)) next fix S :: "'a set set" let ?eq = "{(x, x'). ignore_Inl x x'}" let ?in = "\A. {x. Basic_BNFs.setl x \ Basic_BNFs.setr x \ A}" assume "S \ {}" "\ S \ {}" show "(\A\S. ?eq `` ?in A) \ ?eq `` ?in (\ S)" proof (intro subsetI) fix x assume "x \ (\A\S. ?eq `` ?in A)" with \\ S \ {}\ show "x \ ?eq `` ?in (\ S)" by (cases x) (fastforce(*<*) dest!: spec[where x="Inl _"](*>*))+ qed (*<*)next fix a :: 'a assume "a \ (\mx\{mx. ignore_Inl (map_sum Inr Inr (Inl undefined)) mx}. \ (Basic_BNFs.setr ` (Basic_BNFs.setl mx \ Basic_BNFs.setr mx)))" then show False by (auto simp: setr.simps)(*>*) qed text \ The next example declares a BNF axiomatically. This can be convenient for reasoning abstractly about an arbitrary BNF. The @{command bnf_axiomatization} command below introduces a type \('a, 'b, 'c) F\, three set constants, a map function, a predicator, a relator, and a nonemptiness witness that depends only on \<^typ>\'a\. The type \'a \ ('a, 'b, 'c) F\ of the witness can be read as an implication: Given a witness for \<^typ>\'a\, we can construct a witness for \('a, 'b, 'c) F\. The BNF properties are postulated as axioms. \ bnf_axiomatization (setA: 'a, setB: 'b, setC: 'c) F [wits: "'a \ ('a, 'b, 'c) F"] text \\blankline\ print_theorems print_bnfs subsection \Command Syntax \label{ssec:bnf-command-syntax}\ subsubsection \\keyw{bnf} \label{sssec:bnf}\ text \ \begin{matharray}{rcl} @{command_def "bnf"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ @@{command bnf} target? (name ':')? type \ 'map:' term ('sets:' (term +))? 'bd:' term \ ('wits:' (term +))? ('rel:' term)? \ ('pred:' term)? @{syntax plugins}? \ \medskip \noindent The @{command bnf} command registers an existing type as a bounded natural functor (BNF). The type must be equipped with an appropriate map function (functorial action). In addition, custom set functions, predicators, relators, and nonemptiness witnesses can be specified; otherwise, default versions are used. The syntactic entity \synt{target} can be used to specify a local context, \synt{type} denotes a HOL type, and \synt{term} denotes a HOL term -@{cite "isabelle-isar-ref"}. +\<^cite>\"isabelle-isar-ref"\. The @{syntax plugins} option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. %%% TODO: elaborate on proof obligations \ subsubsection \\keyw{lift_bnf} \label{sssec:lift-bnf}\ text \ \begin{matharray}{rcl} @{command_def "lift_bnf"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ @@{command lift_bnf} target? lb_options? \ @{syntax tyargs} name wit_terms? \ ('via' thm thm?)? @{syntax map_rel_pred}? ; @{syntax_def lb_options}: '(' ((@{syntax plugins} | 'no_warn_wits' | 'no_warn_transfer') + ',') ')' ; @{syntax_def wit_terms}: '[' 'wits' ':' terms ']' \ \medskip \noindent The @{command lift_bnf} command registers as a BNF an existing type (the \emph{abstract type}) that was defined as a subtype of a BNF (the \emph{raw type}) using the @{command typedef} command or as a quotient type of a BNF (also, the \emph{raw type}) using the @{command quotient_type}. To achieve this, it lifts the BNF structure on the raw type to the abstract type following a \<^term>\type_definition\ or a \<^term>\Quotient\ theorem. The theorem is usually inferred from the type, but can also be explicitly supplied by means of the optional \via\ clause. In case of quotients, it is sometimes also necessary to supply a second theorem of the form \<^term>\reflp eq\, that expresses the reflexivity (and thus totality) of the equivalence relation. In addition, custom names for the set functions, the map function, the predicator, and the relator, as well as nonemptiness witnesses can be specified. Nonemptiness witnesses are not lifted from the raw type's BNF, as this would be incomplete. They must be given as terms (on the raw type) and proved to be witnesses. The command warns about witness types that are present in the raw type's BNF but not supplied by the user. The warning can be disabled by specifying the \no_warn_wits\ option. \ subsubsection \\keyw{copy_bnf} \label{sssec:copy-bnf}\ text \ \begin{matharray}{rcl} @{command_def "copy_bnf"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command copy_bnf} target? cb_options? \ @{syntax tyargs} name ('via' thm)? @{syntax map_rel_pred}? ; @{syntax_def cb_options}: '(' ((@{syntax plugins} | 'no_warn_transfer') + ',') ')' \ \medskip \noindent The @{command copy_bnf} command performs the same lifting as @{command lift_bnf} for type copies (@{command typedef}s with \<^term>\UNIV\ as the representing set), without requiring the user to discharge any proof obligations or provide nonemptiness witnesses. \ subsubsection \\keyw{bnf_axiomatization} \label{sssec:bnf-axiomatization}\ text \ \begin{matharray}{rcl} @{command_def "bnf_axiomatization"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command bnf_axiomatization} target? ('(' @{syntax plugins} ')')? \ @{syntax tyargs}? name @{syntax wit_types}? \ mixfix? @{syntax map_rel_pred}? ; @{syntax_def wit_types}: '[' 'wits' ':' types ']' \ \medskip \noindent The @{command bnf_axiomatization} command declares a new type and associated constants (map, set, predicator, relator, and cardinal bound) and asserts the BNF properties for these constants as axioms. The syntactic entity \synt{target} can be used to specify a local context, \synt{name} denotes an identifier, \synt{typefree} denotes fixed type variable (\<^typ>\'a\, \<^typ>\'b\, \ldots), \synt{mixfix} denotes the usual parenthesized mixfix notation, and \synt{types} denotes a space-separated list of types -@{cite "isabelle-isar-ref"}. +\<^cite>\"isabelle-isar-ref"\. The @{syntax plugins} option indicates which plugins should be enabled (\only\) or disabled (\del\). By default, all plugins are enabled. Type arguments are live by default; they can be marked as dead by entering \dead\ in front of the type variable (e.g., \(dead 'a)\) instead of an identifier for the corresponding set function. Witnesses can be specified by their types. Otherwise, the syntax of @{command bnf_axiomatization} is identical to the left-hand side of a @{command datatype} or @{command codatatype} definition. The command is useful to reason abstractly about BNFs. The axioms are safe because there exist BNFs of arbitrary large arities. Applications must import the \<^file>\~~/src/HOL/Library/BNF_Axiomatization.thy\ theory to use this functionality. \ subsubsection \\keyw{print_bnfs} \label{sssec:print-bnfs}\ text \ \begin{matharray}{rcl} @{command_def "print_bnfs"} & : & \local_theory \\ \end{matharray} \<^rail>\ @@{command print_bnfs} \ \ section \Deriving Destructors and Constructor Theorems \label{sec:deriving-destructors-and-constructor-theorems}\ text \ The derivation of convenience theorems for types equipped with free constructors, as performed internally by @{command datatype} and @{command codatatype}, is available as a stand-alone command called @{command free_constructors}. % * need for this is rare but may arise if you want e.g. to add destructors to % a type not introduced by ... % % * @{command free_constructors} % * \plugins\, \discs_sels\ % * hack to have both co and nonco view via locale (cf. ext nats) \ (* subsection \Introductory Example \label{ssec:ctors-introductory-example}\ *) subsection \Command Syntax \label{ssec:ctors-command-syntax}\ subsubsection \\keyw{free_constructors} \label{sssec:free-constructors}\ text \ \begin{matharray}{rcl} @{command_def "free_constructors"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ @@{command free_constructors} target? @{syntax dt_options} \ name 'for' (@{syntax fc_ctor} + '|') \ (@'where' (prop + '|'))? ; @{syntax_def fc_ctor}: (name ':')? term (name * ) \ \medskip \noindent The @{command free_constructors} command generates destructor constants for freely constructed types as well as properties about constructors and destructors. It also registers the constants and theorems in a data structure that is queried by various tools (e.g., \keyw{function}). The syntactic entity \synt{target} can be used to specify a local context, \synt{name} denotes an identifier, \synt{prop} denotes a HOL proposition, and -\synt{term} denotes a HOL term @{cite "isabelle-isar-ref"}. +\synt{term} denotes a HOL term \<^cite>\"isabelle-isar-ref"\. The syntax resembles that of @{command datatype} and @{command codatatype} definitions (Sections \ref{ssec:datatype-command-syntax}~and~\ref{ssec:codatatype-command-syntax}). A constructor is specified by an optional name for the discriminator, the constructor itself (as a term), and a list of optional names for the selectors. Section~\ref{ssec:datatype-generated-theorems} lists the generated theorems. For bootstrapping reasons, the generally useful \[fundef_cong]\ attribute is not set on the generated \case_cong\ theorem. It can be added manually using \keyw{declare}. \ subsubsection \\keyw{simps_of_case} \label{sssec:simps-of-case}\ text \ \begin{matharray}{rcl} @{command_def "simps_of_case"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command simps_of_case} target? (name ':')? \ (thm + ) (@'splits' ':' (thm + ))? \ \medskip \noindent The @{command simps_of_case} command provided by theory \<^file>\~~/src/HOL/Library/Simps_Case_Conv.thy\ converts a single equation with a complex case expression on the right-hand side into a set of pattern-matching equations. For example, \ (*<*) context late begin (*>*) simps_of_case lapp_simps: lapp.code text \ \noindent translates @{thm lapp.code[no_vars]} into % \begin{gather*} @{thm lapp_simps(1)[no_vars]} \\ @{thm lapp_simps(2)[no_vars]} \end{gather*} \ subsubsection \\keyw{case_of_simps} \label{sssec:case-of-simps}\ text \ \begin{matharray}{rcl} @{command_def "case_of_simps"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command case_of_simps} target? (name ':')? \ (thm + ) \ \medskip \noindent The \keyw{case_of_simps} command provided by theory \<^file>\~~/src/HOL/Library/Simps_Case_Conv.thy\ converts a set of pattern-matching equations into single equation with a complex case expression on the right-hand side (cf.\ @{command simps_of_case}). For example, \ case_of_simps lapp_case: lapp_simps text \ \noindent translates % \begin{gather*} @{thm lapp_simps(1)[no_vars]} \\ @{thm lapp_simps(2)[no_vars]} \end{gather*} % into @{thm lapp_case[no_vars]}. \ (*<*) end (*>*) (* section \Using the Standard ML Interface \label{sec:using-the-standard-ml-interface}\ text \ The package's programmatic interface. \ *) section \Selecting Plugins \label{sec:selecting-plugins}\ text \ Plugins extend the (co)datatype package to interoperate with other Isabelle packages and tools, such as the code generator, Transfer, Lifting, and Quickcheck. They can be enabled or disabled individually using the @{syntax plugins} option to the commands @{command datatype}, @{command primrec}, @{command codatatype}, @{command primcorec}, @{command primcorecursive}, @{command bnf}, @{command bnf_axiomatization}, and @{command free_constructors}. For example: \ datatype (plugins del: code "quickcheck") color = Red | Black text \ Beyond the standard plugins, the \emph{Archive of Formal Proofs} includes a \keyw{derive} command that derives class instances of datatypes -@{cite "sternagel-thiemann-2015"}. +\<^cite>\"sternagel-thiemann-2015"\. \ subsection \Code Generator \label{ssec:code-generator}\ text \ The \hthm{code} plugin registers freely generated types, including (co)datatypes, and (co)recursive functions for code generation. No distinction is made between datatypes and codatatypes. This means that for target languages with a strict evaluation strategy (e.g., Standard ML), programs that attempt to produce infinite codatatype values will not terminate. For types, the plugin derives the following properties: \begin{indentblock} \begin{description} \item[\t.\\hthm{eq.refl} \[code nbe]\\rm:] ~ \\ @{thm list.eq.refl[no_vars]} \item[\t.\\hthm{eq.simps} \[code]\\rm:] ~ \\ @{thm list.eq.simps(1)[no_vars]} \\ @{thm list.eq.simps(2)[no_vars]} \\ @{thm list.eq.simps(3)[no_vars]} \\ @{thm list.eq.simps(4)[no_vars]} \\ @{thm list.eq.simps(5)[no_vars]} \\ @{thm list.eq.simps(6)[no_vars]} \end{description} \end{indentblock} In addition, the plugin sets the \[code]\ attribute on a number of properties of freely generated types and of (co)recursive functions, as documented in Sections \ref{ssec:datatype-generated-theorems}, \ref{ssec:primrec-generated-theorems}, \ref{ssec:codatatype-generated-theorems}, and~\ref{ssec:primcorec-generated-theorems}. \ subsection \Size \label{ssec:size}\ text \ For each datatype \t\, the \hthm{size} plugin generates a generic size function \t.size_t\ as well as a specific instance \size :: t \ nat\ belonging to the \size\ type class. The \keyw{fun} command relies on \<^const>\size\ to prove termination of recursive functions on datatypes. The plugin derives the following properties: \begin{indentblock} \begin{description} \item[\t.\\hthm{size} \[simp, code]\\rm:] ~ \\ @{thm list.size(1)[no_vars]} \\ @{thm list.size(2)[no_vars]} \\ @{thm list.size(3)[no_vars]} \\ @{thm list.size(4)[no_vars]} \item[\t.\\hthm{size_gen}\rm:] ~ \\ @{thm list.size_gen(1)[no_vars]} \\ @{thm list.size_gen(2)[no_vars]} \item[\t.\\hthm{size_gen_o_map}\rm:] ~ \\ @{thm list.size_gen_o_map[no_vars]} \item[\t.\\hthm{size_neq}\rm:] ~ \\ This property is missing for \<^typ>\'a list\. If the \<^term>\size\ function always evaluates to a non-zero value, this theorem has the form \<^prop>\\ size x = 0\. \end{description} \end{indentblock} The \t.size\ and \t.size_t\ functions generated for datatypes defined by nested recursion through a datatype \u\ depend on \u.size_u\. If the recursion is through a non-datatype \u\ with type arguments \'a\<^sub>1, \, 'a\<^sub>m\, by default \u\ values are given a size of 0. This can be improved upon by registering a custom size function of type \('a\<^sub>1 \ nat) \ \ \ ('a\<^sub>m \ nat) \ u \ nat\ using the ML function \<^ML>\BNF_LFP_Size.register_size\ or \<^ML>\BNF_LFP_Size.register_size_global\. See theory \<^file>\~~/src/HOL/Library/Multiset.thy\ for an example. \ subsection \Transfer \label{ssec:transfer}\ text \ For each (co)datatype with live type arguments and each manually registered BNF, the \hthm{transfer} plugin generates a predicator \t.pred_t\ and properties that guide the Transfer tool. For types with at least one live type argument and \emph{no dead type arguments}, the plugin derives the following properties: \begin{indentblock} \begin{description} \item[\t.\\hthm{Domainp_rel} \[relator_domain]\\rm:] ~ \\ @{thm list.Domainp_rel[no_vars]} \item[\t.\\hthm{left_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.left_total_rel[no_vars]} \item[\t.\\hthm{left_unique_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.left_unique_rel[no_vars]} \item[\t.\\hthm{right_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.right_total_rel[no_vars]} \item[\t.\\hthm{right_unique_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.right_unique_rel[no_vars]} \item[\t.\\hthm{bi_total_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.bi_total_rel[no_vars]} \item[\t.\\hthm{bi_unique_rel} \[transfer_rule]\\rm:] ~ \\ @{thm list.bi_unique_rel[no_vars]} \end{description} \end{indentblock} For (co)datatypes with at least one live type argument, the plugin sets the \[transfer_rule]\ attribute on the following (co)datatypes properties: \t.case_\\allowbreak \transfer\, \t.sel_\\allowbreak \transfer\, \t.ctr_\\allowbreak \transfer\, \t.disc_\\allowbreak \transfer\, \t.rec_\\allowbreak \transfer\, and \t.corec_\\allowbreak \transfer\. For (co)datatypes that further have \emph{no dead type arguments}, the plugin sets \[transfer_rule]\ on \t.set_\\allowbreak \transfer\, \t.map_\\allowbreak \transfer\, and \t.rel_\\allowbreak \transfer\. For @{command primrec}, @{command primcorec}, and @{command primcorecursive}, the plugin implements the generation of the \f.transfer\ property, conditioned by the \transfer\ option, and sets the \[transfer_rule]\ attribute on these. \ subsection \Lifting \label{ssec:lifting}\ text \ For each (co)datatype and each manually registered BNF with at least one live type argument \emph{and no dead type arguments}, the \hthm{lifting} plugin generates properties and attributes that guide the Lifting tool. The plugin derives the following property: \begin{indentblock} \begin{description} \item[\t.\\hthm{Quotient} \[quot_map]\\rm:] ~ \\ @{thm list.Quotient[no_vars]} \end{description} \end{indentblock} In addition, the plugin sets the \[relator_eq]\ attribute on a variant of the \t.rel_eq_onp\ property, the \[relator_mono]\ attribute on \t.rel_mono\, and the \[relator_distr]\ attribute on \t.rel_compp\. \ subsection \Quickcheck \label{ssec:quickcheck}\ text \ The integration of datatypes with Quickcheck is accomplished by the \hthm{quick\-check} plugin. It combines a number of subplugins that instantiate specific type classes. The subplugins can be enabled or disabled individually. They are listed below: \begin{indentblock} \hthm{quickcheck_random} \\ \hthm{quickcheck_exhaustive} \\ \hthm{quickcheck_bounded_forall} \\ \hthm{quickcheck_full_exhaustive} \\ \hthm{quickcheck_narrowing} \end{indentblock} \ subsection \Program Extraction \label{ssec:program-extraction}\ text \ The \hthm{extraction} plugin provides realizers for induction and case analysis, to enable program extraction from proofs involving datatypes. This functionality is only available with full proof objects, i.e., with the \emph{HOL-Proofs} session. \ section \Known Bugs and Limitations \label{sec:known-bugs-and-limitations}\ text \ This section lists the known bugs and limitations of the (co)datatype package at the time of this writing. \begin{enumerate} \setlength{\itemsep}{0pt} \item \emph{Defining mutually (co)recursive (co)datatypes can be slow.} Fortunately, it is always possible to recast mutual specifications to nested ones, which are processed more efficiently. \item \emph{Locally fixed types and terms cannot be used in type specifications.} The limitation on types can be circumvented by adding type arguments to the local (co)datatypes to abstract over the locally fixed types. \item \emph{The \emph{\keyw{primcorec}} command does not allow user-specified names and attributes next to the entered formulas.} The less convenient syntax, using the \keyw{lemmas} command, is available as an alternative. \item \emph{The \emph{\keyw{primcorec}} command does not allow corecursion under \case\--\of\ for datatypes that are defined without discriminators and selectors.} \item \emph{There is no way to use an overloaded constant from a syntactic type class, such as \0\, as a constructor.} \item \emph{There is no way to register the same type as both a datatype and a codatatype.} This affects types such as the extended natural numbers, for which both views would make sense (for a different set of constructors). \item \emph{The names of variables are often suboptimal in the properties generated by the package.} \item \emph{The compatibility layer sometimes produces induction principles with a slightly different ordering of the premises than the old package.} \end{enumerate} \ text \ \section*{Acknowledgment} Tobias Nipkow and Makarius Wenzel encouraged us to implement the new (co)datatype package. Andreas Lochbihler provided lots of comments on earlier versions of the package, especially on the coinductive part. Brian Huffman suggested major simplifications to the internal constructions. Ond\v{r}ej Kun\v{c}ar implemented the \transfer\ and \lifting\ plugins. Christian Sternagel and Ren\'e Thiemann ported the \keyw{derive} command from the \emph{Archive of Formal Proofs} to the new datatypes. Gerwin Klein and Lars Noschinski implemented the @{command simps_of_case} and @{command case_of_simps} commands. Florian Haftmann, Christian Urban, and Makarius Wenzel provided general advice on Isabelle and package writing. Stefan Milius and Lutz Schröder found an elegant proof that eliminated one of the BNF proof obligations. Mamoun Filali-Amine, Gerwin Klein, Andreas Lochbihler, Tobias Nipkow, and Christian Sternagel suggested many textual improvements to this tutorial. \ end diff --git a/src/Doc/Demo_EPTCS/Document.thy b/src/Doc/Demo_EPTCS/Document.thy --- a/src/Doc/Demo_EPTCS/Document.thy +++ b/src/Doc/Demo_EPTCS/Document.thy @@ -1,64 +1,64 @@ theory Document imports Main begin section \Some section\ subsection \Some subsection\ subsection \Some subsubsection\ subsubsection \Some subsubsubsection\ paragraph \A paragraph.\ text \Informal bla bla.\ definition "foo = True" \ \side remark on \<^const>\foo\\ definition "bar = False" \ \side remark on \<^const>\bar\\ lemma foo unfolding foo_def .. paragraph \Another paragraph.\ -text \See also @{cite \\S3\ "isabelle-system"}.\ +text \See also \<^cite>\\\S3\ in "isabelle-system"\.\ section \Formal proof of Cantor's theorem\ text_raw \\isakeeptag{proof}\ text \ Cantor's Theorem states that there is no surjection from a set to its powerset. The proof works by diagonalization. E.g.\ see \<^item> \<^url>\http://mathworld.wolfram.com/CantorDiagonalMethod.html\ \<^item> \<^url>\https://en.wikipedia.org/wiki/Cantor's_diagonal_argument\ \ theorem Cantor: "\f :: 'a \ 'a set. \A. \x. A = f x" proof assume "\f :: 'a \ 'a set. \A. \x. A = f x" then obtain f :: "'a \ 'a set" where *: "\A. \x. A = f x" .. let ?D = "{x. x \ f x}" from * obtain a where "?D = f a" by blast moreover have "a \ ?D \ a \ f a" by blast ultimately show False by blast qed subsection \Lorem ipsum dolor\ text \ Lorem ipsum dolor sit amet, consectetur adipiscing elit. Donec id ipsum sapien. Vivamus malesuada enim nibh, a tristique nisi sodales ac. Praesent ut sem consectetur, interdum tellus ac, sodales nulla. Quisque vel diam at risus tempus tempor eget a tortor. Suspendisse potenti. Nulla erat lacus, dignissim sed volutpat nec, feugiat non leo. Nunc blandit et justo sed venenatis. Donec scelerisque placerat magna, et congue nulla convallis vel. Cras tristique dolor consequat dolor tristique rutrum. Suspendisse ultrices sem nibh, et suscipit felis ultricies at. Aliquam venenatis est vel nulla efficitur ornare. Lorem ipsum dolor sit amet, consectetur adipiscing elit. \ end diff --git a/src/Doc/Demo_Easychair/Document.thy b/src/Doc/Demo_Easychair/Document.thy --- a/src/Doc/Demo_Easychair/Document.thy +++ b/src/Doc/Demo_Easychair/Document.thy @@ -1,64 +1,64 @@ theory Document imports Main begin section \Some section\ subsection \Some subsection\ subsection \Some subsubsection\ subsubsection \Some subsubsubsection\ paragraph \A paragraph.\ text \Informal bla bla.\ definition "foo = True" \ \side remark on \<^const>\foo\\ definition "bar = False" \ \side remark on \<^const>\bar\\ lemma foo unfolding foo_def .. paragraph \Another paragraph.\ -text \See also @{cite \\S3\ "isabelle-system"}.\ +text \See also \<^cite>\\\S3\ in "isabelle-system"\.\ section \Formal proof of Cantor's theorem\ text_raw \\isakeeptag{proof}\ text \ Cantor's Theorem states that there is no surjection from a set to its powerset. The proof works by diagonalization. E.g.\ see \<^item> \<^url>\http://mathworld.wolfram.com/CantorDiagonalMethod.html\ \<^item> \<^url>\https://en.wikipedia.org/wiki/Cantor's_diagonal_argument\ \ theorem Cantor: "\f :: 'a \ 'a set. \A. \x. A = f x" proof assume "\f :: 'a \ 'a set. \A. \x. A = f x" then obtain f :: "'a \ 'a set" where *: "\A. \x. A = f x" .. let ?D = "{x. x \ f x}" from * obtain a where "?D = f a" by blast moreover have "a \ ?D \ a \ f a" by blast ultimately show False by blast qed subsection \Lorem ipsum dolor\ text \ Lorem ipsum dolor sit amet, consectetur adipiscing elit. Donec id ipsum sapien. Vivamus malesuada enim nibh, a tristique nisi sodales ac. Praesent ut sem consectetur, interdum tellus ac, sodales nulla. Quisque vel diam at risus tempus tempor eget a tortor. Suspendisse potenti. Nulla erat lacus, dignissim sed volutpat nec, feugiat non leo. Nunc blandit et justo sed venenatis. Donec scelerisque placerat magna, et congue nulla convallis vel. Cras tristique dolor consequat dolor tristique rutrum. Suspendisse ultrices sem nibh, et suscipit felis ultricies at. Aliquam venenatis est vel nulla efficitur ornare. Lorem ipsum dolor sit amet, consectetur adipiscing elit. \ end diff --git a/src/Doc/Demo_LIPIcs/Document.thy b/src/Doc/Demo_LIPIcs/Document.thy --- a/src/Doc/Demo_LIPIcs/Document.thy +++ b/src/Doc/Demo_LIPIcs/Document.thy @@ -1,64 +1,64 @@ theory Document imports Main begin section \Some section\ subsection \Some subsection\ subsection \Some subsubsection\ subsubsection \Some subsubsubsection\ paragraph \A paragraph.\ text \Informal bla bla.\ definition "foo = True" \ \side remark on \<^const>\foo\\ definition "bar = False" \ \side remark on \<^const>\bar\\ lemma foo unfolding foo_def .. paragraph \Another paragraph.\ -text \See also @{cite \\S3\ "isabelle-system"}.\ +text \See also \<^cite>\\\S3\ in "isabelle-system"\.\ section \Formal proof of Cantor's theorem\ text_raw \\isakeeptag{proof}\ text \ Cantor's Theorem states that there is no surjection from a set to its powerset. The proof works by diagonalization. E.g.\ see \<^item> \<^url>\http://mathworld.wolfram.com/CantorDiagonalMethod.html\ \<^item> \<^url>\https://en.wikipedia.org/wiki/Cantor's_diagonal_argument\ \ theorem Cantor: "\f :: 'a \ 'a set. \A. \x. A = f x" proof assume "\f :: 'a \ 'a set. \A. \x. A = f x" then obtain f :: "'a \ 'a set" where *: "\A. \x. A = f x" .. let ?D = "{x. x \ f x}" from * obtain a where "?D = f a" by blast moreover have "a \ ?D \ a \ f a" by blast ultimately show False by blast qed subsection \Lorem ipsum dolor\ text \ Lorem ipsum dolor sit amet, consectetur adipiscing elit. Donec id ipsum sapien. Vivamus malesuada enim nibh, a tristique nisi sodales ac. Praesent ut sem consectetur, interdum tellus ac, sodales nulla. Quisque vel diam at risus tempus tempor eget a tortor. Suspendisse potenti. Nulla erat lacus, dignissim sed volutpat nec, feugiat non leo. Nunc blandit et justo sed venenatis. Donec scelerisque placerat magna, et congue nulla convallis vel. Cras tristique dolor consequat dolor tristique rutrum. Suspendisse ultrices sem nibh, et suscipit felis ultricies at. Aliquam venenatis est vel nulla efficitur ornare. Lorem ipsum dolor sit amet, consectetur adipiscing elit. \ end diff --git a/src/Doc/Demo_LLNCS/Document.thy b/src/Doc/Demo_LLNCS/Document.thy --- a/src/Doc/Demo_LLNCS/Document.thy +++ b/src/Doc/Demo_LLNCS/Document.thy @@ -1,64 +1,64 @@ theory Document imports Main begin section \Some section\ subsection \Some subsection\ subsection \Some subsubsection\ subsubsection \Some subsubsubsection\ paragraph \A paragraph.\ text \Informal bla bla.\ definition "foo = True" \ \side remark on \<^const>\foo\\ definition "bar = False" \ \side remark on \<^const>\bar\\ lemma foo unfolding foo_def .. paragraph \Another paragraph.\ -text \See also @{cite \\S3\ "isabelle-system"}.\ +text \See also \<^cite>\\\S3\ in "isabelle-system"\.\ section \Formal proof of Cantor's theorem\ text_raw \\isakeeptag{proof}\ text \ Cantor's Theorem states that there is no surjection from a set to its powerset. The proof works by diagonalization. E.g.\ see \<^item> \<^url>\http://mathworld.wolfram.com/CantorDiagonalMethod.html\ \<^item> \<^url>\https://en.wikipedia.org/wiki/Cantor%27s%5fdiagonal%5fargument\ \ theorem Cantor: "\f :: 'a \ 'a set. \A. \x. A = f x" proof assume "\f :: 'a \ 'a set. \A. \x. A = f x" then obtain f :: "'a \ 'a set" where *: "\A. \x. A = f x" .. let ?D = "{x. x \ f x}" from * obtain a where "?D = f a" by blast moreover have "a \ ?D \ a \ f a" by blast ultimately show False by blast qed subsection \Lorem ipsum dolor\ text \ Lorem ipsum dolor sit amet, consectetur adipiscing elit. Donec id ipsum sapien. Vivamus malesuada enim nibh, a tristique nisi sodales ac. Praesent ut sem consectetur, interdum tellus ac, sodales nulla. Quisque vel diam at risus tempus tempor eget a tortor. Suspendisse potenti. Nulla erat lacus, dignissim sed volutpat nec, feugiat non leo. Nunc blandit et justo sed venenatis. Donec scelerisque placerat magna, et congue nulla convallis vel. Cras tristique dolor consequat dolor tristique rutrum. Suspendisse ultrices sem nibh, et suscipit felis ultricies at. Aliquam venenatis est vel nulla efficitur ornare. Lorem ipsum dolor sit amet, consectetur adipiscing elit. \ end diff --git a/src/Doc/Eisbach/Manual.thy b/src/Doc/Eisbach/Manual.thy --- a/src/Doc/Eisbach/Manual.thy +++ b/src/Doc/Eisbach/Manual.thy @@ -1,957 +1,957 @@ (*:maxLineLen=78:*) theory Manual imports Base "HOL-Eisbach.Eisbach_Tools" begin chapter \The method command\ text \ The @{command_def method} command provides the ability to write proof methods by combining existing ones with their usual syntax. Specifically it allows compound proof methods to be named, and to extend the name space of basic methods accordingly. Method definitions may abstract over parameters: terms, facts, or other methods. \<^medskip> The syntax diagram below refers to some syntactic categories that are - further defined in @{cite "isabelle-isar-ref"}. + further defined in \<^cite>\"isabelle-isar-ref"\. \<^rail>\ @@{command method} name args @'=' method ; args: term_args? method_args? \ fact_args? decl_args? ; term_args: @'for' @{syntax "fixes"} ; method_args: @'methods' (name+) ; fact_args: @'uses' (name+) ; decl_args: @'declares' (name+) \ \ section \Basic method definitions\ text \ Consider the following proof that makes use of usual Isar method combinators. \ lemma "P \ Q \ P" by ((rule impI, (erule conjE)?) | assumption)+ text \ It is clear that this compound method will be applicable in more cases than this proof alone. With the @{command method} command we can define a proof method that makes the above functionality available generally. \ method prop_solver\<^sub>1 = ((rule impI, (erule conjE)?) | assumption)+ lemma "P \ Q \ R \ P" by prop_solver\<^sub>1 text \ In this example, the facts \impI\ and \conjE\ are static. They are evaluated once when the method is defined and cannot be changed later. This makes the method stable in the sense of \<^emph>\static scoping\: naming another fact \impI\ in a later context won't affect the behaviour of \prop_solver\<^sub>1\. \ section \Term abstraction\ text \ Methods can also abstract over terms using the @{keyword_def "for"} keyword, optionally providing type constraints. For instance, the following proof method \intro_ex\ takes a term \<^term>\y\ of any type, which it uses to instantiate the \<^term>\x\-variable of \exI\ (existential introduction) before applying the result as a rule. The instantiation is performed here by Isar's @{attribute_ref "where"} attribute. If the current subgoal is to find a witness for the given predicate \<^term>\Q\, then this has the effect of committing to \<^term>\y\. \ method intro_ex for Q :: "'a \ bool" and y :: 'a = (rule exI ["where" P = Q and x = y]) text \ The term parameters \<^term>\y\ and \<^term>\Q\ can be used arbitrarily inside the method body, as part of attribute applications or arguments to other methods. The expression is type-checked as far as possible when the method is defined, however dynamic type errors can still occur when it is invoked (e.g.\ when terms are instantiated in a parameterized fact). Actual term arguments are supplied positionally, in the same order as in the method definition. \ lemma "P a \ \x. P x" by (intro_ex P a) section \Fact abstraction\ subsection \Named theorems\ text \ A \<^emph>\named theorem\ is a fact whose contents are produced dynamically within the current proof context. The Isar command @{command_ref "named_theorems"} declares a dynamic fact with a corresponding \<^emph>\attribute\ of the same name. This allows to maintain a collection of facts in the context as follows: \ named_theorems intros text \ So far \intros\ refers to the empty fact. Using the Isar command @{command_ref "declare"} we may apply declaration attributes to the context. Below we declare both \conjI\ and \impI\ as \intros\, adding them to the named theorem slot. \ declare conjI [intros] and impI [intros] text \ We can refer to named theorems as dynamic facts within a particular proof context, which are evaluated whenever the method is invoked. Instead of having facts hard-coded into the method, as in \prop_solver\<^sub>1\, we can instead refer to these named theorems. \ named_theorems elims declare conjE [elims] method prop_solver\<^sub>3 = ((rule intros, (erule elims)?) | assumption)+ lemma "P \ Q \ P" by prop_solver\<^sub>3 text \ Often these named theorems need to be augmented on the spot, when a method is invoked. The @{keyword_def "declares"} keyword in the signature of @{command method} adds the common method syntax \method decl: facts\ for each named theorem \decl\. \ method prop_solver\<^sub>4 declares intros elims = ((rule intros, (erule elims)?) | assumption)+ lemma "P \ (P \ Q) \ Q \ P" by (prop_solver\<^sub>4 elims: impE intros: conjI) subsection \Simple fact abstraction\ text \ The @{keyword "declares"} keyword requires that a corresponding dynamic fact has been declared with @{command_ref named_theorems}. This is useful for managing collections of facts which are to be augmented with declarations, but is overkill if we simply want to pass a fact to a method. We may use the @{keyword_def "uses"} keyword in the method header to provide a simple fact parameter. In contrast to @{keyword "declares"}, these facts are always implicitly empty unless augmented when the method is invoked. \ method rule_twice uses my_rule = (rule my_rule, rule my_rule) lemma "P \ Q \ (P \ Q) \ Q" by (rule_twice my_rule: conjI) section \Higher-order methods\ text \ The \<^emph>\structured concatenation\ combinator ``\method\<^sub>1 ; method\<^sub>2\'' was introduced in Isabelle2015, motivated by the development of Eisbach. It is similar to ``\method\<^sub>1, method\<^sub>2\'', but \method\<^sub>2\ is invoked on \<^emph>\all\ subgoals that have newly emerged from \method\<^sub>1\. This is useful to handle cases where the number of subgoals produced by a method is determined dynamically at run-time. \ method conj_with uses rule = (intro conjI ; intro rule) lemma assumes A: "P" shows "P \ P \ P" by (conj_with rule: A) text \ Method definitions may take other methods as arguments, and thus implement method combinators with prefix syntax. For example, to more usefully exploit Isabelle's backtracking, the explicit requirement that a method solve all produced subgoals is frequently useful. This can easily be written as a \<^emph>\higher-order method\ using ``\;\''. The @{keyword "methods"} keyword denotes method parameters that are other proof methods to be invoked by the method being defined. \ method solve methods m = (m ; fail) text \ Given some method-argument \m\, \solve \m\\ applies the method \m\ and then fails whenever \m\ produces any new unsolved subgoals --- i.e. when \m\ fails to completely discharge the goal it was applied to. \ section \Example\ text \ With these simple features we are ready to write our first non-trivial proof method. Returning to the first-order logic example, the following method definition applies various rules with their canonical methods. \ named_theorems subst method prop_solver declares intros elims subst = (assumption | (rule intros) | erule elims | subst subst | subst (asm) subst | (erule notE ; solve \prop_solver\))+ text \ The only non-trivial part above is the final alternative \(erule notE ; solve \prop_solver\)\. Here, in the case that all other alternatives fail, the method takes one of the assumptions \<^term>\\ P\ of the current goal and eliminates it with the rule \notE\, causing the goal to be proved to become \<^term>\P\. The method then recursively invokes itself on the remaining goals. The job of the recursive call is to demonstrate that there is a contradiction in the original assumptions (i.e.\ that \<^term>\P\ can be derived from them). Note this recursive invocation is applied with the @{method solve} method combinator to ensure that a contradiction will indeed be shown. In the case where a contradiction cannot be found, backtracking will occur and a different assumption \<^term>\\ Q\ will be chosen for elimination. Note that the recursive call to @{method prop_solver} does not have any parameters passed to it. Recall that fact parameters, e.g.\ \intros\, \elims\, and \subst\, are managed by declarations in the current proof context. They will therefore be passed to any recursive call to @{method prop_solver} and, more generally, any invocation of a method which declares these named theorems. \<^medskip> After declaring some standard rules to the context, the @{method prop_solver} becomes capable of solving non-trivial propositional tautologies. \ lemmas [intros] = conjI \ \@{thm conjI}\ impI \ \@{thm impI}\ disjCI \ \@{thm disjCI}\ iffI \ \@{thm iffI}\ notI \ \@{thm notI}\ lemmas [elims] = impCE \ \@{thm impCE}\ conjE \ \@{thm conjE}\ disjE \ \@{thm disjE}\ lemma "(A \ B) \ (A \ C) \ (B \ C) \ C" by prop_solver chapter \The match method \label{s:matching}\ text \ So far we have seen methods defined as simple combinations of other methods. Some familiar programming language concepts have been introduced (i.e.\ abstraction and recursion). The only control flow has been implicitly the result of backtracking. When designing more sophisticated proof methods this proves too restrictive and difficult to manage conceptually. To address this, we introduce the @{method_def match} method, which provides more direct access to the higher-order matching facility at the core of Isabelle. It is implemented as a separate proof method (in Isabelle/ML), and thus can be directly applied to proofs, however it is most useful when applied in the context of writing Eisbach method definitions. \<^medskip> The syntax diagram below refers to some syntactic categories that are - further defined in @{cite "isabelle-isar-ref"}. + further defined in \<^cite>\"isabelle-isar-ref"\. \<^rail>\ @@{method match} kind @'in' (pattern '\' @{syntax text} + '\') ; kind: (@'conclusion' | @'premises' ('(' 'local' ')')? | '(' term ')' | @{syntax thms}) ; pattern: fact_name? term args? \ (@'for' fixes)? ; fact_name: @{syntax name} @{syntax attributes}? ':' ; args: '(' (('multi' | 'cut' nat?) + ',') ')' \ Matching allows methods to introspect the goal state, and to implement more explicit control flow. In the basic case, a term or fact \ts\ is given to match against as a \<^emph>\match target\, along with a collection of pattern-method pairs \(p, m)\: roughly speaking, when the pattern \p\ matches any member of \ts\, the \<^emph>\inner\ method \m\ will be executed. \ lemma assumes X: "Q \ P" "Q" shows P by (match X in I: "Q \ P" and I': "Q" \ \insert mp [OF I I']\) text \ In this example we have a structured Isar proof, with the named assumption \X\ and a conclusion \<^term>\P\. With the match method we can find the local facts \<^term>\Q \ P\ and \<^term>\Q\, binding them to separately as \I\ and \I'\. We then specialize the modus-ponens rule @{thm mp [of Q P]} to these facts to solve the goal. \ section \Subgoal focus\ text\ In the previous example we were able to match against an assumption out of the Isar proof state. In general, however, proof subgoals can be \<^emph>\unstructured\, with goal parameters and premises arising from rule application. To address this, @{method match} uses \<^emph>\subgoal focusing\ to produce structured goals out of unstructured ones. In place of fact or term, we may give the keyword @{keyword_def "premises"} as the match target. This causes a subgoal focus on the first subgoal, lifting local goal parameters to fixed term variables and premises into hypothetical theorems. The match is performed against these theorems, naming them and binding them as appropriate. Similarly giving the keyword @{keyword_def "conclusion"} matches against the conclusion of the first subgoal. An unstructured version of the previous example can then be similarly solved through focusing. \ lemma "Q \ P \ Q \ P" by (match premises in I: "Q \ P" and I': "Q" \ \insert mp [OF I I']\) text \ Match variables may be specified by giving a list of @{keyword_ref "for"}-fixes after the pattern description. This marks those terms as bound variables, which may be used in the method body. \ lemma "Q \ P \ Q \ P" by (match premises in I: "Q \ A" and I': "Q" for A \ \match conclusion in A \ \insert mp [OF I I']\\) text \ In this example \<^term>\A\ is a match variable which is bound to \<^term>\P\ upon a successful match. The inner @{method match} then matches the now-bound \<^term>\A\ (bound to \<^term>\P\) against the conclusion (also \<^term>\P\), finally applying the specialized rule to solve the goal. Schematic terms like \?P\ may also be used to specify match variables, but the result of the match is not bound, and thus cannot be used in the inner method body. \<^medskip> In the following example we extract the predicate of an existentially quantified conclusion in the current subgoal and search the current premises for a matching fact. If both matches are successful, we then instantiate the existential introduction rule with both the witness and predicate, solving with the matched premise. \ method solve_ex = (match conclusion in "\x. Q x" for Q \ \match premises in U: "Q y" for y \ \rule exI [where P = Q and x = y, OF U]\\) text \ The first @{method match} matches the pattern \<^term>\\x. Q x\ against the current conclusion, binding the term \<^term>\Q\ in the inner match. Next the pattern \Q y\ is matched against all premises of the current subgoal. In this case \<^term>\Q\ is fixed and \<^term>\y\ may be instantiated. Once a match is found, the local fact \U\ is bound to the matching premise and the variable \<^term>\y\ is bound to the matching witness. The existential introduction rule \exI:\~@{thm exI} is then instantiated with \<^term>\y\ as the witness and \<^term>\Q\ as the predicate, with its proof obligation solved by the local fact U (using the Isar attribute @{attribute OF}). The following example is a trivial use of this method. \ lemma "halts p \ \x. halts x" by solve_ex subsection \Operating within a focus\ text \ Subgoal focusing provides a structured form of a subgoal, allowing for more expressive introspection of the goal state. This requires some consideration in order to be used effectively. When the keyword @{keyword "premises"} is given as the match target, the premises of the subgoal are lifted into hypothetical theorems, which can be found and named via match patterns. Additionally these premises are stripped from the subgoal, leaving only the conclusion. This renders them inaccessible to standard proof methods which operate on the premises, such as @{method frule} or @{method erule}. Naive usage of these methods within a match will most likely not function as the method author intended. \ method my_allE_bad for y :: 'a = (match premises in I: "\x :: 'a. ?Q x" \ \erule allE [where x = y]\) text \ Here we take a single parameter \<^term>\y\ and specialize the universal elimination rule (@{thm allE}) to it, then attempt to apply this specialized rule with @{method erule}. The method @{method erule} will attempt to unify with a universal quantifier in the premises that matches the type of \<^term>\y\. Since @{keyword "premises"} causes a focus, however, there are no subgoal premises to be found and thus @{method my_allE_bad} will always fail. If focusing instead left the premises in place, using methods like @{method erule} would lead to unintended behaviour, specifically during backtracking. In our example, @{method erule} could choose an alternate premise while backtracking, while leaving \I\ bound to the original match. In the case of more complex inner methods, where either \I\ or bound terms are used, this would almost certainly not be the intended behaviour. An alternative implementation would be to specialize the elimination rule to the bound term and apply it directly. \ method my_allE_almost for y :: 'a = (match premises in I: "\x :: 'a. ?Q x" \ \rule allE [where x = y, OF I]\) lemma "\x. P x \ P y" by (my_allE_almost y) text \ This method will insert a specialized duplicate of a universally quantified premise. Although this will successfully apply in the presence of such a premise, it is not likely the intended behaviour. Repeated application of this method will produce an infinite stream of duplicate specialized premises, due to the original premise never being removed. To address this, matched premises may be declared with the @{attribute thin} attribute. This will hide the premise from subsequent inner matches, and remove it from the list of premises when the inner method has finished and the subgoal is unfocused. It can be considered analogous to the existing \thin_tac\. To complete our example, the correct implementation of the method will @{attribute thin} the premise from the match and then apply it to the specialized elimination rule.\ method my_allE for y :: 'a = (match premises in I [thin]: "\x :: 'a. ?Q x" \ \rule allE [where x = y, OF I]\) lemma "\x. P x \ \x. Q x \ P y \ Q y" by (my_allE y)+ (rule conjI) subsubsection \Inner focusing\ text \ Premises are \<^emph>\accumulated\ for the purposes of subgoal focusing. In contrast to using standard methods like @{method frule} within focused match, another @{method match} will have access to all the premises of the outer focus. \ lemma "A \ B \ A \ B" by (match premises in H: A \ \intro conjI, rule H, match premises in H': B \ \rule H'\\) text \ In this example, the inner @{method match} can find the focused premise \<^term>\B\. In contrast, the @{method assumption} method would fail here due to \<^term>\B\ not being logically accessible. \ lemma "A \ A \ (B \ B)" by (match premises in H: A \ \intro conjI, rule H, rule impI, match premises (local) in A \ \fail\ \ H': B \ \rule H'\\) text \ In this example, the only premise that exists in the first focus is \<^term>\A\. Prior to the inner match, the rule \impI\ changes the goal \<^term>\B \ B\ into \<^term>\B \ B\. A standard premise match would also include \<^term>\A\ as an original premise of the outer match. The \local\ argument limits the match to newly focused premises. \ section \Attributes\ text \ Attributes may throw errors when applied to a given fact. For example, rule instantiation will fail if there is a type mismatch or if a given variable doesn't exist. Within a match or a method definition, it isn't generally possible to guarantee that applied attributes won't fail. For example, in the following method there is no guarantee that the two provided facts will necessarily compose. \ method my_compose uses rule1 rule2 = (rule rule1 [OF rule2]) text \ Some attributes (like @{attribute OF}) have been made partially Eisbach-aware. This means that they are able to form a closure despite not necessarily always being applicable. In the case of @{attribute OF}, it is up to the proof author to guard attribute application with an appropriate @{method match}, but there are still no static guarantees. In contrast to @{attribute OF}, the @{attribute "where"} and @{attribute of} attributes attempt to provide static guarantees that they will apply whenever possible. Within a match pattern for a fact, each outermost quantifier specifies the requirement that a matching fact must have a schematic variable at that point. This gives a corresponding name to this ``slot'' for the purposes of forming a static closure, allowing the @{attribute "where"} attribute to perform an instantiation at run-time. \ lemma assumes A: "Q \ False" shows "\ Q" by (match intros in X: "\P. (P \ False) \ \ P" \ \rule X [where P = Q, OF A]\) text \ Subgoal focusing converts the outermost quantifiers of premises into schematics when lifting them to hypothetical facts. This allows us to instantiate them with @{attribute "where"} when using an appropriate match pattern. \ lemma "(\x :: 'a. A x \ B x) \ A y \ B y" by (match premises in I: "\x :: 'a. ?P x \ ?Q x" \ \rule I [where x = y]\) text \ The @{attribute of} attribute behaves similarly. It is worth noting, however, that the positional instantiation of @{attribute of} occurs against the position of the variables as they are declared \<^emph>\in the match pattern\. \ lemma fixes A B and x :: 'a and y :: 'b assumes asm: "(\x y. A y x \ B x y )" shows "A y x \ B x y" by (match asm in I: "\(x :: 'a) (y :: 'b). ?P x y \ ?Q x y" \ \rule I [of x y]\) text \ In this example, the order of schematics in \asm\ is actually \?y ?x\, but we instantiate our matched rule in the opposite order. This is because the effective rule \<^term>\I\ was bound from the match, which declared the \<^typ>\'a\ slot first and the \<^typ>\'b\ slot second. To get the dynamic behaviour of @{attribute of} we can choose to invoke it \<^emph>\unchecked\. This avoids trying to do any type inference for the provided parameters, instead storing them as their most general type and doing type matching at run-time. This, like @{attribute OF}, will throw errors if the expected slots don't exist or there is a type mismatch. \ lemma fixes A B and x :: 'a and y :: 'b assumes asm: "\x y. A y x \ B x y" shows "A y x \ B x y" by (match asm in I: "PROP ?P" \ \rule I [of (unchecked) y x]\) text \ Attributes may be applied to matched facts directly as they are matched. Any declarations will therefore be applied in the context of the inner method, as well as any transformations to the rule. \ lemma "(\x :: 'a. A x \ B x) \ A y \ B y" by (match premises in I [of y, intros]: "\x :: 'a. ?P x \ ?Q x" \ \prop_solver\) text \ In this example, the pattern \\x :: 'a. ?P x \ ?Q x\ matches against the only premise, giving an appropriately typed slot for \<^term>\y\. After the match, the resulting rule is instantiated to \<^term>\y\ and then declared as an @{attribute intros} rule. This is then picked up by @{method prop_solver} to solve the goal. \ section \Multi-match \label{sec:multi}\ text \ In all previous examples, @{method match} was only ever searching for a single rule or premise. Each local fact would therefore always have a length of exactly one. We may, however, wish to find \<^emph>\all\ matching results. To achieve this, we can simply mark a given pattern with the \(multi)\ argument. \ lemma assumes asms: "A \ B" "A \ D" shows "(A \ B) \ (A \ D)" apply (match asms in I [intros]: "?P \ ?Q" \ \solves \prop_solver\\)? apply (match asms in I [intros]: "?P \ ?Q" (multi) \ \prop_solver\) done text \ In the first @{method match}, without the \(multi)\ argument, \<^term>\I\ is only ever be bound to one of the members of \asms\. This backtracks over both possibilities (see next section), however neither assumption in isolation is sufficient to solve to goal. The use of the @{method solves} combinator ensures that @{method prop_solver} has no effect on the goal when it doesn't solve it, and so the first match leaves the goal unchanged. In the second @{method match}, \I\ is bound to all of \asms\, declaring both results as \intros\. With these rules @{method prop_solver} is capable of solving the goal. Using for-fixed variables in patterns imposes additional constraints on the results. In all previous examples, the choice of using \?P\ or a for-fixed \<^term>\P\ only depended on whether or not \<^term>\P\ was mentioned in another pattern or the inner method. When using a multi-match, however, all for-fixed terms must agree in the results. \ lemma assumes asms: "A \ B" "A \ D" "D \ B" shows "(A \ B) \ (A \ D)" apply (match asms in I [intros]: "?P \ Q" (multi) for Q \ \solves \prop_solver\\)? apply (match asms in I [intros]: "P \ ?Q" (multi) for P \ \prop_solver\) done text \ Here we have two seemingly-equivalent applications of @{method match}, however only the second one is capable of solving the goal. The first @{method match} selects the first and third members of \asms\ (those that agree on their conclusion), which is not sufficient. The second @{method match} selects the first and second members of \asms\ (those that agree on their assumption), which is enough for @{method prop_solver} to solve the goal. \ section \Dummy patterns\ text \ Dummy patterns may be given as placeholders for unique schematics in patterns. They implicitly receive all currently bound variables as arguments, and are coerced into the \<^typ>\prop\ type whenever possible. For example, the trivial dummy pattern \_\ will match any proposition. In contrast, by default the pattern \?P\ is considered to have type \<^typ>\bool\. It will not bind anything with meta-logical connectives (e.g. \_ \ _\ or \_ &&& _\). \ lemma assumes asms: "A &&& B \ D" shows "(A \ B \ D)" by (match asms in I: _ \ \prop_solver intros: I conjunctionI\) section \Backtracking\ text \ Patterns are considered top-down, executing the inner method \m\ of the first pattern which is satisfied by the current match target. By default, matching performs extensive backtracking by attempting all valid variable and fact bindings according to the given pattern. In particular, all unifiers for a given pattern will be explored, as well as each matching fact. The inner method \m\ will be re-executed for each different variable/fact binding during backtracking. A successful match is considered a cut-point for backtracking. Specifically, once a match is made no other pattern-method pairs will be considered. The method \foo\ below fails for all goals that are conjunctions. Any such goal will match the first pattern, causing the second pattern (that would otherwise match all goals) to never be considered. \ method foo = (match conclusion in "?P \ ?Q" \ \fail\ \ "?R" \ \prop_solver\) text \ The failure of an inner method that is executed after a successful match will cause the entire match to fail. This distinction is important due to the pervasive use of backtracking. When a method is used in a combinator chain, its failure becomes significant because it signals previously applied methods to move to the next result. Therefore, it is necessary for @{method match} to not mask such failure. One can always rewrite a match using the combinators ``\?\'' and ``\|\'' to try subsequent patterns in the case of an inner-method failure. The following proof method, for example, always invokes @{method prop_solver} for all goals because its first alternative either never matches or (if it does match) always fails. \ method foo\<^sub>1 = (match conclusion in "?P \ ?Q" \ \fail\) | (match conclusion in "?R" \ \prop_solver\) subsection \Cut\ text \ Backtracking may be controlled more precisely by marking individual patterns as \cut\. This causes backtracking to not progress beyond this pattern: once a match is found no others will be considered. \ method foo\<^sub>2 = (match premises in I: "P \ Q" (cut) and I': "P \ ?U" for P Q \ \rule mp [OF I' I [THEN conjunct1]]\) text \ In this example, once a conjunction is found (\<^term>\P \ Q\), all possible implications of \<^term>\P\ in the premises are considered, evaluating the inner @{method rule} with each consequent. No other conjunctions will be considered, with method failure occurring once all implications of the form \P \ ?U\ have been explored. Here the left-right processing of individual patterns is important, as all patterns after of the cut will maintain their usual backtracking behaviour. \ lemma "A \ B \ A \ D \ A \ C \ C" by foo\<^sub>2 lemma "C \ D \ A \ B \ A \ C \ C" by (foo\<^sub>2 | prop_solver) text \ In this example, the first lemma is solved by \foo\<^sub>2\, by first picking \<^term>\A \ D\ for \I'\, then backtracking and ultimately succeeding after picking \<^term>\A \ C\. In the second lemma, however, \<^term>\C \ D\ is matched first, the second pattern in the match cannot be found and so the method fails, falling through to @{method prop_solver}. More precise control is also possible by giving a positive number \n\ as an argument to \cut\. This will limit the number of backtracking results of that match to be at most \n\. The match argument \(cut 1)\ is the same as simply \(cut)\. \ subsection \Multi-match revisited\ text \ A multi-match will produce a sequence of potential bindings for for-fixed variables, where each binding environment is the result of matching against at least one element from the match target. For each environment, the match result will be all elements of the match target which agree with the pattern under that environment. This can result in unexpected behaviour when giving very general patterns. \ lemma assumes asms: "\x. A x \ B x" "\y. A y \ C y" "\z. B z \ C z" shows "A x \ C x" by (match asms in I: "\x. P x \ ?Q x" (multi) for P \ \match (P) in "A" \ \fail\ \ _ \ \match I in "\x. A x \ B x" \ \fail\ \ _ \ \rule I\\\) text \ Intuitively it seems like this proof should fail to check. The first match result, which binds \<^term>\I\ to the first two members of \asms\, fails the second inner match due to binding \<^term>\P\ to \<^term>\A\. Backtracking then attempts to bind \<^term>\I\ to the third member of \asms\. This passes all inner matches, but fails when @{method rule} cannot successfully apply this to the current goal. After this, a valid match that is produced by the unifier is one which binds \<^term>\P\ to simply \\a. A ?x\. The first inner match succeeds because \\a. A ?x\ does not match \<^term>\A\. The next inner match succeeds because \<^term>\I\ has only been bound to the first member of \asms\. This is due to @{method match} considering \\a. A ?x\ and \\a. A ?y\ as distinct terms. The simplest way to address this is to explicitly disallow term bindings which we would consider invalid. \ method abs_used for P = (match (P) in "\a. ?P" \ \fail\ \ _ \ \-\) text \ This method has no effect on the goal state, but instead serves as a filter on the environment produced from match. \ section \Uncurrying\ text \ The @{method match} method is not aware of the logical content of match targets. Each pattern is simply matched against the shallow structure of a fact or term. Most facts are in \<^emph>\normal form\, which curries premises via meta-implication \_ \ _\. \ lemma assumes asms: "D \ B \ C" "D \ A" shows "D \ B \ C \ A" by (match asms in H: "D \ _" (multi) \ \prop_solver elims: H\) text \ For the first member of \asms\ the dummy pattern successfully matches against \<^term>\B \ C\ and so the proof is successful. \ lemma assumes asms: "A \ B \ C" "D \ C" shows "D \ (A \ B) \ C" apply (match asms in H: "_ \ C" (multi) \ \prop_solver elims: H\)(*<*)? apply (prop_solver elims: asms) done(*>*) text \ This proof will fail to solve the goal. Our match pattern will only match rules which have a single premise, and conclusion \<^term>\C\, so the first member of \asms\ is not bound and thus the proof fails. Matching a pattern of the form \<^term>\P \ Q\ against this fact will bind \<^term>\P\ to \<^term>\A\ and \<^term>\Q\ to \<^term>\B \ C\. Our pattern, with a concrete \<^term>\C\ in the conclusion, will fail to match this fact. To express our desired match, we may \<^emph>\uncurry\ our rules before matching against them. This forms a meta-conjunction of all premises in a fact, so that only one implication remains. For example the uncurried version of \<^term>\A \ B \ C\ is \<^term>\A &&& B \ C\. This will now match our desired pattern \_ \ C\, and can be \<^emph>\curried\ after the match to put it back into normal form. \ lemma assumes asms: "A \ B \ C" "D \ C" shows "D \ (A \ B) \ C" by (match asms [uncurry] in H [curry]: "_ \ C" (multi) \ \prop_solver elims: H\) section \Reverse matching\ text \ The @{method match} method only attempts to perform matching of the pattern against the match target. Specifically this means that it will not instantiate schematic terms in the match target. \ lemma assumes asms: "\x :: 'a. A x" shows "A y" apply (match asms in H: "A y" \ \rule H\)? apply (match asms in H: P for P \ \match ("A y") in P \ \rule H\\) done text \ In the first @{method match} we attempt to find a member of \asms\ which matches our goal precisely. This fails due to no such member existing. The second match reverses the role of the fact in the match, by first giving a general pattern \<^term>\P\. This bound pattern is then matched against \<^term>\A y\. In this case, \<^term>\P\ is bound to \A ?x\ and so it successfully matches. \ section \Type matching\ text \ The rule instantiation attributes @{attribute "where"} and @{attribute "of"} attempt to guarantee type-correctness wherever possible. This can require additional invocations of @{method match} in order to statically ensure that instantiation will succeed. \ lemma assumes asms: "\x :: 'a. A x" shows "A y" by (match asms in H: "\z :: 'b. P z" for P \ \match (y) in "y :: 'b" for y \ \rule H [where z = y]\\) text \ In this example the type \'b\ is matched to \'a\, however statically they are formally distinct types. The first match binds \'b\ while the inner match serves to coerce \<^term>\y\ into having the type \'b\. This allows the rule instantiation to successfully apply. \ chapter \Method development\ section \Tracing methods\ text \ Method tracing is supported by auxiliary print methods provided by \<^theory>\HOL-Eisbach.Eisbach_Tools\. These include @{method print_fact}, @{method print_term} and @{method print_type}. Whenever a print method is evaluated it leaves the goal unchanged and writes its argument as tracing output. Print methods can be combined with the @{method fail} method to investigate the backtracking behaviour of a method. \ lemma assumes asms: A B C D shows D apply (match asms in H: _ \ \print_fact H, fail\)(*<*)? apply (simp add: asms) done(*>*) text \ This proof will fail, but the tracing output will show the order that the assumptions are attempted. \ section \Integrating with Isabelle/ML\ subsubsection \Attributes\ text \ A custom rule attribute is a simple way to extend the functionality of Eisbach methods. The dummy rule attribute notation (\[[ _ ]]\) invokes the given attribute against a dummy fact and evaluates to the result of that attribute. When used as a match target, this can serve as an effective auxiliary function. \ attribute_setup get_split_rule = \Args.term >> (fn t => Thm.rule_attribute [] (fn context => fn _ => (case get_split_rule (Context.proof_of context) t of SOME thm => thm | NONE => Drule.dummy_thm)))\ text \ In this example, the new attribute @{attribute get_split_rule} lifts the ML function of the same name into an attribute. When applied to a case distinction over a datatype, it retrieves its corresponding split rule. We can then integrate this into a method that applies the split rule, first matching to ensure that fetching the rule was successful. \ (*<*)declare TrueI [intros](*>*) method splits = (match conclusion in "?P f" for f \ \match [[get_split_rule f]] in U: "(_ :: bool) = _" \ \rule U [THEN iffD2]\\) lemma "L \ [] \ case L of [] \ False | _ \ True" apply splits apply (prop_solver intros: allI) done text \ Here the new @{method splits} method transforms the goal to use only logical connectives: \<^term>\L = [] \ False \ (\x y. L = x # y \ True)\. This goal is then in a form solvable by @{method prop_solver} when given the universal quantifier introduction rule \allI\. \ end diff --git a/src/Doc/Eisbach/Preface.thy b/src/Doc/Eisbach/Preface.thy --- a/src/Doc/Eisbach/Preface.thy +++ b/src/Doc/Eisbach/Preface.thy @@ -1,42 +1,42 @@ (*:maxLineLen=78:*) theory Preface imports Base "HOL-Eisbach.Eisbach_Tools" begin text \ \<^emph>\Eisbach\ is a collection of tools which form the basis for defining new - proof methods in Isabelle/Isar~@{cite "Wenzel-PhD"}. It can be thought of as + proof methods in Isabelle/Isar~\<^cite>\"Wenzel-PhD"\. It can be thought of as a ``proof method language'', but is more precisely an infrastructure for defining new proof methods out of existing ones. The core functionality of Eisbach is provided by the Isar @{command method} command. Here users may define new methods by combining existing ones with the usual Isar syntax. These methods can be abstracted over terms, facts and other methods, as one might expect in any higher-order functional language. Additional functionality is provided by extending the space of methods and attributes. The new @{method match} method allows for explicit control-flow, by taking a match target and a list of pattern-method pairs. By using the functionality provided by Eisbach, additional support methods can be easily written. For example, the @{method catch} method, which provides basic try-catch functionality, only requires a few lines of ML. Eisbach enables users to implement automated proof tools conveniently via Isabelle/Isar syntax. This is in contrast to the traditional approach to use Isabelle/ML (via @{command_ref method_setup}), which poses a higher barrier-to-entry for casual users. \<^medskip> This manual is written for readers familiar with Isabelle/Isar, but not necessarily Isabelle/ML. It covers the usage of the @{command method} as well as the @{method match} method, as well as discussing their integration with existing Isar concepts such as @{command named_theorems}. These commands are provided by theory \<^theory>\HOL-Eisbach.Eisbach\: it needs to be imported by all Eisbach applications. Theory \<^theory>\HOL-Eisbach.Eisbach_Tools\ provides additional proof methods and attributes that are occasionally useful. \ end diff --git a/src/Doc/Functions/Functions.thy b/src/Doc/Functions/Functions.thy --- a/src/Doc/Functions/Functions.thy +++ b/src/Doc/Functions/Functions.thy @@ -1,1242 +1,1242 @@ (* Title: Doc/Functions/Functions.thy Author: Alexander Krauss, TU Muenchen Tutorial for function definitions with the new "function" package. *) theory Functions imports Main begin section \Function Definitions for Dummies\ text \ In most cases, defining a recursive function is just as simple as other definitions: \ fun fib :: "nat \ nat" where "fib 0 = 1" | "fib (Suc 0) = 1" | "fib (Suc (Suc n)) = fib n + fib (Suc n)" text \ The syntax is rather self-explanatory: We introduce a function by giving its name, its type, and a set of defining recursive equations. If we leave out the type, the most general type will be inferred, which can sometimes lead to surprises: Since both \<^term>\1::nat\ and \+\ are overloaded, we would end up with \fib :: nat \ 'a::{one,plus}\. \ text \ The function always terminates, since its argument gets smaller in every recursive call. Since HOL is a logic of total functions, termination is a fundamental requirement to prevent inconsistencies\footnote{From the \qt{definition} \f(n) = f(n) + 1\ we could prove \0 = 1\ by subtracting \f(n)\ on both sides.}. Isabelle tries to prove termination automatically when a definition is made. In \S\ref{termination}, we will look at cases where this fails and see what to do then. \ subsection \Pattern matching\ text \\label{patmatch} Like in functional programming, we can use pattern matching to define functions. At the moment we will only consider \emph{constructor patterns}, which only consist of datatype constructors and variables. Furthermore, patterns must be linear, i.e.\ all variables on the left hand side of an equation must be distinct. In \S\ref{genpats} we discuss more general pattern matching. If patterns overlap, the order of the equations is taken into account. The following function inserts a fixed element between any two elements of a list: \ fun sep :: "'a \ 'a list \ 'a list" where "sep a (x#y#xs) = x # a # sep a (y # xs)" | "sep a xs = xs" text \ Overlapping patterns are interpreted as \qt{increments} to what is already there: The second equation is only meant for the cases where the first one does not match. Consequently, Isabelle replaces it internally by the remaining cases, making the patterns disjoint: \ thm sep.simps text \@{thm [display] sep.simps[no_vars]}\ text \ \noindent The equations from function definitions are automatically used in simplification: \ lemma "sep 0 [1, 2, 3] = [1, 0, 2, 0, 3]" by simp subsection \Induction\ text \ Isabelle provides customized induction rules for recursive functions. These rules follow the recursive structure of the definition. Here is the rule @{thm [source] sep.induct} arising from the above definition of \<^const>\sep\: @{thm [display] sep.induct} We have a step case for list with at least two elements, and two base cases for the zero- and the one-element list. Here is a simple proof about \<^const>\sep\ and \<^const>\map\ \ lemma "map f (sep x ys) = sep (f x) (map f ys)" apply (induct x ys rule: sep.induct) text \ We get three cases, like in the definition. @{subgoals [display]} \ apply auto done text \ With the \cmd{fun} command, you can define about 80\% of the functions that occur in practice. The rest of this tutorial explains the remaining 20\%. \ section \fun vs.\ function\ text \ The \cmd{fun} command provides a convenient shorthand notation for simple function definitions. In this mode, Isabelle tries to solve all the necessary proof obligations automatically. If any proof fails, the definition is rejected. This can either mean that the definition is indeed faulty, or that the default proof procedures are just not smart enough (or rather: not designed) to handle the definition. By expanding the abbreviation to the more verbose \cmd{function} command, these proof obligations become visible and can be analyzed or solved manually. The expansion from \cmd{fun} to \cmd{function} is as follows: \end{isamarkuptext} \[\left[\;\begin{minipage}{0.25\textwidth}\vspace{6pt} \cmd{fun} \f :: \\\\% \cmd{where}\\% \hspace*{2ex}{\it equations}\\% \hspace*{2ex}\vdots\vspace*{6pt} \end{minipage}\right] \quad\equiv\quad \left[\;\begin{minipage}{0.48\textwidth}\vspace{6pt} \cmd{function} \(\\cmd{sequential}\) f :: \\\\% \cmd{where}\\% \hspace*{2ex}{\it equations}\\% \hspace*{2ex}\vdots\\% \cmd{by} \pat_completeness auto\\\% \cmd{termination by} \lexicographic_order\\vspace{6pt} \end{minipage} \right]\] \begin{isamarkuptext} \vspace*{1em} \noindent Some details have now become explicit: \begin{enumerate} \item The \cmd{sequential} option enables the preprocessing of pattern overlaps which we already saw. Without this option, the equations must already be disjoint and complete. The automatic completion only works with constructor patterns. \item A function definition produces a proof obligation which expresses completeness and compatibility of patterns (we talk about this later). The combination of the methods \pat_completeness\ and \auto\ is used to solve this proof obligation. \item A termination proof follows the definition, started by the \cmd{termination} command. This will be explained in \S\ref{termination}. \end{enumerate} Whenever a \cmd{fun} command fails, it is usually a good idea to expand the syntax to the more verbose \cmd{function} form, to see what is actually going on. \ section \Termination\ text \\label{termination} The method \lexicographic_order\ is the default method for termination proofs. It can prove termination of a certain class of functions by searching for a suitable lexicographic combination of size measures. Of course, not all functions have such a simple termination argument. For them, we can specify the termination relation manually. \ subsection \The {\tt relation} method\ text\ Consider the following function, which sums up natural numbers up to \N\, using a counter \i\: \ function sum :: "nat \ nat \ nat" where "sum i N = (if i > N then 0 else i + sum (Suc i) N)" by pat_completeness auto text \ \noindent The \lexicographic_order\ method fails on this example, because none of the arguments decreases in the recursive call, with respect to the standard size ordering. To prove termination manually, we must provide a custom wellfounded relation. The termination argument for \sum\ is based on the fact that the \emph{difference} between \i\ and \N\ gets smaller in every step, and that the recursion stops when \i\ is greater than \N\. Phrased differently, the expression \N + 1 - i\ always decreases. We can use this expression as a measure function suitable to prove termination. \ termination sum apply (relation "measure (\(i,N). N + 1 - i)") text \ The \cmd{termination} command sets up the termination goal for the specified function \sum\. If the function name is omitted, it implicitly refers to the last function definition. The \relation\ method takes a relation of type \<^typ>\('a \ 'a) set\, where \<^typ>\'a\ is the argument type of the function. If the function has multiple curried arguments, then these are packed together into a tuple, as it happened in the above example. The predefined function @{term[source] "measure :: ('a \ nat) \ ('a \ 'a) set"} constructs a wellfounded relation from a mapping into the natural numbers (a \emph{measure function}). After the invocation of \relation\, we must prove that (a) the relation we supplied is wellfounded, and (b) that the arguments of recursive calls indeed decrease with respect to the relation: @{subgoals[display,indent=0]} These goals are all solved by \auto\: \ apply auto done text \ Let us complicate the function a little, by adding some more recursive calls: \ function foo :: "nat \ nat \ nat" where "foo i N = (if i > N then (if N = 0 then 0 else foo 0 (N - 1)) else i + foo (Suc i) N)" by pat_completeness auto text \ When \i\ has reached \N\, it starts at zero again and \N\ is decremented. This corresponds to a nested loop where one index counts up and the other down. Termination can be proved using a lexicographic combination of two measures, namely the value of \N\ and the above difference. The \<^const>\measures\ combinator generalizes \measure\ by taking a list of measure functions. \ termination by (relation "measures [\(i, N). N, \(i,N). N + 1 - i]") auto subsection \How \lexicographic_order\ works\ (*fun fails :: "nat \ nat list \ nat" where "fails a [] = a" | "fails a (x#xs) = fails (x + a) (x # xs)" *) text \ To see how the automatic termination proofs work, let's look at an example where it fails\footnote{For a detailed discussion of the - termination prover, see @{cite bulwahnKN07}}: + termination prover, see \<^cite>\bulwahnKN07\}: \end{isamarkuptext} \cmd{fun} \fails :: "nat \ nat list \ nat"\\\% \cmd{where}\\% \hspace*{2ex}\"fails a [] = a"\\\% |\hspace*{1.5ex}\"fails a (x#xs) = fails (x + a) (x#xs)"\\\ \begin{isamarkuptext} \noindent Isabelle responds with the following error: \begin{isabelle} *** Unfinished subgoals:\newline *** (a, 1, <):\newline *** \ 1.~\\x. x = 0\\newline *** (a, 1, <=):\newline *** \ 1.~False\newline *** (a, 2, <):\newline *** \ 1.~False\newline *** Calls:\newline *** a) \(a, x # xs) -->> (x + a, x # xs)\\newline *** Measures:\newline *** 1) \\x. size (fst x)\\newline *** 2) \\x. size (snd x)\\newline *** Result matrix:\newline *** \ \ \ \ 1\ \ 2 \newline *** a: ? <= \newline *** Could not find lexicographic termination order.\newline *** At command "fun".\newline \end{isabelle} \ text \ The key to this error message is the matrix at the bottom. The rows of that matrix correspond to the different recursive calls (In our case, there is just one). The columns are the function's arguments (expressed through different measure functions, which map the argument tuple to a natural number). The contents of the matrix summarize what is known about argument descents: The second argument has a weak descent (\<=\) at the recursive call, and for the first argument nothing could be proved, which is expressed by \?\. In general, there are the values \<\, \<=\ and \?\. For the failed proof attempts, the unfinished subgoals are also printed. Looking at these will often point to a missing lemma. \ subsection \The \size_change\ method\ text \ Some termination goals that are beyond the powers of \lexicographic_order\ can be solved automatically by the more powerful \size_change\ method, which uses a variant of the size-change principle, together with some other techniques. While the details are discussed - elsewhere @{cite krauss_phd}, + elsewhere \<^cite>\krauss_phd\, here are a few typical situations where \lexicographic_order\ has difficulties and \size_change\ may be worth a try: \begin{itemize} \item Arguments are permuted in a recursive call. \item Several mutually recursive functions with multiple arguments. \item Unusual control flow (e.g., when some recursive calls cannot occur in sequence). \end{itemize} Loading the theory \Multiset\ makes the \size_change\ method a bit stronger: it can then use multiset orders internally. \ subsection \Configuring simplification rules for termination proofs\ text \ Since both \lexicographic_order\ and \size_change\ rely on the simplifier internally, there can sometimes be the need for adding additional simp rules to them. This can be done either as arguments to the methods themselves, or globally via the theorem attribute \termination_simp\, which is useful in rare cases. \ section \Mutual Recursion\ text \ If two or more functions call one another mutually, they have to be defined in one step. Here are \even\ and \odd\: \ function even :: "nat \ bool" and odd :: "nat \ bool" where "even 0 = True" | "odd 0 = False" | "even (Suc n) = odd n" | "odd (Suc n) = even n" by pat_completeness auto text \ To eliminate the mutual dependencies, Isabelle internally creates a single function operating on the sum type \<^typ>\nat + nat\. Then, \<^const>\even\ and \<^const>\odd\ are defined as projections. Consequently, termination has to be proved simultaneously for both functions, by specifying a measure on the sum type: \ termination by (relation "measure (\x. case x of Inl n \ n | Inr n \ n)") auto text \ We could also have used \lexicographic_order\, which supports mutual recursive termination proofs to a certain extent. \ subsection \Induction for mutual recursion\ text \ When functions are mutually recursive, proving properties about them generally requires simultaneous induction. The induction rule @{thm [source] "even_odd.induct"} generated from the above definition reflects this. Let us prove something about \<^const>\even\ and \<^const>\odd\: \ lemma even_odd_mod2: "even n = (n mod 2 = 0)" "odd n = (n mod 2 = 1)" text \ We apply simultaneous induction, specifying the induction variable for both goals, separated by \cmd{and}:\ apply (induct n and n rule: even_odd.induct) text \ We get four subgoals, which correspond to the clauses in the definition of \<^const>\even\ and \<^const>\odd\: @{subgoals[display,indent=0]} Simplification solves the first two goals, leaving us with two statements about the \mod\ operation to prove: \ apply simp_all text \ @{subgoals[display,indent=0]} \noindent These can be handled by Isabelle's arithmetic decision procedures. \ apply arith apply arith done text \ In proofs like this, the simultaneous induction is really essential: Even if we are just interested in one of the results, the other one is necessary to strengthen the induction hypothesis. If we leave out the statement about \<^const>\odd\ and just write \<^term>\True\ instead, the same proof fails: \ lemma failed_attempt: "even n = (n mod 2 = 0)" "True" apply (induct n rule: even_odd.induct) text \ \noindent Now the third subgoal is a dead end, since we have no useful induction hypothesis available: @{subgoals[display,indent=0]} \ oops section \Elimination\ text \ A definition of function \f\ gives rise to two kinds of elimination rules. Rule \f.cases\ simply describes case analysis according to the patterns used in the definition: \ fun list_to_option :: "'a list \ 'a option" where "list_to_option [x] = Some x" | "list_to_option _ = None" thm list_to_option.cases text \ @{thm[display] list_to_option.cases} Note that this rule does not mention the function at all, but only describes the cases used for defining it. In contrast, the rule @{thm[source] list_to_option.elims} also tell us what the function value will be in each case: \ thm list_to_option.elims text \ @{thm[display] list_to_option.elims} \noindent This lets us eliminate an assumption of the form \<^prop>\list_to_option xs = y\ and replace it with the two cases, e.g.: \ lemma "list_to_option xs = y \ P" proof (erule list_to_option.elims) fix x assume "xs = [x]" "y = Some x" thus P sorry next assume "xs = []" "y = None" thus P sorry next fix a b xs' assume "xs = a # b # xs'" "y = None" thus P sorry qed text \ Sometimes it is convenient to derive specialized versions of the \elim\ rules above and keep them around as facts explicitly. For example, it is natural to show that if \<^prop>\list_to_option xs = Some y\, then \<^term>\xs\ must be a singleton. The command \cmd{fun\_cases} derives such facts automatically, by instantiating and simplifying the general elimination rules given some pattern: \ fun_cases list_to_option_SomeE[elim]: "list_to_option xs = Some y" thm list_to_option_SomeE text \ @{thm[display] list_to_option_SomeE} \ section \General pattern matching\ text\\label{genpats}\ subsection \Avoiding automatic pattern splitting\ text \ Up to now, we used pattern matching only on datatypes, and the patterns were always disjoint and complete, and if they weren't, they were made disjoint automatically like in the definition of \<^const>\sep\ in \S\ref{patmatch}. This automatic splitting can significantly increase the number of equations involved, and this is not always desirable. The following example shows the problem: Suppose we are modeling incomplete knowledge about the world by a three-valued datatype, which has values \<^term>\T\, \<^term>\F\ and \<^term>\X\ for true, false and uncertain propositions, respectively. \ datatype P3 = T | F | X text \\noindent Then the conjunction of such values can be defined as follows:\ fun And :: "P3 \ P3 \ P3" where "And T p = p" | "And p T = p" | "And p F = F" | "And F p = F" | "And X X = X" text \ This definition is useful, because the equations can directly be used as simplification rules. But the patterns overlap: For example, the expression \<^term>\And T T\ is matched by both the first and the second equation. By default, Isabelle makes the patterns disjoint by splitting them up, producing instances: \ thm And.simps text \ @{thm[indent=4] And.simps} \vspace*{1em} \noindent There are several problems with this: \begin{enumerate} \item If the datatype has many constructors, there can be an explosion of equations. For \<^const>\And\, we get seven instead of five equations, which can be tolerated, but this is just a small example. \item Since splitting makes the equations \qt{less general}, they do not always match in rewriting. While the term \<^term>\And x F\ can be simplified to \<^term>\F\ with the original equations, a (manual) case split on \<^term>\x\ is now necessary. \item The splitting also concerns the induction rule @{thm [source] "And.induct"}. Instead of five premises it now has seven, which means that our induction proofs will have more cases. \item In general, it increases clarity if we get the same definition back which we put in. \end{enumerate} If we do not want the automatic splitting, we can switch it off by leaving out the \cmd{sequential} option. However, we will have to prove that our pattern matching is consistent\footnote{This prevents us from defining something like \<^term>\f x = True\ and \<^term>\f x = False\ simultaneously.}: \ function And2 :: "P3 \ P3 \ P3" where "And2 T p = p" | "And2 p T = p" | "And2 p F = F" | "And2 F p = F" | "And2 X X = X" text \ \noindent Now let's look at the proof obligations generated by a function definition. In this case, they are: @{subgoals[display,indent=0]}\vspace{-1.2em}\hspace{3cm}\vdots\vspace{1.2em} The first subgoal expresses the completeness of the patterns. It has the form of an elimination rule and states that every \<^term>\x\ of the function's input type must match at least one of the patterns\footnote{Completeness could be equivalently stated as a disjunction of existential statements: \<^term>\(\p. x = (T, p)) \ (\p. x = (p, T)) \ (\p. x = (p, F)) \ (\p. x = (F, p)) \ (x = (X, X))\, and you can use the method \atomize_elim\ to get that form instead.}. If the patterns just involve datatypes, we can solve it with the \pat_completeness\ method: \ apply pat_completeness text \ The remaining subgoals express \emph{pattern compatibility}. We do allow that an input value matches multiple patterns, but in this case, the result (i.e.~the right hand sides of the equations) must also be equal. For each pair of two patterns, there is one such subgoal. Usually this needs injectivity of the constructors, which is used automatically by \auto\. \ by auto termination by (relation "{}") simp subsection \Non-constructor patterns\ text \ Most of Isabelle's basic types take the form of inductive datatypes, and usually pattern matching works on the constructors of such types. However, this need not be always the case, and the \cmd{function} command handles other kind of patterns, too. One well-known instance of non-constructor patterns are so-called \emph{$n+k$-patterns}, which are a little controversial in the functional programming world. Here is the initial fibonacci example with $n+k$-patterns: \ function fib2 :: "nat \ nat" where "fib2 0 = 1" | "fib2 1 = 1" | "fib2 (n + 2) = fib2 n + fib2 (Suc n)" text \ This kind of matching is again justified by the proof of pattern completeness and compatibility. The proof obligation for pattern completeness states that every natural number is either \<^term>\0::nat\, \<^term>\1::nat\ or \<^term>\n + (2::nat)\: @{subgoals[display,indent=0,goals_limit=1]} This is an arithmetic triviality, but unfortunately the \arith\ method cannot handle this specific form of an elimination rule. However, we can use the method \atomize_elim\ to do an ad-hoc conversion to a disjunction of existentials, which can then be solved by the arithmetic decision procedure. Pattern compatibility and termination are automatic as usual. \ apply atomize_elim apply arith apply auto done termination by lexicographic_order text \ We can stretch the notion of pattern matching even more. The following function is not a sensible functional program, but a perfectly valid mathematical definition: \ function ev :: "nat \ bool" where "ev (2 * n) = True" | "ev (2 * n + 1) = False" apply atomize_elim by arith+ termination by (relation "{}") simp text \ This general notion of pattern matching gives you a certain freedom in writing down specifications. However, as always, such freedom should be used with care: If we leave the area of constructor patterns, we have effectively departed from the world of functional programming. This means that it is no longer possible to use the code generator, and expect it to generate ML code for our definitions. Also, such a specification might not work very well together with simplification. Your mileage may vary. \ subsection \Conditional equations\ text \ The function package also supports conditional equations, which are similar to guards in a language like Haskell. Here is Euclid's algorithm written with conditional patterns\footnote{Note that the patterns are also overlapping in the base case}: \ function gcd :: "nat \ nat \ nat" where "gcd x 0 = x" | "gcd 0 y = y" | "x < y \ gcd (Suc x) (Suc y) = gcd (Suc x) (y - x)" | "\ x < y \ gcd (Suc x) (Suc y) = gcd (x - y) (Suc y)" by (atomize_elim, auto, arith) termination by lexicographic_order text \ By now, you can probably guess what the proof obligations for the pattern completeness and compatibility look like. Again, functions with conditional patterns are not supported by the code generator. \ subsection \Pattern matching on strings\ text \ As strings (as lists of characters) are normal datatypes, pattern matching on them is possible, but somewhat problematic. Consider the following definition: \end{isamarkuptext} \noindent\cmd{fun} \check :: "string \ bool"\\\% \cmd{where}\\% \hspace*{2ex}\"check (''good'') = True"\\\% \| "check s = False"\ \begin{isamarkuptext} \noindent An invocation of the above \cmd{fun} command does not terminate. What is the problem? Strings are lists of characters, and characters are a datatype with a lot of constructors. Splitting the catch-all pattern thus leads to an explosion of cases, which cannot be handled by Isabelle. There are two things we can do here. Either we write an explicit \if\ on the right hand side, or we can use conditional patterns: \ function check :: "string \ bool" where "check (''good'') = True" | "s \ ''good'' \ check s = False" by auto termination by (relation "{}") simp section \Partiality \label{sec:partiality}\ text \ In HOL, all functions are total. A function \<^term>\f\ applied to \<^term>\x\ always has the value \<^term>\f x\, and there is no notion of undefinedness. This is why we have to do termination proofs when defining functions: The proof justifies that the function can be defined by wellfounded recursion. However, the \cmd{function} package does support partiality to a certain extent. Let's look at the following function which looks for a zero of a given function f. \ function (*<*)(domintros)(*>*)findzero :: "(nat \ nat) \ nat \ nat" where "findzero f n = (if f n = 0 then n else findzero f (Suc n))" by pat_completeness auto text \ \noindent Clearly, any attempt of a termination proof must fail. And without that, we do not get the usual rules \findzero.simps\ and \findzero.induct\. So what was the definition good for at all? \ subsection \Domain predicates\ text \ The trick is that Isabelle has not only defined the function \<^const>\findzero\, but also a predicate \<^term>\findzero_dom\ that characterizes the values where the function terminates: the \emph{domain} of the function. If we treat a partial function just as a total function with an additional domain predicate, we can derive simplification and induction rules as we do for total functions. They are guarded by domain conditions and are called \psimps\ and \pinduct\: \ text \ \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.psimps}\end{minipage} \hfill(@{thm [source] "findzero.psimps"}) \vspace{1em} \noindent\begin{minipage}{0.79\textwidth}@{thm[display,margin=85] findzero.pinduct}\end{minipage} \hfill(@{thm [source] "findzero.pinduct"}) \ text \ Remember that all we are doing here is use some tricks to make a total function appear as if it was partial. We can still write the term \<^term>\findzero (\x. 1) 0\ and like any other term of type \<^typ>\nat\ it is equal to some natural number, although we might not be able to find out which one. The function is \emph{underdefined}. But it is defined enough to prove something interesting about it. We can prove that if \<^term>\findzero f n\ terminates, it indeed returns a zero of \<^term>\f\: \ lemma findzero_zero: "findzero_dom (f, n) \ f (findzero f n) = 0" text \\noindent We apply induction as usual, but using the partial induction rule:\ apply (induct f n rule: findzero.pinduct) text \\noindent This gives the following subgoals: @{subgoals[display,indent=0]} \noindent The hypothesis in our lemma was used to satisfy the first premise in the induction rule. However, we also get \<^term>\findzero_dom (f, n)\ as a local assumption in the induction step. This allows unfolding \<^term>\findzero f n\ using the \psimps\ rule, and the rest is trivial. \ apply (simp add: findzero.psimps) done text \ Proofs about partial functions are often not harder than for total functions. Fig.~\ref{findzero_isar} shows a slightly more complicated proof written in Isar. It is verbose enough to show how partiality comes into play: From the partial induction, we get an additional domain condition hypothesis. Observe how this condition is applied when calls to \<^term>\findzero\ are unfolded. \ text_raw \ \begin{figure} \hrule\vspace{6pt} \begin{minipage}{0.8\textwidth} \isabellestyle{it} \isastyle\isamarkuptrue \ lemma "\findzero_dom (f, n); x \ {n ..< findzero f n}\ \ f x \ 0" proof (induct rule: findzero.pinduct) fix f n assume dom: "findzero_dom (f, n)" and IH: "\f n \ 0; x \ {Suc n ..< findzero f (Suc n)}\ \ f x \ 0" and x_range: "x \ {n ..< findzero f n}" have "f n \ 0" proof assume "f n = 0" with dom have "findzero f n = n" by (simp add: findzero.psimps) with x_range show False by auto qed from x_range have "x = n \ x \ {Suc n ..< findzero f n}" by auto thus "f x \ 0" proof assume "x = n" with \f n \ 0\ show ?thesis by simp next assume "x \ {Suc n ..< findzero f n}" with dom and \f n \ 0\ have "x \ {Suc n ..< findzero f (Suc n)}" by (simp add: findzero.psimps) with IH and \f n \ 0\ show ?thesis by simp qed qed text_raw \ \isamarkupfalse\isabellestyle{tt} \end{minipage}\vspace{6pt}\hrule \caption{A proof about a partial function}\label{findzero_isar} \end{figure} \ subsection \Partial termination proofs\ text \ Now that we have proved some interesting properties about our function, we should turn to the domain predicate and see if it is actually true for some values. Otherwise we would have just proved lemmas with \<^term>\False\ as a premise. Essentially, we need some introduction rules for \findzero_dom\. The function package can prove such domain introduction rules automatically. But since they are not used very often (they are almost never needed if the function is total), this functionality is disabled by default for efficiency reasons. So we have to go back and ask for them explicitly by passing the \(domintros)\ option to the function package: \vspace{1ex} \noindent\cmd{function} \(domintros) findzero :: "(nat \ nat) \ nat \ nat"\\\% \cmd{where}\isanewline% \ \ \ldots\\ \noindent Now the package has proved an introduction rule for \findzero_dom\: \ thm findzero.domintros text \ @{thm[display] findzero.domintros} Domain introduction rules allow to show that a given value lies in the domain of a function, if the arguments of all recursive calls are in the domain as well. They allow to do a \qt{single step} in a termination proof. Usually, you want to combine them with a suitable induction principle. Since our function increases its argument at recursive calls, we need an induction principle which works \qt{backwards}. We will use @{thm [source] inc_induct}, which allows to do induction from a fixed number \qt{downwards}: \begin{center}@{thm inc_induct}\hfill(@{thm [source] "inc_induct"})\end{center} Figure \ref{findzero_term} gives a detailed Isar proof of the fact that \findzero\ terminates if there is a zero which is greater or equal to \<^term>\n\. First we derive two useful rules which will solve the base case and the step case of the induction. The induction is then straightforward, except for the unusual induction principle. \ text_raw \ \begin{figure} \hrule\vspace{6pt} \begin{minipage}{0.8\textwidth} \isabellestyle{it} \isastyle\isamarkuptrue \ lemma findzero_termination: assumes "x \ n" and "f x = 0" shows "findzero_dom (f, n)" proof - have base: "findzero_dom (f, x)" by (rule findzero.domintros) (simp add:\f x = 0\) have step: "\i. findzero_dom (f, Suc i) \ findzero_dom (f, i)" by (rule findzero.domintros) simp from \x \ n\ show ?thesis proof (induct rule:inc_induct) show "findzero_dom (f, x)" by (rule base) next fix i assume "findzero_dom (f, Suc i)" thus "findzero_dom (f, i)" by (rule step) qed qed text_raw \ \isamarkupfalse\isabellestyle{tt} \end{minipage}\vspace{6pt}\hrule \caption{Termination proof for \findzero\}\label{findzero_term} \end{figure} \ text \ Again, the proof given in Fig.~\ref{findzero_term} has a lot of detail in order to explain the principles. Using more automation, we can also have a short proof: \ lemma findzero_termination_short: assumes zero: "x >= n" assumes [simp]: "f x = 0" shows "findzero_dom (f, n)" using zero by (induct rule:inc_induct) (auto intro: findzero.domintros) text \ \noindent It is simple to combine the partial correctness result with the termination lemma: \ lemma findzero_total_correctness: "f x = 0 \ f (findzero f 0) = 0" by (blast intro: findzero_zero findzero_termination) subsection \Definition of the domain predicate\ text \ Sometimes it is useful to know what the definition of the domain predicate looks like. Actually, \findzero_dom\ is just an abbreviation: @{abbrev[display] findzero_dom} The domain predicate is the \emph{accessible part} of a relation \<^const>\findzero_rel\, which was also created internally by the function package. \<^const>\findzero_rel\ is just a normal inductive predicate, so we can inspect its definition by looking at the introduction rules @{thm [source] findzero_rel.intros}. In our case there is just a single rule: @{thm[display] findzero_rel.intros} The predicate \<^const>\findzero_rel\ describes the \emph{recursion relation} of the function definition. The recursion relation is a binary relation on the arguments of the function that relates each argument to its recursive calls. In general, there is one introduction rule for each recursive call. The predicate \<^term>\Wellfounded.accp findzero_rel\ is the accessible part of that relation. An argument belongs to the accessible part, if it can be reached in a finite number of steps (cf.~its definition in \Wellfounded.thy\). Since the domain predicate is just an abbreviation, you can use lemmas for \<^const>\Wellfounded.accp\ and \<^const>\findzero_rel\ directly. Some lemmas which are occasionally useful are @{thm [source] accpI}, @{thm [source] accp_downward}, and of course the introduction and elimination rules for the recursion relation @{thm [source] "findzero_rel.intros"} and @{thm [source] "findzero_rel.cases"}. \ section \Nested recursion\ text \ Recursive calls which are nested in one another frequently cause complications, since their termination proof can depend on a partial correctness property of the function itself. As a small example, we define the \qt{nested zero} function: \ function nz :: "nat \ nat" where "nz 0 = 0" | "nz (Suc n) = nz (nz n)" by pat_completeness auto text \ If we attempt to prove termination using the identity measure on naturals, this fails: \ termination apply (relation "measure (\n. n)") apply auto text \ We get stuck with the subgoal @{subgoals[display]} Of course this statement is true, since we know that \<^const>\nz\ is the zero function. And in fact we have no problem proving this property by induction. \ (*<*)oops(*>*) lemma nz_is_zero: "nz_dom n \ nz n = 0" by (induct rule:nz.pinduct) (auto simp: nz.psimps) text \ We formulate this as a partial correctness lemma with the condition \<^term>\nz_dom n\. This allows us to prove it with the \pinduct\ rule before we have proved termination. With this lemma, the termination proof works as expected: \ termination by (relation "measure (\n. n)") (auto simp: nz_is_zero) text \ As a general strategy, one should prove the statements needed for termination as a partial property first. Then they can be used to do the termination proof. This also works for less trivial examples. Figure \ref{f91} defines the 91-function, a well-known challenge problem due to John McCarthy, and proves its termination. \ text_raw \ \begin{figure} \hrule\vspace{6pt} \begin{minipage}{0.8\textwidth} \isabellestyle{it} \isastyle\isamarkuptrue \ function f91 :: "nat \ nat" where "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))" by pat_completeness auto lemma f91_estimate: assumes trm: "f91_dom n" shows "n < f91 n + 11" using trm by induct (auto simp: f91.psimps) termination proof let ?R = "measure (\x. 101 - x)" show "wf ?R" .. fix n :: nat assume "\ 100 < n" \ \Assumptions for both calls\ thus "(n + 11, n) \ ?R" by simp \ \Inner call\ assume inner_trm: "f91_dom (n + 11)" \ \Outer call\ with f91_estimate have "n + 11 < f91 (n + 11) + 11" . with \\ 100 < n\ show "(f91 (n + 11), n) \ ?R" by simp qed text_raw \ \isamarkupfalse\isabellestyle{tt} \end{minipage} \vspace{6pt}\hrule \caption{McCarthy's 91-function}\label{f91} \end{figure} \ section \Higher-Order Recursion\ text \ Higher-order recursion occurs when recursive calls are passed as arguments to higher-order combinators such as \<^const>\map\, \<^term>\filter\ etc. As an example, imagine a datatype of n-ary trees: \ datatype 'a tree = Leaf 'a | Branch "'a tree list" text \\noindent We can define a function which swaps the left and right subtrees recursively, using the list functions \<^const>\rev\ and \<^const>\map\:\ fun mirror :: "'a tree \ 'a tree" where "mirror (Leaf n) = Leaf n" | "mirror (Branch l) = Branch (rev (map mirror l))" text \ Although the definition is accepted without problems, let us look at the termination proof: \ termination proof text \ As usual, we have to give a wellfounded relation, such that the arguments of the recursive calls get smaller. But what exactly are the arguments of the recursive calls when mirror is given as an argument to \<^const>\map\? Isabelle gives us the subgoals @{subgoals[display,indent=0]} So the system seems to know that \<^const>\map\ only applies the recursive call \<^term>\mirror\ to elements of \<^term>\l\, which is essential for the termination proof. This knowledge about \<^const>\map\ is encoded in so-called congruence rules, which are special theorems known to the \cmd{function} command. The rule for \<^const>\map\ is @{thm[display] map_cong} You can read this in the following way: Two applications of \<^const>\map\ are equal, if the list arguments are equal and the functions coincide on the elements of the list. This means that for the value \<^term>\map f l\ we only have to know how \<^term>\f\ behaves on the elements of \<^term>\l\. Usually, one such congruence rule is needed for each higher-order construct that is used when defining new functions. In fact, even basic functions like \<^const>\If\ and \<^const>\Let\ are handled by this mechanism. The congruence rule for \<^const>\If\ states that the \then\ branch is only relevant if the condition is true, and the \else\ branch only if it is false: @{thm[display] if_cong} Congruence rules can be added to the function package by giving them the \<^term>\fundef_cong\ attribute. The constructs that are predefined in Isabelle, usually come with the respective congruence rules. But if you define your own higher-order functions, you may have to state and prove the required congruence rules yourself, if you want to use your functions in recursive definitions. \ (*<*)oops(*>*) subsection \Congruence Rules and Evaluation Order\ text \ Higher order logic differs from functional programming languages in that it has no built-in notion of evaluation order. A program is just a set of equations, and it is not specified how they must be evaluated. However for the purpose of function definition, we must talk about evaluation order implicitly, when we reason about termination. Congruence rules express that a certain evaluation order is consistent with the logical definition. Consider the following function. \ function f :: "nat \ bool" where "f n = (n = 0 \ f (n - 1))" (*<*)by pat_completeness auto(*>*) text \ For this definition, the termination proof fails. The default configuration specifies no congruence rule for disjunction. We have to add a congruence rule that specifies left-to-right evaluation order: \vspace{1ex} \noindent @{thm disj_cong}\hfill(@{thm [source] "disj_cong"}) \vspace{1ex} Now the definition works without problems. Note how the termination proof depends on the extra condition that we get from the congruence rule. However, as evaluation is not a hard-wired concept, we could just turn everything around by declaring a different congruence rule. Then we can make the reverse definition: \ lemma disj_cong2[fundef_cong]: "(\ Q' \ P = P') \ (Q = Q') \ (P \ Q) = (P' \ Q')" by blast fun f' :: "nat \ bool" where "f' n = (f' (n - 1) \ n = 0)" text \ \noindent These examples show that, in general, there is no \qt{best} set of congruence rules. However, such tweaking should rarely be necessary in practice, as most of the time, the default set of congruence rules works well. \ end diff --git a/src/Doc/Implementation/Eq.thy b/src/Doc/Implementation/Eq.thy --- a/src/Doc/Implementation/Eq.thy +++ b/src/Doc/Implementation/Eq.thy @@ -1,139 +1,139 @@ (*:maxLineLen=78:*) theory Eq imports Base begin chapter \Equational reasoning\ text \ Equality is one of the most fundamental concepts of mathematics. The Isabelle/Pure logic (\chref{ch:logic}) provides a builtin relation \\ :: \ \ \ \ prop\ that expresses equality of arbitrary terms (or propositions) at the framework level, as expressed by certain basic inference rules (\secref{sec:eq-rules}). Equational reasoning means to replace equals by equals, using reflexivity and transitivity to form chains of replacement steps, and congruence rules to access sub-structures. Conversions (\secref{sec:conv}) provide a convenient framework to compose basic equational steps to build specific equational reasoning tools. Higher-order matching is able to provide suitable instantiations for giving equality rules, which leads to the versatile concept of \\\-term rewriting (\secref{sec:rewriting}). Internally this is based on the general-purpose Simplifier engine of Isabelle, which is more specific and more efficient than plain conversions. Object-logics usually introduce specific notions of equality or equivalence, and relate it with the Pure equality. This enables to re-use the Pure tools for equational reasoning for particular object-logic connectives as well. \ section \Basic equality rules \label{sec:eq-rules}\ text \ Isabelle/Pure uses \\\ for equality of arbitrary terms, which includes equivalence of propositions of the logical framework. The conceptual axiomatization of the constant \\ :: \ \ \ \ prop\ is given in \figref{fig:pure-equality}. The inference kernel presents slightly different equality rules, which may be understood as derived rules from this minimal axiomatization. The Pure theory also provides some theorems that express the same reasoning schemes as theorems that can be composed like object-level rules as explained in \secref{sec:obj-rules}. For example, \<^ML>\Thm.symmetric\ as Pure inference is an ML function that maps a theorem \th\ stating \t \ u\ to one stating \u \ t\. In contrast, @{thm [source] Pure.symmetric} as Pure theorem expresses the same reasoning in declarative form. If used like \th [THEN Pure.symmetric]\ in Isar source notation, it achieves a similar effect as the ML inference function, although the rule attribute @{attribute THEN} or ML operator \<^ML>\op RS\ involve the full machinery of higher-order unification (modulo \\\\-conversion) and lifting of \\/\\ contexts. \ text %mlref \ \begin{mldecls} @{define_ML Thm.reflexive: "cterm -> thm"} \\ @{define_ML Thm.symmetric: "thm -> thm"} \\ @{define_ML Thm.transitive: "thm -> thm -> thm"} \\ @{define_ML Thm.abstract_rule: "string -> cterm -> thm -> thm"} \\ @{define_ML Thm.combination: "thm -> thm -> thm"} \\[0.5ex] @{define_ML Thm.equal_intr: "thm -> thm -> thm"} \\ @{define_ML Thm.equal_elim: "thm -> thm -> thm"} \\ \end{mldecls} See also \<^file>\~~/src/Pure/thm.ML\ for further description of these inference rules, and a few more for primitive \\\ and \\\ conversions. Note that \\\ conversion is implicit due to the representation of terms with de-Bruijn indices (\secref{sec:terms}). \ section \Conversions \label{sec:conv}\ text \ - The classic article @{cite "paulson:1983"} introduces the concept of + The classic article \<^cite>\"paulson:1983"\ introduces the concept of conversion for Cambridge LCF. This was historically important to implement all kinds of ``simplifiers'', but the Isabelle Simplifier is done quite - differently, see @{cite \\S9.3\ "isabelle-isar-ref"}. + differently, see \<^cite>\\\S9.3\ in "isabelle-isar-ref"\. \ text %mlref \ \begin{mldecls} @{define_ML_structure Conv} \\ @{define_ML_type conv} \\ @{define_ML Simplifier.asm_full_rewrite : "Proof.context -> conv"} \\ \end{mldecls} \<^descr> \<^ML_structure>\Conv\ is a library of combinators to build conversions, over type \<^ML_type>\conv\ (see also \<^file>\~~/src/Pure/conv.ML\). This is one of the few Isabelle/ML modules that are usually used with \<^verbatim>\open\: finding examples works by searching for ``\<^verbatim>\open Conv\'' instead of ``\<^verbatim>\Conv.\''. \<^descr> \<^ML>\Simplifier.asm_full_rewrite\ invokes the Simplifier as a conversion. There are a few related operations, corresponding to the various modes of simplification. \ section \Rewriting \label{sec:rewriting}\ text \ Rewriting normalizes a given term (theorem or goal) by replacing instances of given equalities \t \ u\ in subterms. Rewriting continues until no rewrites are applicable to any subterm. This may be used to unfold simple definitions of the form \f x\<^sub>1 \ x\<^sub>n \ u\, but is slightly more general than that. \ text %mlref \ \begin{mldecls} @{define_ML rewrite_rule: "Proof.context -> thm list -> thm -> thm"} \\ @{define_ML rewrite_goals_rule: "Proof.context -> thm list -> thm -> thm"} \\ @{define_ML rewrite_goal_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML rewrite_goals_tac: "Proof.context -> thm list -> tactic"} \\ @{define_ML fold_goals_tac: "Proof.context -> thm list -> tactic"} \\ \end{mldecls} \<^descr> \<^ML>\rewrite_rule\~\ctxt rules thm\ rewrites the whole theorem by the given rules. \<^descr> \<^ML>\rewrite_goals_rule\~\ctxt rules thm\ rewrites the outer premises of the given theorem. Interpreting the same as a goal state (\secref{sec:tactical-goals}) it means to rewrite all subgoals (in the same manner as \<^ML>\rewrite_goals_tac\). \<^descr> \<^ML>\rewrite_goal_tac\~\ctxt rules i\ rewrites subgoal \i\ by the given rewrite rules. \<^descr> \<^ML>\rewrite_goals_tac\~\ctxt rules\ rewrites all subgoals by the given rewrite rules. \<^descr> \<^ML>\fold_goals_tac\~\ctxt rules\ essentially uses \<^ML>\rewrite_goals_tac\ with the symmetric form of each member of \rules\, re-ordered to fold longer expression first. This supports to idea to fold primitive definitions that appear in expended form in the proof state. \ end diff --git a/src/Doc/Implementation/Integration.thy b/src/Doc/Implementation/Integration.thy --- a/src/Doc/Implementation/Integration.thy +++ b/src/Doc/Implementation/Integration.thy @@ -1,174 +1,174 @@ (*:maxLineLen=78:*) theory Integration imports Base begin chapter \System integration\ section \Isar toplevel \label{sec:isar-toplevel}\ text \ The Isar \<^emph>\toplevel state\ represents the outermost configuration that is transformed by a sequence of transitions (commands) within a theory body. This is a pure value with pure functions acting on it in a timeless and stateless manner. Historically, the sequence of transitions was wrapped up as sequential command loop, such that commands are applied one-by-one. In contemporary Isabelle/Isar, processing toplevel commands usually works in - parallel in multi-threaded Isabelle/ML @{cite "Wenzel:2009" and - "Wenzel:2013:ITP"}. + parallel in multi-threaded Isabelle/ML \<^cite>\"Wenzel:2009" and + "Wenzel:2013:ITP"\. \ subsection \Toplevel state\ text \ The toplevel state is a disjoint sum of empty \toplevel\, or \theory\, or \proof\. The initial toplevel is empty; a theory is commenced by a @{command theory} header; within a theory we may use theory commands such as @{command definition}, or state a @{command theorem} to be proven. A proof state accepts a rich collection of Isar proof commands for structured proof composition, or unstructured proof scripts. When the proof is concluded we get back to the (local) theory, which is then updated by defining the resulting fact. Further theory declarations or theorem statements with proofs may follow, until we eventually conclude the theory development by issuing @{command end} to get back to the empty toplevel. \ text %mlref \ \begin{mldecls} @{define_ML_type Toplevel.state} \\ @{define_ML_exception Toplevel.UNDEF} \\ @{define_ML Toplevel.is_toplevel: "Toplevel.state -> bool"} \\ @{define_ML Toplevel.theory_of: "Toplevel.state -> theory"} \\ @{define_ML Toplevel.proof_of: "Toplevel.state -> Proof.state"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Toplevel.state\ represents Isar toplevel states, which are normally manipulated through the concept of toplevel transitions only (\secref{sec:toplevel-transition}). \<^descr> \<^ML>\Toplevel.UNDEF\ is raised for undefined toplevel operations. Many operations work only partially for certain cases, since \<^ML_type>\Toplevel.state\ is a sum type. \<^descr> \<^ML>\Toplevel.is_toplevel\~\state\ checks for an empty toplevel state. \<^descr> \<^ML>\Toplevel.theory_of\~\state\ selects the background theory of \state\, it raises \<^ML>\Toplevel.UNDEF\ for an empty toplevel state. \<^descr> \<^ML>\Toplevel.proof_of\~\state\ selects the Isar proof state if available, otherwise it raises an error. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "Isar.state"} & : & \ML_antiquotation\ \\ \end{matharray} \<^descr> \@{Isar.state}\ refers to Isar toplevel state at that point --- as abstract value. This only works for diagnostic ML commands, such as @{command ML_val} or @{command ML_command}. \ subsection \Toplevel transitions \label{sec:toplevel-transition}\ text \ An Isar toplevel transition consists of a partial function on the toplevel state, with additional information for diagnostics and error reporting: there are fields for command name, source position, and other meta-data. The operational part is represented as the sequential union of a list of partial functions, which are tried in turn until the first one succeeds. This acts like an outer case-expression for various alternative state transitions. For example, \<^theory_text>\qed\ works differently for a local proofs vs.\ the global ending of an outermost proof. Transitions are composed via transition transformers. Internally, Isar commands are put together from an empty transition extended by name and source position. It is then left to the individual command parser to turn the given concrete syntax into a suitable transition transformer that adjoins actual operations on a theory or proof state. \ text %mlref \ \begin{mldecls} @{define_ML Toplevel.keep: "(Toplevel.state -> unit) -> Toplevel.transition -> Toplevel.transition"} \\ @{define_ML Toplevel.theory: "(theory -> theory) -> Toplevel.transition -> Toplevel.transition"} \\ @{define_ML Toplevel.theory_to_proof: "(theory -> Proof.state) -> Toplevel.transition -> Toplevel.transition"} \\ @{define_ML Toplevel.proof: "(Proof.state -> Proof.state) -> Toplevel.transition -> Toplevel.transition"} \\ @{define_ML Toplevel.proofs: "(Proof.state -> Proof.state Seq.result Seq.seq) -> Toplevel.transition -> Toplevel.transition"} \\ @{define_ML Toplevel.end_proof: "(bool -> Proof.state -> Proof.context) -> Toplevel.transition -> Toplevel.transition"} \\ \end{mldecls} \<^descr> \<^ML>\Toplevel.keep\~\tr\ adjoins a diagnostic function. \<^descr> \<^ML>\Toplevel.theory\~\tr\ adjoins a theory transformer. \<^descr> \<^ML>\Toplevel.theory_to_proof\~\tr\ adjoins a global goal function, which turns a theory into a proof state. The theory may be changed before entering the proof; the generic Isar goal setup includes an \<^verbatim>\after_qed\ argument that specifies how to apply the proven result to the enclosing context, when the proof is finished. \<^descr> \<^ML>\Toplevel.proof\~\tr\ adjoins a deterministic proof command, with a singleton result. \<^descr> \<^ML>\Toplevel.proofs\~\tr\ adjoins a general proof command, with zero or more result states (represented as a lazy list). \<^descr> \<^ML>\Toplevel.end_proof\~\tr\ adjoins a concluding proof command, that returns the resulting theory, after applying the resulting facts to the target context. \ text %mlex \ The file \<^file>\~~/src/HOL/Examples/Commands.thy\ shows some example Isar command definitions, with the all-important theory header declarations for outer syntax keywords. \ section \Theory loader database\ text \ In batch mode and within dumped logic images, the theory database maintains a collection of theories as a directed acyclic graph. A theory may refer to other theories as @{keyword "imports"}, or to auxiliary files via special \<^emph>\load commands\ (e.g.\ @{command ML_file}). For each theory, the base directory of its own theory file is called \<^emph>\master directory\: this is used as the relative location to refer to other files from that theory. \ text %mlref \ \begin{mldecls} @{define_ML use_thy: "string -> unit"} \\ @{define_ML Thy_Info.get_theory: "string -> theory"} \\ @{define_ML Thy_Info.remove_thy: "string -> unit"} \\ @{define_ML Thy_Info.register_thy: "theory -> unit"} \\ \end{mldecls} \<^descr> \<^ML>\use_thy\~\A\ ensures that theory \A\ is fully up-to-date wrt.\ the external file store; outdated ancestors are reloaded on demand. \<^descr> \<^ML>\Thy_Info.get_theory\~\A\ retrieves the theory value presently associated with name \A\. Note that the result might be outdated wrt.\ the file-system content. \<^descr> \<^ML>\Thy_Info.remove_thy\~\A\ deletes theory \A\ and all descendants from the theory database. \<^descr> \<^ML>\Thy_Info.register_thy\~\text thy\ registers an existing theory value with the theory loader database and updates source version information according to the file store. \ end diff --git a/src/Doc/Implementation/Isar.thy b/src/Doc/Implementation/Isar.thy --- a/src/Doc/Implementation/Isar.thy +++ b/src/Doc/Implementation/Isar.thy @@ -1,547 +1,547 @@ (*:maxLineLen=78:*) theory Isar imports Base begin chapter \Isar language elements\ text \ - The Isar proof language (see also @{cite \\S2\ "isabelle-isar-ref"}) + The Isar proof language (see also \<^cite>\\\S2\ in "isabelle-isar-ref"\) consists of three main categories of language elements: \<^enum> Proof \<^emph>\commands\ define the primary language of transactions of the underlying Isar/VM interpreter. Typical examples are @{command "fix"}, @{command "assume"}, @{command "show"}, @{command "proof"}, and @{command "qed"}. Composing proof commands according to the rules of the Isar/VM leads to expressions of structured proof text, such that both the machine and the human reader can give it a meaning as formal reasoning. \<^enum> Proof \<^emph>\methods\ define a secondary language of mixed forward-backward refinement steps involving facts and goals. Typical examples are @{method rule}, @{method unfold}, and @{method simp}. Methods can occur in certain well-defined parts of the Isar proof language, say as arguments to @{command "proof"}, @{command "qed"}, or @{command "by"}. \<^enum> \<^emph>\Attributes\ define a tertiary language of small annotations to theorems being defined or referenced. Attributes can modify both the context and the theorem. Typical examples are @{attribute intro} (which affects the context), and @{attribute symmetric} (which affects the theorem). \ section \Proof commands\ text \ A \<^emph>\proof command\ is state transition of the Isar/VM proof interpreter. In principle, Isar proof commands could be defined in user-space as well. The system is built like that in the first place: one part of the commands are primitive, the other part is defined as derived elements. Adding to the genuine structured proof language requires profound understanding of the Isar/VM machinery, though, so this is beyond the scope of this manual. What can be done realistically is to define some diagnostic commands that inspect the general state of the Isar/VM, and report some feedback to the user. Typically this involves checking of the linguistic \<^emph>\mode\ of a proof state, or peeking at the pending goals (if available). Another common application is to define a toplevel command that poses a problem to the user as Isar proof state and processes the final result relatively to the context. Thus a proof can be incorporated into the context of some user-space tool, without modifying the Isar proof language itself. \ text %mlref \ \begin{mldecls} @{define_ML_type Proof.state} \\ @{define_ML Proof.assert_forward: "Proof.state -> Proof.state"} \\ @{define_ML Proof.assert_chain: "Proof.state -> Proof.state"} \\ @{define_ML Proof.assert_backward: "Proof.state -> Proof.state"} \\ @{define_ML Proof.simple_goal: "Proof.state -> {context: Proof.context, goal: thm}"} \\ @{define_ML Proof.goal: "Proof.state -> {context: Proof.context, facts: thm list, goal: thm}"} \\ @{define_ML Proof.raw_goal: "Proof.state -> {context: Proof.context, facts: thm list, goal: thm}"} \\ @{define_ML Proof.theorem: "Method.text option -> (thm list list -> Proof.context -> Proof.context) -> (term * term list) list list -> Proof.context -> Proof.state"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Proof.state\ represents Isar proof states. This is a block-structured configuration with proof context, linguistic mode, and optional goal. The latter consists of goal context, goal facts (``\using\''), and tactical goal state (see \secref{sec:tactical-goals}). The general idea is that the facts shall contribute to the refinement of some parts of the tactical goal --- how exactly is defined by the proof method that is applied in that situation. \<^descr> \<^ML>\Proof.assert_forward\, \<^ML>\Proof.assert_chain\, \<^ML>\Proof.assert_backward\ are partial identity functions that fail unless a certain linguistic mode is active, namely ``\proof(state)\'', ``\proof(chain)\'', ``\proof(prove)\'', respectively (using the terminology - of @{cite "isabelle-isar-ref"}). + of \<^cite>\"isabelle-isar-ref"\). It is advisable study the implementations of existing proof commands for suitable modes to be asserted. \<^descr> \<^ML>\Proof.simple_goal\~\state\ returns the structured Isar goal (if available) in the form seen by ``simple'' methods (like @{method simp} or @{method blast}). The Isar goal facts are already inserted as premises into the subgoals, which are presented individually as in \<^ML>\Proof.goal\. \<^descr> \<^ML>\Proof.goal\~\state\ returns the structured Isar goal (if available) in the form seen by regular methods (like @{method rule}). The auxiliary internal encoding of Pure conjunctions is split into individual subgoals as usual. \<^descr> \<^ML>\Proof.raw_goal\~\state\ returns the structured Isar goal (if available) in the raw internal form seen by ``raw'' methods (like @{method induct}). This form is rarely appropriate for diagnostic tools; \<^ML>\Proof.simple_goal\ or \<^ML>\Proof.goal\ should be used in most situations. \<^descr> \<^ML>\Proof.theorem\~\before_qed after_qed statement ctxt\ initializes a toplevel Isar proof state within a given context. The optional \before_qed\ method is applied at the end of the proof, just before extracting the result (this feature is rarely used). The \after_qed\ continuation receives the extracted result in order to apply it to the final context in a suitable way (e.g.\ storing named facts). Note that at this generic level the target context is specified as \<^ML_type>\Proof.context\, but the usual wrapping of toplevel proofs into command transactions will provide a \<^ML_type>\local_theory\ here (\chref{ch:local-theory}). This affects the way how results are stored. The \statement\ is given as a nested list of terms, each associated with optional @{keyword "is"} patterns as usual in the Isar source language. The original nested list structure over terms is turned into one over theorems when \after_qed\ is invoked. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "Isar.goal"} & : & \ML_antiquotation\ \\ \end{matharray} \<^descr> \@{Isar.goal}\ refers to the regular goal state (if available) of the current proof state managed by the Isar toplevel --- as abstract value. This only works for diagnostic ML commands, such as @{command ML_val} or @{command ML_command}. \ text %mlex \ The following example peeks at a certain goal configuration. \ notepad begin have A and B and C ML_val \val n = Thm.nprems_of (#goal @{Isar.goal}); \<^assert> (n = 3);\ sorry end text \ Here we see 3 individual subgoals in the same way as regular proof methods would do. \ section \Proof methods\ text \ A \method\ is a function \thm\<^sup>* \ context * goal \ (context \ goal)\<^sup>*\<^sup>*\ that operates on the full Isar goal configuration with context, goal facts, and tactical goal state and enumerates possible follow-up goal states. Under normal circumstances, the goal context remains unchanged, but it is also possible to declare named extensions of the proof context (\<^emph>\cases\). This means a proof method is like a structurally enhanced tactic (cf.\ \secref{sec:tactics}). The well-formedness conditions for tactics need to hold for methods accordingly, with the following additions. \<^item> Goal addressing is further limited either to operate uniformly on \<^emph>\all\ subgoals, or specifically on the \<^emph>\first\ subgoal. Exception: old-style tactic emulations that are embedded into the method space, e.g.\ @{method rule_tac}. \<^item> A non-trivial method always needs to make progress: an identical follow-up goal state has to be avoided.\<^footnote>\This enables the user to write method expressions like \meth\<^sup>+\ without looping, while the trivial do-nothing case can be recovered via \meth\<^sup>?\.\ Exception: trivial stuttering steps, such as ``@{method -}'' or @{method succeed}. \<^item> Goal facts passed to the method must not be ignored. If there is no sensible use of facts outside the goal state, facts should be inserted into the subgoals that are addressed by the method. \<^medskip> Syntactically, the language of proof methods appears as arguments to Isar commands like @{command "by"} or @{command apply}. User-space additions are reasonably easy by plugging suitable method-valued parser functions into the framework, using the @{command method_setup} command, for example. To get a better idea about the range of possibilities, consider the following Isar proof schemes. This is the general form of structured proof text: \<^medskip> \begin{tabular}{l} @{command from}~\facts\<^sub>1\~@{command have}~\props\~@{command using}~\facts\<^sub>2\ \\ @{command proof}~\(initial_method)\ \\ \quad\body\ \\ @{command qed}~\(terminal_method)\ \\ \end{tabular} \<^medskip> The goal configuration consists of \facts\<^sub>1\ and \facts\<^sub>2\ appended in that order, and various \props\ being claimed. The \initial_method\ is invoked with facts and goals together and refines the problem to something that is handled recursively in the proof \body\. The \terminal_method\ has another chance to finish any remaining subgoals, but it does not see the facts of the initial step. \<^medskip> This pattern illustrates unstructured proof scripts: \<^medskip> \begin{tabular}{l} @{command have}~\props\ \\ \quad@{command using}~\facts\<^sub>1\~@{command apply}~\method\<^sub>1\ \\ \quad@{command apply}~\method\<^sub>2\ \\ \quad@{command using}~\facts\<^sub>3\~@{command apply}~\method\<^sub>3\ \\ \quad@{command done} \\ \end{tabular} \<^medskip> The \method\<^sub>1\ operates on the original claim while using \facts\<^sub>1\. Since the @{command apply} command structurally resets the facts, the \method\<^sub>2\ will operate on the remaining goal state without facts. The \method\<^sub>3\ will see again a collection of \facts\<^sub>3\ that has been inserted into the script explicitly. \<^medskip> Empirically, any Isar proof method can be categorized as follows. \<^enum> \<^emph>\Special method with cases\ with named context additions associated with the follow-up goal state. Example: @{method "induct"}, which is also a ``raw'' method since it operates on the internal representation of simultaneous claims as Pure conjunction (\secref{sec:logic-aux}), instead of separate subgoals (\secref{sec:tactical-goals}). \<^enum> \<^emph>\Structured method\ with strong emphasis on facts outside the goal state. Example: @{method "rule"}, which captures the key ideas behind structured reasoning in Isar in its purest form. \<^enum> \<^emph>\Simple method\ with weaker emphasis on facts, which are inserted into subgoals to emulate old-style tactical ``premises''. Examples: @{method "simp"}, @{method "blast"}, @{method "auto"}. \<^enum> \<^emph>\Old-style tactic emulation\ with detailed numeric goal addressing and explicit references to entities of the internal goal state (which are otherwise invisible from proper Isar proof text). The naming convention \foo_tac\ makes this special non-standard status clear. Example: @{method "rule_tac"}. When implementing proof methods, it is advisable to study existing implementations carefully and imitate the typical ``boiler plate'' for context-sensitive parsing and further combinators to wrap-up tactic expressions as methods.\<^footnote>\Aliases or abbreviations of the standard method combinators should be avoided. Note that from Isabelle99 until Isabelle2009 the system did provide various odd combinations of method syntax wrappers that made applications more complicated than necessary.\ \ text %mlref \ \begin{mldecls} @{define_ML_type Proof.method} \\ @{define_ML CONTEXT_METHOD: "(thm list -> context_tactic) -> Proof.method"} \\ @{define_ML METHOD: "(thm list -> tactic) -> Proof.method"} \\ @{define_ML SIMPLE_METHOD: "tactic -> Proof.method"} \\ @{define_ML SIMPLE_METHOD': "(int -> tactic) -> Proof.method"} \\ @{define_ML Method.insert_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML Method.setup: "binding -> (Proof.context -> Proof.method) context_parser -> string -> theory -> theory"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Proof.method\ represents proof methods as abstract type. \<^descr> \<^ML>\CONTEXT_METHOD\~\(fn facts => context_tactic)\ wraps \context_tactic\ depending on goal facts as a general proof method that may change the proof context dynamically. A typical operation is \<^ML>\Proof_Context.update_cases\, which is wrapped up as combinator @{define_ML CONTEXT_CASES} for convenience. \<^descr> \<^ML>\METHOD\~\(fn facts => tactic)\ wraps \tactic\ depending on goal facts as regular proof method; the goal context is passed via method syntax. \<^descr> \<^ML>\SIMPLE_METHOD\~\tactic\ wraps a tactic that addresses all subgoals uniformly as simple proof method. Goal facts are already inserted into all subgoals before \tactic\ is applied. \<^descr> \<^ML>\SIMPLE_METHOD'\~\tactic\ wraps a tactic that addresses a specific subgoal as simple proof method that operates on subgoal 1. Goal facts are inserted into the subgoal then the \tactic\ is applied. \<^descr> \<^ML>\Method.insert_tac\~\ctxt facts i\ inserts \facts\ into subgoal \i\. This is convenient to reproduce part of the \<^ML>\SIMPLE_METHOD\ or \<^ML>\SIMPLE_METHOD'\ wrapping within regular \<^ML>\METHOD\, for example. \<^descr> \<^ML>\Method.setup\~\name parser description\ provides the functionality of the Isar command @{command method_setup} as ML function. \ text %mlex \ - See also @{command method_setup} in @{cite "isabelle-isar-ref"} which + See also @{command method_setup} in \<^cite>\"isabelle-isar-ref"\ which includes some abstract examples. \<^medskip> The following toy examples illustrate how the goal facts and state are passed to proof methods. The predefined proof method called ``@{method tactic}'' wraps ML source of type \<^ML_type>\tactic\ (abstracted over \<^ML_text>\facts\). This allows immediate experimentation without parsing of concrete syntax. \ notepad begin fix A B :: bool assume a: A and b: B have "A \ B" apply (tactic \resolve_tac \<^context> @{thms conjI} 1\) using a apply (tactic \resolve_tac \<^context> facts 1\) using b apply (tactic \resolve_tac \<^context> facts 1\) done have "A \ B" using a and b ML_val \@{Isar.goal}\ apply (tactic \Method.insert_tac \<^context> facts 1\) apply (tactic \(resolve_tac \<^context> @{thms conjI} THEN_ALL_NEW assume_tac \<^context>) 1\) done end text \ \<^medskip> The next example implements a method that simplifies the first subgoal by rewrite rules that are given as arguments. \ method_setup my_simp = \Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (fn i => CHANGED (asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps thms) i)))\ "rewrite subgoal by given rules" text \ The concrete syntax wrapping of @{command method_setup} always passes-through the proof context at the end of parsing, but it is not used in this example. The \<^ML>\Attrib.thms\ parser produces a list of theorems from the usual Isar syntax involving attribute expressions etc.\ (syntax category @{syntax - thms}) @{cite "isabelle-isar-ref"}. The resulting \<^ML_text>\thms\ are + thms}) \<^cite>\"isabelle-isar-ref"\. The resulting \<^ML_text>\thms\ are added to \<^ML>\HOL_basic_ss\ which already contains the basic Simplifier setup for HOL. The tactic \<^ML>\asm_full_simp_tac\ is the one that is also used in method @{method simp} by default. The extra wrapping by the \<^ML>\CHANGED\ tactical ensures progress of simplification: identical goal states are filtered out explicitly to make the raw tactic conform to standard Isar method behaviour. \<^medskip> Method @{method my_simp} can be used in Isar proofs like this: \ notepad begin fix a b c :: 'a assume a: "a = b" assume b: "b = c" have "a = c" by (my_simp a b) end text \ Here is a similar method that operates on all subgoals, instead of just the first one.\ method_setup my_simp_all = \Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD (CHANGED (ALLGOALS (asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps thms)))))\ "rewrite all subgoals by given rules" notepad begin fix a b c :: 'a assume a: "a = b" assume b: "b = c" have "a = c" and "c = b" by (my_simp_all a b) end text \ \<^medskip> Apart from explicit arguments, common proof methods typically work with a default configuration provided by the context. As a shortcut to rule management we use a cheap solution via the @{command named_theorems} command to declare a dynamic fact in the context. \ named_theorems my_simp text \ The proof method is now defined as before, but we append the explicit arguments and the rules from the context. \ method_setup my_simp' = \Attrib.thms >> (fn thms => fn ctxt => let val my_simps = Named_Theorems.get ctxt \<^named_theorems>\my_simp\ in SIMPLE_METHOD' (fn i => CHANGED (asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (thms @ my_simps)) i)) end)\ "rewrite subgoal by given rules and my_simp rules from the context" text \ \<^medskip> Method @{method my_simp'} can be used in Isar proofs like this: \ notepad begin fix a b c :: 'a assume [my_simp]: "a \ b" assume [my_simp]: "b \ c" have "a \ c" by my_simp' end text \ \<^medskip> The @{method my_simp} variants defined above are ``simple'' methods, i.e.\ the goal facts are merely inserted as goal premises by the \<^ML>\SIMPLE_METHOD'\ or \<^ML>\SIMPLE_METHOD\ wrapper. For proof methods that are similar to the standard collection of @{method simp}, @{method blast}, @{method fast}, @{method auto} there is little more that can be done. Note that using the primary goal facts in the same manner as the method arguments obtained via concrete syntax or the context does not meet the requirement of ``strong emphasis on facts'' of regular proof methods, because rewrite rules as used above can be easily ignored. A proof text ``@{command using}~\foo\~@{command "by"}~\my_simp\'' where \foo\ is not used would deceive the reader. \<^medskip> The technical treatment of rules from the context requires further attention. Above we rebuild a fresh \<^ML_type>\simpset\ from the arguments and \<^emph>\all\ rules retrieved from the context on every invocation of the method. This does not scale to really large collections of rules, which easily emerges in the context of a big theory library, for example. This is an inherent limitation of the simplistic rule management via @{command named_theorems}, because it lacks tool-specific storage and retrieval. More realistic applications require efficient index-structures that organize theorems in a customized manner, such as a discrimination net that is indexed by the left-hand sides of rewrite rules. For variations on the Simplifier, re-use of the existing type \<^ML_type>\simpset\ is adequate, but scalability would require it be maintained statically within the context data, not dynamically on each tool invocation. \ section \Attributes \label{sec:attributes}\ text \ An \<^emph>\attribute\ is a function \context \ thm \ context \ thm\, which means both a (generic) context and a theorem can be modified simultaneously. In practice this mixed form is very rare, instead attributes are presented either as \<^emph>\declaration attribute:\ \thm \ context \ context\ or \<^emph>\rule attribute:\ \context \ thm \ thm\. Attributes can have additional arguments via concrete syntax. There is a collection of context-sensitive parsers for various logical entities (types, terms, theorems). These already take care of applying morphisms to the arguments when attribute expressions are moved into a different context (see also \secref{sec:morphisms}). When implementing declaration attributes, it is important to operate exactly on the variant of the generic context that is provided by the system, which is either global theory context or local proof context. In particular, the background theory of a local context must not be modified in this situation! \ text %mlref \ \begin{mldecls} @{define_ML_type attribute} \\ @{define_ML Thm.rule_attribute: "thm list -> (Context.generic -> thm -> thm) -> attribute"} \\ @{define_ML Thm.declaration_attribute: " (thm -> Context.generic -> Context.generic) -> attribute"} \\ @{define_ML Attrib.setup: "binding -> attribute context_parser -> string -> theory -> theory"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\attribute\ represents attributes as concrete type alias. \<^descr> \<^ML>\Thm.rule_attribute\~\thms (fn context => rule)\ wraps a context-dependent rule (mapping on \<^ML_type>\thm\) as attribute. The \thms\ are additional parameters: when forming an abstract closure, the system may provide dummy facts that are propagated according to strict evaluation discipline. In that case, \rule\ is bypassed. \<^descr> \<^ML>\Thm.declaration_attribute\~\(fn thm => decl)\ wraps a theorem-dependent declaration (mapping on \<^ML_type>\Context.generic\) as attribute. When forming an abstract closure, the system may provide a dummy fact as \thm\. In that case, \decl\ is bypassed. \<^descr> \<^ML>\Attrib.setup\~\name parser description\ provides the functionality of the Isar command @{command attribute_setup} as ML function. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def attributes} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation attributes} attributes \ \<^descr> \@{attributes [\]}\ embeds attribute source representation into the ML text, which is particularly useful with declarations like \<^ML>\Local_Theory.note\. Attribute names are internalized at compile time, but the source is unevaluated. This means attributes with formal arguments (types, terms, theorems) may be subject to odd effects of dynamic scoping! \ text %mlex \ - See also @{command attribute_setup} in @{cite "isabelle-isar-ref"} which + See also @{command attribute_setup} in \<^cite>\"isabelle-isar-ref"\ which includes some abstract examples. \ end diff --git a/src/Doc/Implementation/Local_Theory.thy b/src/Doc/Implementation/Local_Theory.thy --- a/src/Doc/Implementation/Local_Theory.thy +++ b/src/Doc/Implementation/Local_Theory.thy @@ -1,149 +1,149 @@ (*:maxLineLen=78:*) theory Local_Theory imports Base begin chapter \Local theory specifications \label{ch:local-theory}\ text \ A \<^emph>\local theory\ combines aspects of both theory and proof context (cf.\ \secref{sec:context}), such that definitional specifications may be given relatively to parameters and assumptions. A local theory is represented as a regular proof context, augmented by administrative data about the \<^emph>\target context\. The target is usually derived from the background theory by adding local \\\ and \\\ elements, plus suitable modifications of non-logical context data (e.g.\ a special type-checking discipline). Once initialized, the target is ready to absorb definitional primitives: \\\ for terms and \\\ for theorems. Such definitions may get transformed in a target-specific way, but the programming interface hides such details. Isabelle/Pure provides target mechanisms for locales, type-classes, type-class instantiations, and general overloading. In principle, users can implement new targets as well, but this rather arcane discipline is beyond the scope of this manual. In contrast, implementing derived definitional packages to be used within a local theory context is quite easy: the interfaces are even simpler and more abstract than the underlying primitives for raw theories. Many definitional packages for local theories are available in Isabelle. Although a few old packages only work for global theories, the standard way of implementing definitional packages in Isabelle is via the local theory interface. \ section \Definitional elements\ text \ There are separate elements \\ c \ t\ for terms, and \\ b = thm\ for theorems. Types are treated implicitly, according to Hindley-Milner discipline (cf.\ \secref{sec:variables}). These definitional primitives essentially act like \let\-bindings within a local context that may already contain earlier \let\-bindings and some initial \\\-bindings. Thus we gain \<^emph>\dependent definitions\ that are relative to an initial axiomatic context. The following diagram illustrates this idea of axiomatic elements versus definitional elements: \begin{center} \begin{tabular}{|l|l|l|} \hline & \\\-binding & \let\-binding \\ \hline types & fixed \\\ & arbitrary \\\ \\ terms & \\ x :: \\ & \\ c \ t\ \\ theorems & \\ a: A\ & \\ b = \<^BG>B\<^EN>\ \\ \hline \end{tabular} \end{center} A user package merely needs to produce suitable \\\ and \\\ elements according to the application. For example, a package for inductive definitions might first \\\ a certain predicate as some fixed-point construction, then \\\ a proven result about monotonicity of the functor involved here, and then produce further derived concepts via additional \\\ and \\\ elements. The cumulative sequence of \\\ and \\\ produced at package runtime is managed by the local theory infrastructure by means of an \<^emph>\auxiliary context\. Thus the system holds up the impression of working within a fully abstract situation with hypothetical entities: \\ c \ t\ always results in a literal fact \\<^BG>c \ t\<^EN>\, where \c\ is a fixed variable \c\. The details about global constants, name spaces etc. are handled internally. So the general structure of a local theory is a sandwich of three layers: \begin{center} \framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}} \end{center} When a definitional package is finished, the auxiliary context is reset to the target context. The target now holds definitions for terms and theorems that stem from the hypothetical \\\ and \\\ elements, - transformed by the particular target policy (see @{cite \\S4--5\ - "Haftmann-Wenzel:2009"} for details). + transformed by the particular target policy (see \<^cite>\\\S4--5\ in + "Haftmann-Wenzel:2009"\ for details). \ text %mlref \ \begin{mldecls} @{define_ML_type local_theory = Proof.context} \\ @{define_ML Named_Target.init: "string list -> string -> theory -> local_theory"} \\[1ex] @{define_ML Local_Theory.define: "(binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory"} \\ @{define_ML Local_Theory.note: "Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\local_theory\ represents local theories. Although this is merely an alias for \<^ML_type>\Proof.context\, it is semantically a subtype of the same: a \<^ML_type>\local_theory\ holds target information as special context data. Subtyping means that any value \lthy:\~\<^ML_type>\local_theory\ can be also used with operations on expecting a regular \ctxt:\~\<^ML_type>\Proof.context\. \<^descr> \<^ML>\Named_Target.init\~\includes name thy\ initializes a local theory derived from the given background theory. An empty name refers to a \<^emph>\global theory\ context, and a non-empty name refers to a @{command locale} or @{command class} context (a fully-qualified internal name is expected here). This is useful for experimentation --- normally the Isar toplevel already takes care to initialize the local theory context. \<^descr> \<^ML>\Local_Theory.define\~\((b, mx), (a, rhs)) lthy\ defines a local entity according to the specification that is given relatively to the current \lthy\ context. In particular the term of the RHS may refer to earlier local entities from the auxiliary context, or hypothetical parameters from the target context. The result is the newly defined term (which is always a fixed variable with exactly the same name as specified for the LHS), together with an equational theorem that states the definition as a hypothetical fact. Unless an explicit name binding is given for the RHS, the resulting fact will be called \b_def\. Any given attributes are applied to that same fact --- immediately in the auxiliary context \<^emph>\and\ in any transformed versions stemming from target-specific policies or any later interpretations of results from the target context (think of @{command locale} and @{command interpretation}, for example). This means that attributes should be usually plain declarations such as @{attribute simp}, while non-trivial rules like @{attribute simplified} are better avoided. \<^descr> \<^ML>\Local_Theory.note\~\(a, ths) lthy\ is analogous to \<^ML>\Local_Theory.define\, but defines facts instead of terms. There is also a slightly more general variant \<^ML>\Local_Theory.notes\ that defines several facts (with attribute expressions) simultaneously. This is essentially the internal version of the @{command lemmas} command, or @{command declare} if an empty name binding is given. \ section \Morphisms and declarations \label{sec:morphisms}\ text \ %FIXME - See also @{cite "Chaieb-Wenzel:2007"}. + See also \<^cite>\"Chaieb-Wenzel:2007"\. \ end diff --git a/src/Doc/Implementation/Logic.thy b/src/Doc/Implementation/Logic.thy --- a/src/Doc/Implementation/Logic.thy +++ b/src/Doc/Implementation/Logic.thy @@ -1,1365 +1,1362 @@ (*:maxLineLen=78:*) theory Logic imports Base begin chapter \Primitive logic \label{ch:logic}\ text \ The logical foundations of Isabelle/Isar are that of the Pure logic, which - has been introduced as a Natural Deduction framework in @{cite paulson700}. + has been introduced as a Natural Deduction framework in \<^cite>\paulson700\. This is essentially the same logic as ``\\HOL\'' in the more abstract - setting of Pure Type Systems (PTS) @{cite "Barendregt-Geuvers:2001"}, + setting of Pure Type Systems (PTS) \<^cite>\"Barendregt-Geuvers:2001"\, although there are some key differences in the specific treatment of simple types in Isabelle/Pure. Following type-theoretic parlance, the Pure logic consists of three levels of \\\-calculus with corresponding arrows, \\\ for syntactic function space (terms depending on terms), \\\ for universal quantification (proofs depending on terms), and \\\ for implication (proofs depending on proofs). Derivations are relative to a logical theory, which declares type constructors, constants, and axioms. Theory declarations support schematic polymorphism, which is strictly speaking outside the logic.\<^footnote>\This is the deeper logical reason, why the theory context \\\ is separate from the proof context \\\ of the core calculus: type constructors, term constants, and facts (proof constants) may involve arbitrary type schemes, but the type of a locally fixed term parameter is also fixed!\ \ section \Types \label{sec:types}\ text \ The language of types is an uninterpreted order-sorted first-order algebra; types are qualified by ordered type classes. \<^medskip> A \<^emph>\type class\ is an abstract syntactic entity declared in the theory context. The \<^emph>\subclass relation\ \c\<^sub>1 \ c\<^sub>2\ is specified by stating an acyclic generating relation; the transitive closure is maintained internally. The resulting relation is an ordering: reflexive, transitive, and antisymmetric. A \<^emph>\sort\ is a list of type classes written as \s = {c\<^sub>1, \, c\<^sub>m}\, it represents symbolic intersection. Notationally, the curly braces are omitted for singleton intersections, i.e.\ any class \c\ may be read as a sort \{c}\. The ordering on type classes is extended to sorts according to the meaning of intersections: \{c\<^sub>1, \ c\<^sub>m} \ {d\<^sub>1, \, d\<^sub>n}\ iff \\j. \i. c\<^sub>i \ d\<^sub>j\. The empty intersection \{}\ refers to the universal sort, which is the largest element wrt.\ the sort order. Thus \{}\ represents the ``full sort'', not the empty one! The intersection of all (finitely many) classes declared in the current theory is the least element wrt.\ the sort ordering. \<^medskip> A \<^emph>\fixed type variable\ is a pair of a basic name (starting with a \'\ character) and a sort constraint, e.g.\ \('a, s)\ which is usually printed as \\\<^sub>s\. A \<^emph>\schematic type variable\ is a pair of an indexname and a sort constraint, e.g.\ \(('a, 0), s)\ which is usually printed as \?\\<^sub>s\. Note that \<^emph>\all\ syntactic components contribute to the identity of type variables: basic name, index, and sort constraint. The core logic handles type variables with the same name but different sorts as different, although the type-inference layer (which is outside the core) rejects anything like that. A \<^emph>\type constructor\ \\\ is a \k\-ary operator on types declared in the theory. Type constructor application is written postfix as \(\\<^sub>1, \, \\<^sub>k)\\. For \k = 0\ the argument tuple is omitted, e.g.\ \prop\ instead of \()prop\. For \k = 1\ the parentheses are omitted, e.g.\ \\ list\ instead of \(\)list\. Further notation is provided for specific constructors, notably the right-associative infix \\ \ \\ instead of \(\, \)fun\. The logical category \<^emph>\type\ is defined inductively over type variables and type constructors as follows: \\ = \\<^sub>s | ?\\<^sub>s | (\\<^sub>1, \, \\<^sub>k)\\. A \<^emph>\type abbreviation\ is a syntactic definition \(\<^vec>\)\ = \\ of an arbitrary type expression \\\ over variables \\<^vec>\\. Type abbreviations appear as type constructors in the syntax, but are expanded before entering the logical core. A \<^emph>\type arity\ declares the image behavior of a type constructor wrt.\ the algebra of sorts: \\ :: (s\<^sub>1, \, s\<^sub>k)s\ means that \(\\<^sub>1, \, \\<^sub>k)\\ is of sort \s\ if every argument type \\\<^sub>i\ is of sort \s\<^sub>i\. Arity declarations are implicitly completed, i.e.\ \\ :: (\<^vec>s)c\ entails \\ :: (\<^vec>s)c'\ for any \c' \ c\. \<^medskip> The sort algebra is always maintained as \<^emph>\coregular\, which means that type arities are consistent with the subclass relation: for any type constructor \\\, and classes \c\<^sub>1 \ c\<^sub>2\, and arities \\ :: (\<^vec>s\<^sub>1)c\<^sub>1\ and \\ :: (\<^vec>s\<^sub>2)c\<^sub>2\ holds \\<^vec>s\<^sub>1 \ \<^vec>s\<^sub>2\ component-wise. The key property of a coregular order-sorted algebra is that sort constraints can be solved in a most general fashion: for each type constructor \\\ and sort \s\ there is a most general vector of argument sorts \(s\<^sub>1, \, s\<^sub>k)\ such that a type scheme \(\\<^bsub>s\<^sub>1\<^esub>, \, \\<^bsub>s\<^sub>k\<^esub>)\\ is of sort \s\. Consequently, type unification has most general solutions (modulo equivalence of sorts), so type-inference produces primary types as expected - @{cite "nipkow-prehofer"}. + \<^cite>\"nipkow-prehofer"\. \ text %mlref \ \begin{mldecls} @{define_ML_type class = string} \\ @{define_ML_type sort = "class list"} \\ @{define_ML_type arity = "string * sort list * sort"} \\ @{define_ML_type typ} \\ @{define_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\ @{define_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{define_ML Sign.subsort: "theory -> sort * sort -> bool"} \\ @{define_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\ @{define_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\ @{define_ML Sign.add_type_abbrev: "Proof.context -> binding * string list * typ -> theory -> theory"} \\ @{define_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\ @{define_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\ @{define_ML Sign.primitive_arity: "arity -> theory -> theory"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\class\ represents type classes. \<^descr> Type \<^ML_type>\sort\ represents sorts, i.e.\ finite intersections of classes. The empty list \<^ML>\[]: sort\ refers to the empty class intersection, i.e.\ the ``full sort''. \<^descr> Type \<^ML_type>\arity\ represents type arities. A triple \(\, \<^vec>s, s) : arity\ represents \\ :: (\<^vec>s)s\ as described above. \<^descr> Type \<^ML_type>\typ\ represents types; this is a datatype with constructors \<^ML>\TFree\, \<^ML>\TVar\, \<^ML>\Type\. \<^descr> \<^ML>\Term.map_atyps\~\f \\ applies the mapping \f\ to all atomic types (\<^ML>\TFree\, \<^ML>\TVar\) occurring in \\\. \<^descr> \<^ML>\Term.fold_atyps\~\f \\ iterates the operation \f\ over all occurrences of atomic types (\<^ML>\TFree\, \<^ML>\TVar\) in \\\; the type structure is traversed from left to right. \<^descr> \<^ML>\Sign.subsort\~\thy (s\<^sub>1, s\<^sub>2)\ tests the subsort relation \s\<^sub>1 \ s\<^sub>2\. \<^descr> \<^ML>\Sign.of_sort\~\thy (\, s)\ tests whether type \\\ is of sort \s\. \<^descr> \<^ML>\Sign.add_type\~\ctxt (\, k, mx)\ declares a new type constructors \\\ with \k\ arguments and optional mixfix syntax. \<^descr> \<^ML>\Sign.add_type_abbrev\~\ctxt (\, \<^vec>\, \)\ defines a new type abbreviation \(\<^vec>\)\ = \\. \<^descr> \<^ML>\Sign.primitive_class\~\(c, [c\<^sub>1, \, c\<^sub>n])\ declares a new class \c\, together with class relations \c \ c\<^sub>i\, for \i = 1, \, n\. \<^descr> \<^ML>\Sign.primitive_classrel\~\(c\<^sub>1, c\<^sub>2)\ declares the class relation \c\<^sub>1 \ c\<^sub>2\. \<^descr> \<^ML>\Sign.primitive_arity\~\(\, \<^vec>s, s)\ declares the arity \\ :: (\<^vec>s)s\. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "class"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "sort"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "type_name"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "type_abbrev"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "nonterminal"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "typ"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Type"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Type_fn"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation class} embedded ; @@{ML_antiquotation sort} sort ; (@@{ML_antiquotation type_name} | @@{ML_antiquotation type_abbrev} | @@{ML_antiquotation nonterminal}) embedded ; @@{ML_antiquotation typ} type ; (@@{ML_antiquotation Type} | @@{ML_antiquotation Type_fn}) embedded \ \<^descr> \@{class c}\ inlines the internalized class \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{sort s}\ inlines the internalized sort \s\ --- as \<^ML_type>\string list\ literal. \<^descr> \@{type_name c}\ inlines the internalized type constructor \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{type_abbrev c}\ inlines the internalized type abbreviation \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{nonterminal c}\ inlines the internalized syntactic type~/ grammar nonterminal \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{typ \}\ inlines the internalized type \\\ --- as constructor term for datatype \<^ML_type>\typ\. \<^descr> \@{Type source}\ refers to embedded source text to produce a type constructor with name (formally checked) and arguments (inlined ML text). The embedded \source\ follows the syntax category @{syntax type_const} defined below. \<^descr> \@{Type_fn source}\ is similar to \@{Type source}\, but produces a partial ML function that matches against a type constructor pattern, following syntax category @{syntax type_const_fn} below. \<^rail>\ @{syntax_def type_const}: @{syntax name} (@{syntax embedded_ml}*) ; @{syntax_def type_const_fn}: @{syntax type_const} @'=>' @{syntax embedded_ml} ; @{syntax_def embedded_ml}: @'_' | @{syntax embedded} | @{syntax control_symbol} @{syntax embedded} \ The text provided as @{syntax embedded_ml} is treated as regular Isabelle/ML source, possibly with nested antiquotations. ML text that only consists of a single antiquotation in compact control-cartouche notation, can be written without an extra level of nesting embedded text (see the last example below). \ text %mlex \ Here are some minimal examples for type constructor antiquotations. \ ML \ \ \type constructor without arguments:\ val natT = \<^Type>\nat\; \ \type constructor with arguments:\ fun mk_funT (A, B) = \<^Type>\fun A B\; \ \type constructor decomposition as partial function:\ val dest_funT = \<^Type_fn>\fun A B => \(A, B)\\; \ \nested type constructors:\ fun mk_relT A = \<^Type>\fun A \<^Type>\fun A \<^Type>\bool\\\; \<^assert> (mk_relT natT = \<^typ>\nat \ nat \ bool\); \ section \Terms \label{sec:terms}\ text \ The language of terms is that of simply-typed \\\-calculus with de-Bruijn - indices for bound variables (cf.\ @{cite debruijn72} or @{cite - "paulson-ml2"}), with the types being determined by the corresponding + indices for bound variables (cf.\ \<^cite>\debruijn72\ or \<^cite>\"paulson-ml2"\), with the types being determined by the corresponding binders. In contrast, free variables and constants have an explicit name and type in each occurrence. \<^medskip> A \<^emph>\bound variable\ is a natural number \b\, which accounts for the number of intermediate binders between the variable occurrence in the body and its binding position. For example, the de-Bruijn term \\\<^bsub>bool\<^esub>. \\<^bsub>bool\<^esub>. 1 \ 0\ would correspond to \\x\<^bsub>bool\<^esub>. \y\<^bsub>bool\<^esub>. x \ y\ in a named representation. Note that a bound variable may be represented by different de-Bruijn indices at different occurrences, depending on the nesting of abstractions. A \<^emph>\loose variable\ is a bound variable that is outside the scope of local binders. The types (and names) for loose variables can be managed as a separate context, that is maintained as a stack of hypothetical binders. The core logic operates on closed terms, without any loose variables. A \<^emph>\fixed variable\ is a pair of a basic name and a type, e.g.\ \(x, \)\ which is usually printed \x\<^sub>\\ here. A \<^emph>\schematic variable\ is a pair of an indexname and a type, e.g.\ \((x, 0), \)\ which is likewise printed as \?x\<^sub>\\. \<^medskip> A \<^emph>\constant\ is a pair of a basic name and a type, e.g.\ \(c, \)\ which is usually printed as \c\<^sub>\\ here. Constants are declared in the context as polymorphic families \c :: \\, meaning that all substitution instances \c\<^sub>\\ for \\ = \\\ are valid. The vector of \<^emph>\type arguments\ of constant \c\<^sub>\\ wrt.\ the declaration \c :: \\ is defined as the codomain of the matcher \\ = {?\\<^sub>1 \ \\<^sub>1, \, ?\\<^sub>n \ \\<^sub>n}\ presented in canonical order \(\\<^sub>1, \, \\<^sub>n)\, corresponding to the left-to-right occurrences of the \\\<^sub>i\ in \\\. Within a given theory context, there is a one-to-one correspondence between any constant \c\<^sub>\\ and the application \c(\\<^sub>1, \, \\<^sub>n)\ of its type arguments. For example, with \plus :: \ \ \ \ \\, the instance \plus\<^bsub>nat \ nat \ nat\<^esub>\ corresponds to \plus(nat)\. Constant declarations \c :: \\ may contain sort constraints for type variables in \\\. These are observed by type-inference as expected, but \<^emph>\ignored\ by the core logic. This means the primitive logic is able to reason with instances of polymorphic constants that the user-level type-checker would reject due to violation of type class restrictions. \<^medskip> An \<^emph>\atomic term\ is either a variable or constant. The logical category \<^emph>\term\ is defined inductively over atomic terms, with abstraction and application as follows: \t = b | x\<^sub>\ | ?x\<^sub>\ | c\<^sub>\ | \\<^sub>\. t | t\<^sub>1 t\<^sub>2\. Parsing and printing takes care of converting between an external representation with named bound variables. Subsequently, we shall use the latter notation instead of internal de-Bruijn representation. The inductive relation \t :: \\ assigns a (unique) type to a term according to the structure of atomic terms, abstractions, and applications: \[ \infer{\a\<^sub>\ :: \\}{} \qquad \infer{\(\x\<^sub>\. t) :: \ \ \\}{\t :: \\} \qquad \infer{\t u :: \\}{\t :: \ \ \\ & \u :: \\} \] A \<^emph>\well-typed term\ is a term that can be typed according to these rules. Typing information can be omitted: type-inference is able to reconstruct the most general type of a raw term, while assigning most general types to all of its variables and constants. Type-inference depends on a context of type constraints for fixed variables, and declarations for polymorphic constants. The identity of atomic terms consists both of the name and the type component. This means that different variables \x\<^bsub>\\<^sub>1\<^esub>\ and \x\<^bsub>\\<^sub>2\<^esub>\ may become the same after type instantiation. Type-inference rejects variables of the same name, but different types. In contrast, mixed instances of polymorphic constants occur routinely. \<^medskip> The \<^emph>\hidden polymorphism\ of a term \t :: \\ is the set of type variables occurring in \t\, but not in its type \\\. This means that the term implicitly depends on type arguments that are not accounted in the result type, i.e.\ there are different type instances \t\ :: \\ and \t\' :: \\ with the same type. This slightly pathological situation notoriously demands additional care. \<^medskip> A \<^emph>\term abbreviation\ is a syntactic definition \c\<^sub>\ \ t\ of a closed term \t\ of type \\\, without any hidden polymorphism. A term abbreviation looks like a constant in the syntax, but is expanded before entering the logical core. Abbreviations are usually reverted when printing terms, using \t \ c\<^sub>\\ as rules for higher-order rewriting. \<^medskip> Canonical operations on \\\-terms include \\\\\-conversion: \\\-conversion refers to capture-free renaming of bound variables; \\\-conversion contracts an abstraction applied to an argument term, substituting the argument in the body: \(\x. b)a\ becomes \b[a/x]\; \\\-conversion contracts vacuous application-abstraction: \\x. f x\ becomes \f\, provided that the bound variable does not occur in \f\. Terms are normally treated modulo \\\-conversion, which is implicit in the de-Bruijn representation. Names for bound variables in abstractions are maintained separately as (meaningless) comments, mostly for parsing and printing. Full \\\\\-conversion is commonplace in various standard operations (\secref{sec:obj-rules}) that are based on higher-order unification and matching. \ text %mlref \ \begin{mldecls} @{define_ML_type term} \\ @{define_ML_infix "aconv": "term * term -> bool"} \\ @{define_ML Term.map_types: "(typ -> typ) -> term -> term"} \\ @{define_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ @{define_ML Term.map_aterms: "(term -> term) -> term -> term"} \\ @{define_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{define_ML fastype_of: "term -> typ"} \\ @{define_ML lambda: "term -> term -> term"} \\ @{define_ML betapply: "term * term -> term"} \\ @{define_ML incr_boundvars: "int -> term -> term"} \\ @{define_ML Sign.declare_const: "Proof.context -> (binding * typ) * mixfix -> theory -> term * theory"} \\ @{define_ML Sign.add_abbrev: "string -> binding * term -> theory -> (term * term) * theory"} \\ @{define_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\ @{define_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\term\ represents de-Bruijn terms, with comments in abstractions, and explicitly named free variables and constants; this is a datatype with constructors @{define_ML Bound}, @{define_ML Free}, @{define_ML Var}, @{define_ML Const}, @{define_ML Abs}, @{define_ML_infix "$"}. \<^descr> \t\~\<^ML_text>\aconv\~\u\ checks \\\-equivalence of two terms. This is the basic equality relation on type \<^ML_type>\term\; raw datatype equality should only be used for operations related to parsing or printing! \<^descr> \<^ML>\Term.map_types\~\f t\ applies the mapping \f\ to all types occurring in \t\. \<^descr> \<^ML>\Term.fold_types\~\f t\ iterates the operation \f\ over all occurrences of types in \t\; the term structure is traversed from left to right. \<^descr> \<^ML>\Term.map_aterms\~\f t\ applies the mapping \f\ to all atomic terms (\<^ML>\Bound\, \<^ML>\Free\, \<^ML>\Var\, \<^ML>\Const\) occurring in \t\. \<^descr> \<^ML>\Term.fold_aterms\~\f t\ iterates the operation \f\ over all occurrences of atomic terms (\<^ML>\Bound\, \<^ML>\Free\, \<^ML>\Var\, \<^ML>\Const\) in \t\; the term structure is traversed from left to right. \<^descr> \<^ML>\fastype_of\~\t\ determines the type of a well-typed term. This operation is relatively slow, despite the omission of any sanity checks. \<^descr> \<^ML>\lambda\~\a b\ produces an abstraction \\a. b\, where occurrences of the atomic term \a\ in the body \b\ are replaced by bound variables. \<^descr> \<^ML>\betapply\~\(t, u)\ produces an application \t u\, with topmost \\\-conversion if \t\ is an abstraction. \<^descr> \<^ML>\incr_boundvars\~\j\ increments a term's dangling bound variables by the offset \j\. This is required when moving a subterm into a context where it is enclosed by a different number of abstractions. Bound variables with a matching abstraction are unaffected. \<^descr> \<^ML>\Sign.declare_const\~\ctxt ((c, \), mx)\ declares a new constant \c :: \\ with optional mixfix syntax. \<^descr> \<^ML>\Sign.add_abbrev\~\print_mode (c, t)\ introduces a new term abbreviation \c \ t\. \<^descr> \<^ML>\Sign.const_typargs\~\thy (c, \)\ and \<^ML>\Sign.const_instance\~\thy (c, [\\<^sub>1, \, \\<^sub>n])\ convert between two representations of polymorphic constants: full type instance vs.\ compact type arguments form. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "const_name"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "const_abbrev"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "const"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "term"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "prop"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Const"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Const_"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Const_fn"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ (@@{ML_antiquotation const_name} | @@{ML_antiquotation const_abbrev}) embedded ; @@{ML_antiquotation const} ('(' (type + ',') ')')? ; @@{ML_antiquotation term} term ; @@{ML_antiquotation prop} prop ; (@@{ML_antiquotation Const} | @@{ML_antiquotation Const_} | @@{ML_antiquotation Const_fn}) embedded \ \<^descr> \@{const_name c}\ inlines the internalized logical constant name \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{const_abbrev c}\ inlines the internalized abbreviated constant name \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{const c(\<^vec>\)}\ inlines the internalized constant \c\ with precise type instantiation in the sense of \<^ML>\Sign.const_instance\ --- as \<^ML>\Const\ constructor term for datatype \<^ML_type>\term\. \<^descr> \@{term t}\ inlines the internalized term \t\ --- as constructor term for datatype \<^ML_type>\term\. \<^descr> \@{prop \}\ inlines the internalized proposition \\\ --- as constructor term for datatype \<^ML_type>\term\. \<^descr> \@{Const source}\ refers to embedded source text to produce a term constructor with name (formally checked), type arguments and term arguments (inlined ML text). The embedded \source\ follows the syntax category @{syntax term_const} defined below, using @{syntax embedded_ml} from antiquotation @{ML_antiquotation Type} (\secref{sec:types}). \<^descr> \@{Const_ source}\ is similar to \@{Const source}\ for patterns instead of closed values. This distinction is required due to redundant type information within term constants: subtrees with duplicate ML pattern variable need to be suppressed (replaced by dummy patterns). The embedded \source\ follows the syntax category @{syntax term_const} defined below. \<^descr> \@{Const_fn source}\ is similar to \@{Const_ source}\, but produces a proper \<^verbatim>\fn\ expression with function body. The embedded \source\ follows the syntax category @{syntax term_const_fn} defined below. \<^rail>\ @{syntax_def term_const}: @{syntax name} (@{syntax embedded_ml}*) (@{syntax for_args})? ; @{syntax_def term_const_fn}: @{syntax term_const} @'=>' @{syntax embedded_ml} ; @{syntax_def for_args}: @'for' (@{syntax embedded_ml}+) \ \ text %mlex \ Here are some minimal examples for term constant antiquotations. Type arguments for constants are analogous to type constructors (\secref{sec:types}). Term arguments are different: a flexible number of arguments is inserted via curried applications (\<^ML>\op $\). \ ML \ \ \constant without type argument:\ fun mk_conj (A, B) = \<^Const>\conj for A B\; val dest_conj = \<^Const_fn>\conj for A B => \(A, B)\\; \ \constant with type argument:\ fun mk_eq T (t, u) = \<^Const>\HOL.eq T for t u\; val dest_eq = \<^Const_fn>\HOL.eq T for t u => \(T, (t, u))\\; \ \constant with variable number of term arguments:\ val c = Const (\<^const_name>\conj\, \<^typ>\bool \ bool \ bool\); val P = \<^term>\P::bool\; val Q = \<^term>\Q::bool\; \<^assert> (\<^Const>\conj\ = c); \<^assert> (\<^Const>\conj for P\ = c $ P); \<^assert> (\<^Const>\conj for P Q\ = c $ P $ Q); \ section \Theorems \label{sec:thms}\ text \ A \<^emph>\proposition\ is a well-typed term of type \prop\, a \<^emph>\theorem\ is a proven proposition (depending on a context of hypotheses and the background theory). Primitive inferences include plain Natural Deduction rules for the primary connectives \\\ and \\\ of the framework. There is also a builtin notion of equality/equivalence \\\. \ subsection \Primitive connectives and rules \label{sec:prim-rules}\ text \ The theory \Pure\ contains constant declarations for the primitive connectives \\\, \\\, and \\\ of the logical framework, see \figref{fig:pure-connectives}. The derivability judgment \A\<^sub>1, \, A\<^sub>n \ B\ is defined inductively by the primitive inferences given in \figref{fig:prim-rules}, with the global restriction that the hypotheses must \<^emph>\not\ contain any schematic variables. The builtin equality is conceptually axiomatized as shown in \figref{fig:pure-equality}, although the implementation works directly with derived inferences. \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \all :: (\ \ prop) \ prop\ & universal quantification (binder \\\) \\ \\ :: prop \ prop \ prop\ & implication (right associative infix) \\ \\ :: \ \ \ \ prop\ & equality relation (infix) \\ \end{tabular} \caption{Primitive connectives of Pure}\label{fig:pure-connectives} \end{center} \end{figure} \begin{figure}[htb] \begin{center} \[ \infer[\(axiom)\]{\\ A\}{\A \ \\} \qquad \infer[\(assume)\]{\A \ A\}{} \] \[ \infer[\(\\intro)\]{\\ \ \x. B[x]\}{\\ \ B[x]\ & \x \ \\} \qquad \infer[\(\\elim)\]{\\ \ B[a]\}{\\ \ \x. B[x]\} \] \[ \infer[\(\\intro)\]{\\ - A \ A \ B\}{\\ \ B\} \qquad \infer[\(\\elim)\]{\\\<^sub>1 \ \\<^sub>2 \ B\}{\\\<^sub>1 \ A \ B\ & \\\<^sub>2 \ A\} \] \caption{Primitive inferences of Pure}\label{fig:prim-rules} \end{center} \end{figure} \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \\ (\x. b[x]) a \ b[a]\ & \\\-conversion \\ \\ x \ x\ & reflexivity \\ \\ x \ y \ P x \ P y\ & substitution \\ \\ (\x. f x \ g x) \ f \ g\ & extensionality \\ \\ (A \ B) \ (B \ A) \ A \ B\ & logical equivalence \\ \end{tabular} \caption{Conceptual axiomatization of Pure equality}\label{fig:pure-equality} \end{center} \end{figure} The introduction and elimination rules for \\\ and \\\ are analogous to formation of dependently typed \\\-terms representing the underlying proof objects. Proof terms are irrelevant in the Pure logic, though; they cannot occur within propositions. The system provides a runtime option to record explicit proof terms for primitive inferences, see also \secref{sec:proof-terms}. Thus all three levels of \\\-calculus become - explicit: \\\ for terms, and \\/\\ for proofs (cf.\ @{cite - "Berghofer-Nipkow:2000:TPHOL"}). + explicit: \\\ for terms, and \\/\\ for proofs (cf.\ \<^cite>\"Berghofer-Nipkow:2000:TPHOL"\). Observe that locally fixed parameters (as in \\\intro\) need not be recorded in the hypotheses, because the simple syntactic types of Pure are always inhabitable. ``Assumptions'' \x :: \\ for type-membership are only present as long as some \x\<^sub>\\ occurs in the statement body.\<^footnote>\This is the key - difference to ``\\HOL\'' in the PTS framework @{cite - "Barendregt-Geuvers:2001"}, where hypotheses \x : A\ are treated uniformly + difference to ``\\HOL\'' in the PTS framework \<^cite>\"Barendregt-Geuvers:2001"\, where hypotheses \x : A\ are treated uniformly for propositions and types.\ \<^medskip> The axiomatization of a theory is implicitly closed by forming all instances of type and term variables: \\ A\\ holds for any substitution instance of an axiom \\ A\. By pushing substitutions through derivations inductively, we also get admissible \generalize\ and \instantiate\ rules as shown in \figref{fig:subst-rules}. \begin{figure}[htb] \begin{center} \[ \infer{\\ \ B[?\]\}{\\ \ B[\]\ & \\ \ \\} \quad \infer[\quad\(generalize)\]{\\ \ B[?x]\}{\\ \ B[x]\ & \x \ \\} \] \[ \infer{\\ \ B[\]\}{\\ \ B[?\]\} \quad \infer[\quad\(instantiate)\]{\\ \ B[t]\}{\\ \ B[?x]\} \] \caption{Admissible substitution rules}\label{fig:subst-rules} \end{center} \end{figure} Note that \instantiate\ does not require an explicit side-condition, because \\\ may never contain schematic variables. In principle, variables could be substituted in hypotheses as well, but this would disrupt the monotonicity of reasoning: deriving \\\ \ B\\ from \\ \ B\ is correct, but \\\ \ \\ does not necessarily hold: the result belongs to a different proof context. \<^medskip> An \<^emph>\oracle\ is a function that produces axioms on the fly. Logically, this is an instance of the \axiom\ rule (\figref{fig:prim-rules}), but there is an operational difference. The inference kernel records oracle invocations within derivations of theorems by a unique tag. This also includes implicit type-class reasoning via the order-sorted algebra of class relations and type arities (see also @{command_ref instantiation} and @{command_ref instance}). Axiomatizations should be limited to the bare minimum, typically as part of the initial logical basis of an object-logic formalization. Later on, theories are usually developed in a strictly definitional fashion, by stating only certain equalities over new constants. A \<^emph>\simple definition\ consists of a constant declaration \c :: \\ together with an axiom \\ c \ t\, where \t :: \\ is a closed term without any hidden polymorphism. The RHS may depend on further defined constants, but not \c\ itself. Definitions of functions may be presented as \c \<^vec>x \ t\ instead of the puristic \c \ \\<^vec>x. t\. An \<^emph>\overloaded definition\ consists of a collection of axioms for the same constant, with zero or one equations \c((\<^vec>\)\) \ t\ for each type constructor \\\ (for distinct variables \\<^vec>\\). The RHS may mention previously defined constants as above, or arbitrary constants \d(\\<^sub>i)\ for some \\\<^sub>i\ projected from \\<^vec>\\. Thus overloaded definitions essentially work by primitive recursion over the syntactic structure of a - single type argument. See also @{cite \\S4.3\ - "Haftmann-Wenzel:2006:classes"}. + single type argument. See also \<^cite>\\\S4.3\ in + "Haftmann-Wenzel:2006:classes"\. \ text %mlref \ \begin{mldecls} @{define_ML Logic.all: "term -> term -> term"} \\ @{define_ML Logic.mk_implies: "term * term -> term"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type ctyp} \\ @{define_ML_type cterm} \\ @{define_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\ @{define_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\ @{define_ML Thm.apply: "cterm -> cterm -> cterm"} \\ @{define_ML Thm.lambda: "cterm -> cterm -> cterm"} \\ @{define_ML Thm.all: "Proof.context -> cterm -> cterm -> cterm"} \\ @{define_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type thm} \\ @{define_ML Thm.transfer: "theory -> thm -> thm"} \\ @{define_ML Thm.assume: "cterm -> thm"} \\ @{define_ML Thm.forall_intr: "cterm -> thm -> thm"} \\ @{define_ML Thm.forall_elim: "cterm -> thm -> thm"} \\ @{define_ML Thm.implies_intr: "cterm -> thm -> thm"} \\ @{define_ML Thm.implies_elim: "thm -> thm -> thm"} \\ @{define_ML Thm.generalize: "Names.set * Names.set -> int -> thm -> thm"} \\ @{define_ML Thm.instantiate: "ctyp TVars.table * cterm Vars.table -> thm -> thm"} \\ @{define_ML Thm.add_axiom: "Proof.context -> binding * term -> theory -> (string * thm) * theory"} \\ @{define_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory"} \\ @{define_ML Thm.add_def: "Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory"} \\ \end{mldecls} \begin{mldecls} @{define_ML Theory.add_deps: "Defs.context -> string -> Defs.entry -> Defs.entry list -> theory -> theory"} \\ @{define_ML Thm_Deps.all_oracles: "thm list -> Proofterm.oracle list"} \\ \end{mldecls} \<^descr> \<^ML>\Logic.all\~\a B\ produces a Pure quantification \\a. B\, where occurrences of the atomic term \a\ in the body proposition \B\ are replaced by bound variables. (See also \<^ML>\lambda\ on terms.) \<^descr> \<^ML>\Logic.mk_implies\~\(A, B)\ produces a Pure implication \A \ B\. \<^descr> Types \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ represent certified types and terms, respectively. These are abstract datatypes that guarantee that its values have passed the full well-formedness (and well-typedness) checks, relative to the declarations of type constructors, constants etc.\ in the background theory. The abstract types \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ are part of the same inference kernel that is mainly responsible for \<^ML_type>\thm\. Thus syntactic operations on \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ are located in the \<^ML_structure>\Thm\ module, even though theorems are not yet involved at that stage. \<^descr> \<^ML>\Thm.ctyp_of\~\ctxt \\ and \<^ML>\Thm.cterm_of\~\ctxt t\ explicitly check types and terms, respectively. This also involves some basic normalizations, such expansion of type and term abbreviations from the underlying theory context. Full re-certification is relatively slow and should be avoided in tight reasoning loops. \<^descr> \<^ML>\Thm.apply\, \<^ML>\Thm.lambda\, \<^ML>\Thm.all\, \<^ML>\Drule.mk_implies\ etc.\ compose certified terms (or propositions) incrementally. This is equivalent to \<^ML>\Thm.cterm_of\ after unchecked \<^ML_infix>\$\, \<^ML>\lambda\, \<^ML>\Logic.all\, \<^ML>\Logic.mk_implies\ etc., but there can be a big difference in performance when large existing entities are composed by a few extra constructions on top. There are separate operations to decompose certified terms and theorems to produce certified terms again. \<^descr> Type \<^ML_type>\thm\ represents proven propositions. This is an abstract datatype that guarantees that its values have been constructed by basic principles of the \<^ML_structure>\Thm\ module. Every \<^ML_type>\thm\ value refers its background theory, cf.\ \secref{sec:context-theory}. \<^descr> \<^ML>\Thm.transfer\~\thy thm\ transfers the given theorem to a \<^emph>\larger\ theory, see also \secref{sec:context}. This formal adjustment of the background context has no logical significance, but is occasionally required for formal reasons, e.g.\ when theorems that are imported from more basic theories are used in the current situation. \<^descr> \<^ML>\Thm.assume\, \<^ML>\Thm.forall_intr\, \<^ML>\Thm.forall_elim\, \<^ML>\Thm.implies_intr\, and \<^ML>\Thm.implies_elim\ correspond to the primitive inferences of \figref{fig:prim-rules}. \<^descr> \<^ML>\Thm.generalize\~\(\<^vec>\, \<^vec>x)\ corresponds to the \generalize\ rules of \figref{fig:subst-rules}. Here collections of type and term variables are generalized simultaneously, specified by the given sets of basic names. \<^descr> \<^ML>\Thm.instantiate\~\(\<^vec>\\<^sub>s, \<^vec>x\<^sub>\)\ corresponds to the \instantiate\ rules of \figref{fig:subst-rules}. Type variables are substituted before term variables. Note that the types in \\<^vec>x\<^sub>\\ refer to the instantiated versions. \<^descr> \<^ML>\Thm.add_axiom\~\ctxt (name, A)\ declares an arbitrary proposition as axiom, and retrieves it as a theorem from the resulting theory, cf.\ \axiom\ in \figref{fig:prim-rules}. Note that the low-level representation in the axiom table may differ slightly from the returned theorem. \<^descr> \<^ML>\Thm.add_oracle\~\(binding, oracle)\ produces a named oracle rule, essentially generating arbitrary axioms on the fly, cf.\ \axiom\ in \figref{fig:prim-rules}. \<^descr> \<^ML>\Thm.add_def\~\ctxt unchecked overloaded (name, c \<^vec>x \ t)\ states a definitional axiom for an existing constant \c\. Dependencies are recorded via \<^ML>\Theory.add_deps\, unless the \unchecked\ option is set. Note that the low-level representation in the axiom table may differ slightly from the returned theorem. \<^descr> \<^ML>\Theory.add_deps\~\ctxt name c\<^sub>\ \<^vec>d\<^sub>\\ declares dependencies of a named specification for constant \c\<^sub>\\, relative to existing specifications for constants \\<^vec>d\<^sub>\\. This also works for type constructors. \<^descr> \<^ML>\Thm_Deps.all_oracles\~\thms\ returns all oracles used in the internal derivation of the given theorems; this covers the full graph of transitive dependencies. The result contains an authentic oracle name; if @{ML Proofterm.proofs} was at least @{ML 1} during the oracle inference, it also contains the position of the oracle invocation and its proposition. See also the command @{command_ref "thm_oracles"}. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "ctyp"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "cterm"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "cprop"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "thm"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "thms"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "lemma"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "oracle_name"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation ctyp} typ ; @@{ML_antiquotation cterm} term ; @@{ML_antiquotation cprop} prop ; @@{ML_antiquotation thm} thm ; @@{ML_antiquotation thms} thms ; @@{ML_antiquotation lemma} ('(' @'open' ')')? ((prop +) + @'and') \ @{syntax for_fixes} @'by' method method? ; @@{ML_antiquotation oracle_name} embedded \ \<^descr> \@{ctyp \}\ produces a certified type wrt.\ the current background theory --- as abstract value of type \<^ML_type>\ctyp\. \<^descr> \@{cterm t}\ and \@{cprop \}\ produce a certified term wrt.\ the current background theory --- as abstract value of type \<^ML_type>\cterm\. \<^descr> \@{thm a}\ produces a singleton fact --- as abstract value of type \<^ML_type>\thm\. \<^descr> \@{thms a}\ produces a general fact --- as abstract value of type \<^ML_type>\thm list\. \<^descr> \@{lemma \ by meth}\ produces a fact that is proven on the spot according to the minimal proof, which imitates a terminal Isar proof. The result is an abstract value of type \<^ML_type>\thm\ or \<^ML_type>\thm list\, depending on the number of propositions given here. The internal derivation object lacks a proper theorem name, but it is formally closed, unless the \(open)\ option is specified (this may impact performance of applications with proof terms). Since ML antiquotations are always evaluated at compile-time, there is no run-time overhead even for non-trivial proofs. Nonetheless, the justification is syntactically limited to a single @{command "by"} step. More complex Isar proofs should be done in regular theory source, before compiling the corresponding ML text that uses the result. \<^descr> \@{oracle_name a}\ inlines the internalized oracle name \a\ --- as \<^ML_type>\string\ literal. \ subsection \Auxiliary connectives \label{sec:logic-aux}\ text \ Theory \Pure\ provides a few auxiliary connectives that are defined on top of the primitive ones, see \figref{fig:pure-aux}. These special constants are useful in certain internal encodings, and are normally not directly exposed to the user. \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \conjunction :: prop \ prop \ prop\ & (infix \&&&\) \\ \\ A &&& B \ (\C. (A \ B \ C) \ C)\ \\[1ex] \prop :: prop \ prop\ & (prefix \#\, suppressed) \\ \#A \ A\ \\[1ex] \term :: \ \ prop\ & (prefix \TERM\) \\ \term x \ (\A. A \ A)\ \\[1ex] \type :: \ itself\ & (prefix \TYPE\) \\ \(unspecified)\ \\ \end{tabular} \caption{Definitions of auxiliary connectives}\label{fig:pure-aux} \end{center} \end{figure} The introduction \A \ B \ A &&& B\, and eliminations (projections) \A &&& B \ A\ and \A &&& B \ B\ are available as derived rules. Conjunction allows to treat simultaneous assumptions and conclusions uniformly, e.g.\ consider \A \ B \ C &&& D\. In particular, the goal mechanism represents multiple claims as explicit conjunction internally, but this is refined (via backwards introduction) into separate sub-goals before the user commences the proof; the final result is projected into a list of theorems using eliminations (cf.\ \secref{sec:tactical-goals}). The \prop\ marker (\#\) makes arbitrarily complex propositions appear as atomic, without changing the meaning: \\ \ A\ and \\ \ #A\ are interchangeable. See \secref{sec:tactical-goals} for specific operations. The \term\ marker turns any well-typed term into a derivable proposition: \\ TERM t\ holds unconditionally. Although this is logically vacuous, it allows to treat terms and proofs uniformly, similar to a type-theoretic framework. The \TYPE\ constructor is the canonical representative of the unspecified type \\ itself\; it essentially injects the language of types into that of terms. There is specific notation \TYPE(\)\ for \TYPE\<^bsub>\ itself\<^esub>\. Although being devoid of any particular meaning, the term \TYPE(\)\ accounts for the type \\\ within the term language. In particular, \TYPE(\)\ may be used as formal argument in primitive definitions, in order to circumvent hidden polymorphism (cf.\ \secref{sec:terms}). For example, \c TYPE(\) \ A[\]\ defines \c :: \ itself \ prop\ in terms of a proposition \A\ that depends on an additional type argument, which is essentially a predicate on types. \ text %mlref \ \begin{mldecls} @{define_ML Conjunction.intr: "thm -> thm -> thm"} \\ @{define_ML Conjunction.elim: "thm -> thm * thm"} \\ @{define_ML Drule.mk_term: "cterm -> thm"} \\ @{define_ML Drule.dest_term: "thm -> cterm"} \\ @{define_ML Logic.mk_type: "typ -> term"} \\ @{define_ML Logic.dest_type: "term -> typ"} \\ \end{mldecls} \<^descr> \<^ML>\Conjunction.intr\ derives \A &&& B\ from \A\ and \B\. \<^descr> \<^ML>\Conjunction.elim\ derives \A\ and \B\ from \A &&& B\. \<^descr> \<^ML>\Drule.mk_term\ derives \TERM t\. \<^descr> \<^ML>\Drule.dest_term\ recovers term \t\ from \TERM t\. \<^descr> \<^ML>\Logic.mk_type\~\\\ produces the term \TYPE(\)\. \<^descr> \<^ML>\Logic.dest_type\~\TYPE(\)\ recovers the type \\\. \ subsection \Sort hypotheses\ text \ Type variables are decorated with sorts, as explained in \secref{sec:types}. This constrains type instantiation to certain ranges of types: variable \\\<^sub>s\ may only be assigned to types \\\ that belong to sort \s\. Within the logic, sort constraints act like implicit preconditions on the result \\\\<^sub>1 : s\<^sub>1\, \, \\\<^sub>n : s\<^sub>n\, \ \ \\ where the type variables \\\<^sub>1, \, \\<^sub>n\ cover the propositions \\\, \\\, as well as the proof of \\ \ \\. These \<^emph>\sort hypothesis\ of a theorem are passed monotonically through further derivations. They are redundant, as long as the statement of a theorem still contains the type variables that are accounted here. The logical significance of sort hypotheses is limited to the boundary case where type variables disappear from the proposition, e.g.\ \\\\<^sub>s : s\ \ \\. Since such dangling type variables can be renamed arbitrarily without changing the proposition \\\, the inference kernel maintains sort hypotheses in anonymous form \s \ \\. In most practical situations, such extra sort hypotheses may be stripped in a final bookkeeping step, e.g.\ at the end of a proof: they are typically left over from intermediate reasoning with type classes that can be satisfied by some concrete type \\\ of sort \s\ to replace the hypothetical type variable \\\<^sub>s\. \ text %mlref \ \begin{mldecls} @{define_ML Thm.extra_shyps: "thm -> sort list"} \\ @{define_ML Thm.strip_shyps: "thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Thm.extra_shyps\~\thm\ determines the extraneous sort hypotheses of the given theorem, i.e.\ the sorts that are not present within type variables of the statement. \<^descr> \<^ML>\Thm.strip_shyps\~\thm\ removes any extraneous sort hypotheses that can be witnessed from the type signature. \ text %mlex \ The following artificial example demonstrates the derivation of \<^prop>\False\ with a pending sort hypothesis involving a logically empty sort. \ class empty = assumes bad: "\(x::'a) y. x \ y" theorem (in empty) false: False using bad by blast ML_val \\<^assert> (Thm.extra_shyps @{thm false} = [\<^sort>\empty\])\ text \ Thanks to the inference kernel managing sort hypothesis according to their logical significance, this example is merely an instance of \<^emph>\ex falso quodlibet consequitur\ --- not a collapse of the logical framework! \ section \Object-level rules \label{sec:obj-rules}\ text \ The primitive inferences covered so far mostly serve foundational purposes. User-level reasoning usually works via object-level rules that are represented as theorems of Pure. Composition of rules involves \<^emph>\backchaining\, \<^emph>\higher-order unification\ modulo \\\\\-conversion of \\\-terms, and so-called \<^emph>\lifting\ of rules into a context of \\\ and \\\ connectives. Thus the full power of higher-order Natural Deduction in Isabelle/Pure becomes readily available. \ subsection \Hereditary Harrop Formulae\ text \ The idea of object-level rules is to model Natural Deduction inferences in - the style of Gentzen @{cite "Gentzen:1935"}, but we allow arbitrary nesting - similar to @{cite "Schroeder-Heister:1984"}. The most basic rule format is + the style of Gentzen \<^cite>\"Gentzen:1935"\, but we allow arbitrary nesting + similar to \<^cite>\"Schroeder-Heister:1984"\. The most basic rule format is that of a \<^emph>\Horn Clause\: \[ \infer{\A\}{\A\<^sub>1\ & \\\ & \A\<^sub>n\} \] where \A, A\<^sub>1, \, A\<^sub>n\ are atomic propositions of the framework, usually of the form \Trueprop B\, where \B\ is a (compound) object-level statement. This object-level inference corresponds to an iterated implication in Pure like this: \[ \A\<^sub>1 \ \ A\<^sub>n \ A\ \] As an example consider conjunction introduction: \A \ B \ A \ B\. Any parameters occurring in such rule statements are conceptionally treated as arbitrary: \[ \\x\<^sub>1 \ x\<^sub>m. A\<^sub>1 x\<^sub>1 \ x\<^sub>m \ \ A\<^sub>n x\<^sub>1 \ x\<^sub>m \ A x\<^sub>1 \ x\<^sub>m\ \] Nesting of rules means that the positions of \A\<^sub>i\ may again hold compound rules, not just atomic propositions. Propositions of this format are called - \<^emph>\Hereditary Harrop Formulae\ in the literature @{cite "Miller:1991"}. Here + \<^emph>\Hereditary Harrop Formulae\ in the literature \<^cite>\"Miller:1991"\. Here we give an inductive characterization as follows: \<^medskip> \begin{tabular}{ll} \\<^bold>x\ & set of variables \\ \\<^bold>A\ & set of atomic propositions \\ \\<^bold>H = \\<^bold>x\<^sup>*. \<^bold>H\<^sup>* \ \<^bold>A\ & set of Hereditary Harrop Formulas \\ \end{tabular} \<^medskip> Thus we essentially impose nesting levels on propositions formed from \\\ and \\\. At each level there is a prefix of parameters and compound premises, concluding an atomic proposition. Typical examples are \\\-introduction \(A \ B) \ A \ B\ or mathematical induction \P 0 \ (\n. P n \ P (Suc n)) \ P n\. Even deeper nesting occurs in well-founded induction \(\x. (\y. y \ x \ P y) \ P x) \ P x\, but this already marks the limit of rule complexity that is usually seen in practice. \<^medskip> Regular user-level inferences in Isabelle/Pure always maintain the following canonical form of results: \<^item> Normalization by \(A \ (\x. B x)) \ (\x. A \ B x)\, which is a theorem of Pure, means that quantifiers are pushed in front of implication at each level of nesting. The normal form is a Hereditary Harrop Formula. \<^item> The outermost prefix of parameters is represented via schematic variables: instead of \\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x\ we have \\<^vec>H ?\<^vec>x \ A ?\<^vec>x\. Note that this representation looses information about the order of parameters, and vacuous quantifiers vanish automatically. \ text %mlref \ \begin{mldecls} @{define_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Simplifier.norm_hhf\~\ctxt thm\ normalizes the given theorem according to the canonical form specified above. This is occasionally helpful to repair some low-level tools that do not handle Hereditary Harrop Formulae properly. \ subsection \Rule composition\ text \ The rule calculus of Isabelle/Pure provides two main inferences: @{inference resolution} (i.e.\ back-chaining of rules) and @{inference assumption} (i.e.\ closing a branch), both modulo higher-order unification. There are also combined variants, notably @{inference elim_resolution} and @{inference dest_resolution}. To understand the all-important @{inference resolution} principle, we first consider raw @{inference_def composition} (modulo higher-order unification with substitution \\\): \[ \infer[(@{inference_def composition})]{\\<^vec>A\ \ C\\} {\\<^vec>A \ B\ & \B' \ C\ & \B\ = B'\\} \] Here the conclusion of the first rule is unified with the premise of the second; the resulting rule instance inherits the premises of the first and conclusion of the second. Note that \C\ can again consist of iterated implications. We can also permute the premises of the second rule back-and-forth in order to compose with \B'\ in any position (subsequently we shall always refer to position 1 w.l.o.g.). In @{inference composition} the internal structure of the common part \B\ and \B'\ is not taken into account. For proper @{inference resolution} we require \B\ to be atomic, and explicitly observe the structure \\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x\ of the premise of the second rule. The idea is to adapt the first rule by ``lifting'' it into this context, by means of iterated application of the following inferences: \[ \infer[(@{inference_def imp_lift})]{\(\<^vec>H \ \<^vec>A) \ (\<^vec>H \ B)\}{\\<^vec>A \ B\} \] \[ \infer[(@{inference_def all_lift})]{\(\\<^vec>x. \<^vec>A (?\<^vec>a \<^vec>x)) \ (\\<^vec>x. B (?\<^vec>a \<^vec>x))\}{\\<^vec>A ?\<^vec>a \ B ?\<^vec>a\} \] By combining raw composition with lifting, we get full @{inference resolution} as follows: \[ \infer[(@{inference_def resolution})] {\(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (?\<^vec>a \<^vec>x))\ \ C\\} {\begin{tabular}{l} \\<^vec>A ?\<^vec>a \ B ?\<^vec>a\ \\ \(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C\ \\ \(\\<^vec>x. B (?\<^vec>a \<^vec>x))\ = B'\\ \\ \end{tabular}} \] Continued resolution of rules allows to back-chain a problem towards more and sub-problems. Branches are closed either by resolving with a rule of 0 premises, or by producing a ``short-circuit'' within a solved situation (again modulo unification): \[ \infer[(@{inference_def assumption})]{\C\\} {\(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C\ & \A\ = H\<^sub>i\\~~\mbox{(for some~\i\)}} \] %FIXME @{inference_def elim_resolution}, @{inference_def dest_resolution} \ text %mlref \ \begin{mldecls} @{define_ML_infix "RSN": "thm * (int * thm) -> thm"} \\ @{define_ML_infix "RS": "thm * thm -> thm"} \\ @{define_ML_infix "RLN": "thm list * (int * thm list) -> thm list"} \\ @{define_ML_infix "RL": "thm list * thm list -> thm list"} \\ @{define_ML_infix "MRS": "thm list * thm -> thm"} \\ @{define_ML_infix "OF": "thm * thm list -> thm"} \\ \end{mldecls} \<^descr> \rule\<^sub>1 RSN (i, rule\<^sub>2)\ resolves the conclusion of \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, according to the @{inference resolution} principle explained above. Unless there is precisely one resolvent it raises exception \<^ML>\THM\. This corresponds to the rule attribute @{attribute THEN} in Isar source language. \<^descr> \rule\<^sub>1 RS rule\<^sub>2\ abbreviates \rule\<^sub>1 RSN (1, rule\<^sub>2)\. \<^descr> \rules\<^sub>1 RLN (i, rules\<^sub>2)\ joins lists of rules. For every \rule\<^sub>1\ in \rules\<^sub>1\ and \rule\<^sub>2\ in \rules\<^sub>2\, it resolves the conclusion of \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, accumulating multiple results in one big list. Note that such strict enumerations of higher-order unifications can be inefficient compared to the lazy variant seen in elementary tactics like \<^ML>\resolve_tac\. \<^descr> \rules\<^sub>1 RL rules\<^sub>2\ abbreviates \rules\<^sub>1 RLN (1, rules\<^sub>2)\. \<^descr> \[rule\<^sub>1, \, rule\<^sub>n] MRS rule\ resolves \rule\<^sub>i\ against premise \i\ of \rule\, for \i = n, \, 1\. By working from right to left, newly emerging premises are concatenated in the result, without interfering. \<^descr> \rule OF rules\ is an alternative notation for \rules MRS rule\, which makes rule composition look more like function application. Note that the argument \rules\ need not be atomic. This corresponds to the rule attribute @{attribute OF} in Isar source language. \ section \Proof terms \label{sec:proof-terms}\ text \ The Isabelle/Pure inference kernel can record the proof of each theorem as a proof term that contains all logical inferences in detail. Rule composition by resolution (\secref{sec:obj-rules}) and type-class reasoning is broken down to primitive rules of the logical framework. The proof term can be inspected by a separate proof-checker, for example. According to the well-known \<^emph>\Curry-Howard isomorphism\, a proof can be viewed as a \\\-term. Following this idea, proofs in Isabelle are internally represented by a datatype similar to the one for terms described in \secref{sec:terms}. On top of these syntactic terms, two more layers of \\\-calculus are added, which correspond to \\x :: \. B x\ and \A \ B\ according to the propositions-as-types principle. The resulting 3-level \\\-calculus resembles ``\\HOL\'' in the more abstract setting of Pure Type - Systems (PTS) @{cite "Barendregt-Geuvers:2001"}, if some fine points like + Systems (PTS) \<^cite>\"Barendregt-Geuvers:2001"\, if some fine points like schematic polymorphism and type classes are ignored. \<^medskip> \<^emph>\Proof abstractions\ of the form \\<^bold>\x :: \. prf\ or \\<^bold>\p : A. prf\ correspond to introduction of \\\/\\\, and \<^emph>\proof applications\ of the form \p \ t\ or \p \ q\ correspond to elimination of \\\/\\\. Actual types \\\, propositions \A\, and terms \t\ might be suppressed and reconstructed from the overall proof term. \<^medskip> Various atomic proofs indicate special situations within the proof construction as follows. A \<^emph>\bound proof variable\ is a natural number \b\ that acts as de-Bruijn index for proof term abstractions. A \<^emph>\minimal proof\ ``\?\'' is a dummy proof term. This indicates some unrecorded part of the proof. \Hyp A\ refers to some pending hypothesis by giving its proposition. This indicates an open context of implicit hypotheses, similar to loose bound variables or free variables within a term (\secref{sec:terms}). An \<^emph>\axiom\ or \<^emph>\oracle\ \a : A[\<^vec>\]\ refers some postulated \proof constant\, which is subject to schematic polymorphism of theory content, and the particular type instantiation may be given explicitly. The vector of types \\<^vec>\\ refers to the schematic type variables in the generic proposition \A\ in canonical order. A \<^emph>\proof promise\ \a : A[\<^vec>\]\ is a placeholder for some proof of polymorphic proposition \A\, with explicit type instantiation as given by the vector \\<^vec>\\, as above. Unlike axioms or oracles, proof promises may be \<^emph>\fulfilled\ eventually, by substituting \a\ by some particular proof \q\ at the corresponding type instance. This acts like Hindley-Milner \let\-polymorphism: a generic local proof definition may get used at different type instances, and is replaced by the concrete instance eventually. A \<^emph>\named theorem\ wraps up some concrete proof as a closed formal entity, in the manner of constant definitions for proof terms. The \<^emph>\proof body\ of such boxed theorems involves some digest about oracles and promises occurring in the original proof. This allows the inference kernel to manage this critical information without the full overhead of explicit proof terms. \ subsection \Reconstructing and checking proof terms\ text \ Fully explicit proof terms can be large, but most of this information is redundant and can be reconstructed from the context. Therefore, the Isabelle/Pure inference kernel records only \<^emph>\implicit\ proof terms, by omitting all typing information in terms, all term and type labels of proof abstractions, and some argument terms of applications \p \ t\ (if possible). There are separate operations to reconstruct the full proof term later on, - using \<^emph>\higher-order pattern unification\ @{cite "nipkow-patterns" and - "Berghofer-Nipkow:2000:TPHOL"}. + using \<^emph>\higher-order pattern unification\ \<^cite>\"nipkow-patterns" and + "Berghofer-Nipkow:2000:TPHOL"\. The \<^emph>\proof checker\ expects a fully reconstructed proof term, and can turn it into a theorem by replaying its primitive inferences within the kernel. \ subsection \Concrete syntax of proof terms\ text \ The concrete syntax of proof terms is a slight extension of the regular - inner syntax of Isabelle/Pure @{cite "isabelle-isar-ref"}. Its main + inner syntax of Isabelle/Pure \<^cite>\"isabelle-isar-ref"\. Its main syntactic category @{syntax (inner) proof} is defined as follows: \begin{center} \begin{supertabular}{rclr} @{syntax_def (inner) proof} & = & \\<^bold>\\ \params\ \<^verbatim>\.\ \proof\ \\ & \|\ & \proof\ \\\ \any\ \\ & \|\ & \proof\ \\\ \proof\ \\ & \|\ & \id | longid\ \\ \\ \param\ & = & \idt\ \\ & \|\ & \idt\ \<^verbatim>\:\ \prop\ \\ & \|\ & \<^verbatim>\(\ \param\ \<^verbatim>\)\ \\ \\ \params\ & = & \param\ \\ & \|\ & \param\ \params\ \\ \end{supertabular} \end{center} Implicit term arguments in partial proofs are indicated by ``\_\''. Type arguments for theorems and axioms may be specified using \p \ TYPE(type)\ (they must appear before any other term argument of a theorem or axiom, but may be omitted altogether). \<^medskip> There are separate read and print operations for proof terms, in order to avoid conflicts with the regular term language. \ text %mlref \ \begin{mldecls} @{define_ML_type proof} \\ @{define_ML_type proof_body} \\ @{define_ML Proofterm.proofs: "int Unsynchronized.ref"} \\ @{define_ML Proofterm.reconstruct_proof: "theory -> term -> proof -> proof"} \\ @{define_ML Proofterm.expand_proof: "theory -> (Proofterm.thm_header -> string option) -> proof -> proof"} \\ @{define_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\ @{define_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\ @{define_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\proof\ represents proof terms; this is a datatype with constructors @{define_ML Abst}, @{define_ML AbsP}, @{define_ML_infix "%"}, @{define_ML_infix "%%"}, @{define_ML PBound}, @{define_ML MinProof}, @{define_ML Hyp}, @{define_ML PAxm}, @{define_ML Oracle}, @{define_ML PThm} as explained above. %FIXME PClass (!?) \<^descr> Type \<^ML_type>\proof_body\ represents the nested proof information of a named theorem, consisting of a digest of oracles and named theorem over some proof term. The digest only covers the directly visible part of the proof: in order to get the full information, the implicit graph of nested theorems needs to be traversed (e.g.\ using \<^ML>\Proofterm.fold_body_thms\). \<^descr> \<^ML>\Thm.proof_of\~\thm\ and \<^ML>\Thm.proof_body_of\~\thm\ produce the proof term or proof body (with digest of oracles and theorems) from a given theorem. Note that this involves a full join of internal futures that fulfill pending proof promises, and thus disrupts the natural bottom-up construction of proofs by introducing dynamic ad-hoc dependencies. Parallel performance may suffer by inspecting proof terms at run-time. \<^descr> \<^ML>\Proofterm.proofs\ specifies the detail of proof recording within \<^ML_type>\thm\ values produced by the inference kernel: \<^ML>\0\ records only the names of oracles, \<^ML>\1\ records oracle names and propositions, \<^ML>\2\ additionally records full proof terms. Officially named theorems that contribute to a result are recorded in any case. \<^descr> \<^ML>\Proofterm.reconstruct_proof\~\thy prop prf\ turns the implicit proof term \prf\ into a full proof of the given proposition. Reconstruction may fail if \prf\ is not a proof of \prop\, or if it does not contain sufficient information for reconstruction. Failure may only happen for proofs that are constructed manually, but not for those produced automatically by the inference kernel. \<^descr> \<^ML>\Proofterm.expand_proof\~\thy expand prf\ reconstructs and expands the proofs of nested theorems according to the given \expand\ function: a result of @{ML \SOME ""\} means full expansion, @{ML \SOME\}~\name\ means to keep the theorem node but replace its internal name, @{ML NONE} means no change. \<^descr> \<^ML>\Proof_Checker.thm_of_proof\~\thy prf\ turns the given (full) proof into a theorem, by replaying it using only primitive rules of the inference kernel. \<^descr> \<^ML>\Proof_Syntax.read_proof\~\thy b\<^sub>1 b\<^sub>2 s\ reads in a proof term. The Boolean flags indicate the use of sort and type information. Usually, typing information is left implicit and is inferred during proof reconstruction. %FIXME eliminate flags!? \<^descr> \<^ML>\Proof_Syntax.pretty_proof\~\ctxt prf\ pretty-prints the given proof term. \ text %mlex \ \<^item> \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\ provides basic examples involving proof terms. \<^item> \<^file>\~~/src/HOL/Proofs/ex/XML_Data.thy\ demonstrates export and import of proof terms via XML/ML data representation. \ end diff --git a/src/Doc/Implementation/ML.thy b/src/Doc/Implementation/ML.thy --- a/src/Doc/Implementation/ML.thy +++ b/src/Doc/Implementation/ML.thy @@ -1,2240 +1,2238 @@ (*:maxLineLen=78:*) theory "ML" imports Base begin chapter \Isabelle/ML\ text \ Isabelle/ML is best understood as a certain culture based on Standard ML. Thus it is not a new programming language, but a certain way to use SML at an advanced level within the Isabelle environment. This covers a variety of aspects that are geared towards an efficient and robust platform for applications of formal logic with fully foundational proof construction --- according to the well-known \<^emph>\LCF principle\. There is specific infrastructure with library modules to address the needs of this difficult task. For example, the raw parallel programming model of Poly/ML is presented as considerably more abstract concept of \<^emph>\futures\, which is then used to augment the inference kernel, Isar theory and proof interpreter, and PIDE document management. The main aspects of Isabelle/ML are introduced below. These first-hand explanations should help to understand how proper Isabelle/ML is to be read and written, and to get access to the wealth of experience that is expressed in the source text and its history of changes.\<^footnote>\See \<^url>\https://isabelle.in.tum.de/repos/isabelle\ for the full Mercurial history. There are symbolic tags to refer to official Isabelle releases, as opposed to arbitrary \<^emph>\tip\ versions that merely reflect snapshots that are never really up-to-date.\ \ section \Style and orthography\ text \ The sources of Isabelle/Isar are optimized for \<^emph>\readability\ and \<^emph>\maintainability\. The main purpose is to tell an informed reader what is really going on and how things really work. This is a non-trivial aim, but it is supported by a certain style of writing Isabelle/ML that has emerged from long years of system development.\<^footnote>\See also the interesting style guide for OCaml \<^url>\https://caml.inria.fr/resources/doc/guides/guidelines.en.html\ which shares many of our means and ends.\ The main principle behind any coding style is \<^emph>\consistency\. For a single author of a small program this merely means ``choose your style and stick to it''. A complex project like Isabelle, with long years of development and different contributors, requires more standardization. A coding style that is changed every few years or with every new contributor is no style at all, because consistency is quickly lost. Global consistency is hard to achieve, though. Nonetheless, one should always strive at least for local consistency of modules and sub-systems, without deviating from some general principles how to write Isabelle/ML. In a sense, good coding style is like an \<^emph>\orthography\ for the sources: it helps to read quickly over the text and see through the main points, without getting distracted by accidental presentation of free-style code. \ subsection \Header and sectioning\ text \ Isabelle source files have a certain standardized header format (with precise spacing) that follows ancient traditions reaching back to the earliest versions of the system by Larry Paulson. See \<^file>\~~/src/Pure/thm.ML\, for example. The header includes at least \<^verbatim>\Title\ and \<^verbatim>\Author\ entries, followed by a prose description of the purpose of the module. The latter can range from a single line to several paragraphs of explanations. The rest of the file is divided into chapters, sections, subsections, subsubsections, paragraphs etc.\ using a simple layout via ML comments as follows. @{verbatim [display] \ (**** chapter ****) (*** section ***) (** subsection **) (* subsubsection *) (*short paragraph*) (* long paragraph, with more text *)\} As in regular typography, there is some extra space \<^emph>\before\ section headings that are adjacent to plain text, but not other headings as in the example above. \<^medskip> The precise wording of the prose text given in these headings is chosen carefully to introduce the main theme of the subsequent formal ML text. \ subsection \Naming conventions\ text \ Since ML is the primary medium to express the meaning of the source text, naming of ML entities requires special care. \ paragraph \Notation.\ text \ A name consists of 1--3 \<^emph>\words\ (rarely 4, but not more) that are separated by underscore. There are three variants concerning upper or lower case letters, which are used for certain ML categories as follows: \<^medskip> \begin{tabular}{lll} variant & example & ML categories \\\hline lower-case & \<^ML_text>\foo_bar\ & values, types, record fields \\ capitalized & \<^ML_text>\Foo_Bar\ & datatype constructors, structures, functors \\ upper-case & \<^ML_text>\FOO_BAR\ & special values, exception constructors, signatures \\ \end{tabular} \<^medskip> For historical reasons, many capitalized names omit underscores, e.g.\ old-style \<^ML_text>\FooBar\ instead of \<^ML_text>\Foo_Bar\. Genuine mixed-case names are \<^emph>\not\ used, because clear division of words is essential for readability.\<^footnote>\Camel-case was invented to workaround the lack of underscore in some early non-ASCII character sets. Later it became habitual in some language communities that are now strong in numbers.\ A single (capital) character does not count as ``word'' in this respect: some Isabelle/ML names are suffixed by extra markers like this: \<^ML_text>\foo_barT\. Name variants are produced by adding 1--3 primes, e.g.\ \<^ML_text>\foo'\, \<^ML_text>\foo''\, or \<^ML_text>\foo'''\, but not \<^ML_text>\foo''''\ or more. Decimal digits scale better to larger numbers, e.g.\ \<^ML_text>\foo0\, \<^ML_text>\foo1\, \<^ML_text>\foo42\. \ paragraph \Scopes.\ text \ Apart from very basic library modules, ML structures are not ``opened'', but names are referenced with explicit qualification, as in \<^ML>\Syntax.string_of_term\ for example. When devising names for structures and their components it is important to aim at eye-catching compositions of both parts, because this is how they are seen in the sources and documentation. For the same reasons, aliases of well-known library functions should be avoided. Local names of function abstraction or case/let bindings are typically shorter, sometimes using only rudiments of ``words'', while still avoiding cryptic shorthands. An auxiliary function called \<^ML_text>\helper\, \<^ML_text>\aux\, or \<^ML_text>\f\ is considered bad style. Example: @{verbatim [display] \ (* RIGHT *) fun print_foo ctxt foo = let fun print t = ... Syntax.string_of_term ctxt t ... in ... end; (* RIGHT *) fun print_foo ctxt foo = let val string_of_term = Syntax.string_of_term ctxt; fun print t = ... string_of_term t ... in ... end; (* WRONG *) val string_of_term = Syntax.string_of_term; fun print_foo ctxt foo = let fun aux t = ... string_of_term ctxt t ... in ... end;\} \ paragraph \Specific conventions.\ text \ Here are some specific name forms that occur frequently in the sources. \<^item> A function that maps \<^ML_text>\foo\ to \<^ML_text>\bar\ is called \<^ML_text>\foo_to_bar\ or \<^ML_text>\bar_of_foo\ (never \<^ML_text>\foo2bar\, nor \<^ML_text>\bar_from_foo\, nor \<^ML_text>\bar_for_foo\, nor \<^ML_text>\bar4foo\). \<^item> The name component \<^ML_text>\legacy\ means that the operation is about to be discontinued soon. \<^item> The name component \<^ML_text>\global\ means that this works with the background theory instead of the regular local context (\secref{sec:context}), sometimes for historical reasons, sometimes due a genuine lack of locality of the concept involved, sometimes as a fall-back for the lack of a proper context in the application code. Whenever there is a non-global variant available, the application should be migrated to use it with a proper local context. \<^item> Variables of the main context types of the Isabelle/Isar framework (\secref{sec:context} and \chref{ch:local-theory}) have firm naming conventions as follows: \<^item> theories are called \<^ML_text>\thy\, rarely \<^ML_text>\theory\ (never \<^ML_text>\thry\) \<^item> proof contexts are called \<^ML_text>\ctxt\, rarely \<^ML_text>\context\ (never \<^ML_text>\ctx\) \<^item> generic contexts are called \<^ML_text>\context\ \<^item> local theories are called \<^ML_text>\lthy\, except for local theories that are treated as proof context (which is a semantic super-type) Variations with primed or decimal numbers are always possible, as well as semantic prefixes like \<^ML_text>\foo_thy\ or \<^ML_text>\bar_ctxt\, but the base conventions above need to be preserved. This allows to emphasize their data flow via plain regular expressions in the text editor. \<^item> The main logical entities (\secref{ch:logic}) have established naming convention as follows: \<^item> sorts are called \<^ML_text>\S\ \<^item> types are called \<^ML_text>\T\, \<^ML_text>\U\, or \<^ML_text>\ty\ (never \<^ML_text>\t\) \<^item> terms are called \<^ML_text>\t\, \<^ML_text>\u\, or \<^ML_text>\tm\ (never \<^ML_text>\trm\) \<^item> certified types are called \<^ML_text>\cT\, rarely \<^ML_text>\T\, with variants as for types \<^item> certified terms are called \<^ML_text>\ct\, rarely \<^ML_text>\t\, with variants as for terms (never \<^ML_text>\ctrm\) \<^item> theorems are called \<^ML_text>\th\, or \<^ML_text>\thm\ Proper semantic names override these conventions completely. For example, the left-hand side of an equation (as a term) can be called \<^ML_text>\lhs\ (not \<^ML_text>\lhs_tm\). Or a term that is known to be a variable can be called \<^ML_text>\v\ or \<^ML_text>\x\. \<^item> Tactics (\secref{sec:tactics}) are sufficiently important to have specific naming conventions. The name of a basic tactic definition always has a \<^ML_text>\_tac\ suffix, the subgoal index (if applicable) is always called \<^ML_text>\i\, and the goal state (if made explicit) is usually called \<^ML_text>\st\ instead of the somewhat misleading \<^ML_text>\thm\. Any other arguments are given before the latter two, and the general context is given first. Example: @{verbatim [display] \ fun my_tac ctxt arg1 arg2 i st = ...\} Note that the goal state \<^ML_text>\st\ above is rarely made explicit, if tactic combinators (tacticals) are used as usual. A tactic that requires a proof context needs to make that explicit as seen in the \<^verbatim>\ctxt\ argument above. Do not refer to the background theory of \<^verbatim>\st\ -- it is not a proper context, but merely a formal certificate. \ subsection \General source layout\ text \ The general Isabelle/ML source layout imitates regular type-setting conventions, augmented by the requirements for deeply nested expressions that are commonplace in functional programming. \ paragraph \Line length\ text \ is limited to 80 characters according to ancient standards, but we allow as much as 100 characters (not more).\<^footnote>\Readability requires to keep the beginning of a line in view while watching its end. Modern wide-screen displays do not change the way how the human brain works. Sources also need to be printable on plain paper with reasonable font-size.\ The extra 20 characters acknowledge the space requirements due to qualified library references in Isabelle/ML. \ paragraph \White-space\ text \ is used to emphasize the structure of expressions, following mostly standard conventions for mathematical typesetting, as can be seen in plain {\TeX} or {\LaTeX}. This defines positioning of spaces for parentheses, punctuation, and infixes as illustrated here: @{verbatim [display] \ val x = y + z * (a + b); val pair = (a, b); val record = {foo = 1, bar = 2};\} Lines are normally broken \<^emph>\after\ an infix operator or punctuation character. For example: @{verbatim [display] \ val x = a + b + c; val tuple = (a, b, c); \} Some special infixes (e.g.\ \<^ML_text>\|>\) work better at the start of the line, but punctuation is always at the end. Function application follows the tradition of \\\-calculus, not informal mathematics. For example: \<^ML_text>\f a b\ for a curried function, or \<^ML_text>\g (a, b)\ for a tupled function. Note that the space between \<^ML_text>\g\ and the pair \<^ML_text>\(a, b)\ follows the important principle of \<^emph>\compositionality\: the layout of \<^ML_text>\g p\ does not change when \<^ML_text>\p\ is refined to the concrete pair \<^ML_text>\(a, b)\. \ paragraph \Indentation\ text \ uses plain spaces, never hard tabulators.\<^footnote>\Tabulators were invented to move the carriage of a type-writer to certain predefined positions. In software they could be used as a primitive run-length compression of consecutive spaces, but the precise result would depend on non-standardized text editor configuration.\ Each level of nesting is indented by 2 spaces, sometimes 1, very rarely 4, never 8 or any other odd number. Indentation follows a simple logical format that only depends on the nesting depth, not the accidental length of the text that initiates a level of nesting. Example: @{verbatim [display] \ (* RIGHT *) if b then expr1_part1 expr1_part2 else expr2_part1 expr2_part2 (* WRONG *) if b then expr1_part1 expr1_part2 else expr2_part1 expr2_part2\} The second form has many problems: it assumes a fixed-width font when viewing the sources, it uses more space on the line and thus makes it hard to observe its strict length limit (working against \<^emph>\readability\), it requires extra editing to adapt the layout to changes of the initial text (working against \<^emph>\maintainability\) etc. \<^medskip> For similar reasons, any kind of two-dimensional or tabular layouts, ASCII-art with lines or boxes of asterisks etc.\ should be avoided. \ paragraph \Complex expressions\ text \ that consist of multi-clausal function definitions, \<^ML_text>\handle\, \<^ML_text>\case\, \<^ML_text>\let\ (and combinations) require special attention. The syntax of Standard ML is quite ambitious and admits a lot of variance that can distort the meaning of the text. Multiple clauses of \<^ML_text>\fun\, \<^ML_text>\fn\, \<^ML_text>\handle\, \<^ML_text>\case\ get extra indentation to indicate the nesting clearly. Example: @{verbatim [display] \ (* RIGHT *) fun foo p1 = expr1 | foo p2 = expr2 (* WRONG *) fun foo p1 = expr1 | foo p2 = expr2\} Body expressions consisting of \<^ML_text>\case\ or \<^ML_text>\let\ require care to maintain compositionality, to prevent loss of logical indentation where it is especially important to see the structure of the text. Example: @{verbatim [display] \ (* RIGHT *) fun foo p1 = (case e of q1 => ... | q2 => ...) | foo p2 = let ... in ... end (* WRONG *) fun foo p1 = case e of q1 => ... | q2 => ... | foo p2 = let ... in ... end\} Extra parentheses around \<^ML_text>\case\ expressions are optional, but help to analyse the nesting based on character matching in the text editor. \<^medskip> There are two main exceptions to the overall principle of compositionality in the layout of complex expressions. \<^enum> \<^ML_text>\if\ expressions are iterated as if ML had multi-branch conditionals, e.g. @{verbatim [display] \ (* RIGHT *) if b1 then e1 else if b2 then e2 else e3\} \<^enum> \<^ML_text>\fn\ abstractions are often layed-out as if they would lack any structure by themselves. This traditional form is motivated by the possibility to shift function arguments back and forth wrt.\ additional combinators. Example: @{verbatim [display] \ (* RIGHT *) fun foo x y = fold (fn z => expr)\} Here the visual appearance is that of three arguments \<^ML_text>\x\, \<^ML_text>\y\, \<^ML_text>\z\ in a row. Such weakly structured layout should be use with great care. Here are some counter-examples involving \<^ML_text>\let\ expressions: @{verbatim [display] \ (* WRONG *) fun foo x = let val y = ... in ... end (* WRONG *) fun foo x = let val y = ... in ... end (* WRONG *) fun foo x = let val y = ... in ... end (* WRONG *) fun foo x = let val y = ... in ... end\} \<^medskip> In general the source layout is meant to emphasize the structure of complex language expressions, not to pretend that SML had a completely different syntax (say that of Haskell, Scala, Java). \ section \ML embedded into Isabelle/Isar\ text \ ML and Isar are intertwined via an open-ended bootstrap process that provides more and more programming facilities and logical content in an alternating manner. Bootstrapping starts from the raw environment of existing implementations of Standard ML (mainly Poly/ML). Isabelle/Pure marks the point where the raw ML toplevel is superseded by Isabelle/ML within the Isar theory and proof language, with a uniform context for arbitrary ML values (see also \secref{sec:context}). This formal environment holds ML compiler bindings, logical entities, and many other things. Object-logics like Isabelle/HOL are built within the Isabelle/ML/Isar environment by introducing suitable theories with associated ML modules, either inlined within \<^verbatim>\.thy\ files, or as separate \<^verbatim>\.ML\ files that are loading from some theory. Thus Isabelle/HOL is defined as a regular user-space application within the Isabelle framework. Further add-on tools can be implemented in ML within the Isar context in the same manner: ML is part of the standard repertoire of Isabelle, and there is no distinction between ``users'' and ``developers'' in this respect. \ subsection \Isar ML commands\ text \ The primary Isar source language provides facilities to ``open a window'' to the underlying ML compiler. Especially see the Isar commands @{command_ref "ML_file"} and @{command_ref "ML"}: both work the same way, but the source text is provided differently, via a file vs.\ inlined, respectively. Apart from embedding ML into the main theory definition like that, there are many more commands that refer to ML source, such as @{command_ref setup} or @{command_ref declaration}. Even more fine-grained embedding of ML into Isar is encountered in the proof method @{method_ref tactic}, which refines the pending goal state via a given expression of type \<^ML_type>\tactic\. \ text %mlex \ The following artificial example demonstrates some ML toplevel declarations within the implicit Isar theory context. This is regular functional programming without referring to logical entities yet. \ ML \ fun factorial 0 = 1 | factorial n = n * factorial (n - 1) \ text \ Here the ML environment is already managed by Isabelle, i.e.\ the \<^ML>\factorial\ function is not yet accessible in the preceding paragraph, nor in a different theory that is independent from the current one in the import hierarchy. Removing the above ML declaration from the source text will remove any trace of this definition, as expected. The Isabelle/ML toplevel environment is managed in a \<^emph>\stateless\ way: in contrast to the raw ML toplevel, there are no global side-effects involved here.\<^footnote>\Such a stateless compilation environment is also a prerequisite for robust parallel compilation within independent nodes of the implicit theory development graph.\ \<^medskip> The next example shows how to embed ML into Isar proofs, using @{command_ref "ML_prf"} instead of @{command_ref "ML"}. As illustrated below, the effect on the ML environment is local to the whole proof body, but ignoring the block structure. \ notepad begin ML_prf %"ML" \val a = 1\ { ML_prf %"ML" \val b = a + 1\ } \ \Isar block structure ignored by ML environment\ ML_prf %"ML" \val c = b + 1\ end text \ By side-stepping the normal scoping rules for Isar proof blocks, embedded ML code can refer to the different contexts and manipulate corresponding entities, e.g.\ export a fact from a block context. \<^medskip> Two further ML commands are useful in certain situations: @{command_ref ML_val} and @{command_ref ML_command} are \<^emph>\diagnostic\ in the sense that there is no effect on the underlying environment, and can thus be used anywhere. The examples below produce long strings of digits by invoking \<^ML>\factorial\: @{command ML_val} takes care of printing the ML toplevel result, but @{command ML_command} is silent so we produce an explicit output message. \ ML_val \factorial 100\ ML_command \writeln (string_of_int (factorial 100))\ notepad begin ML_val \factorial 100\ ML_command \writeln (string_of_int (factorial 100))\ end subsection \Compile-time context\ text \ Whenever the ML compiler is invoked within Isabelle/Isar, the formal context is passed as a thread-local reference variable. Thus ML code may access the theory context during compilation, by reading or writing the (local) theory under construction. Note that such direct access to the compile-time context is rare. In practice it is typically done via some derived ML functions instead. \ text %mlref \ \begin{mldecls} @{define_ML Context.the_generic_context: "unit -> Context.generic"} \\ @{define_ML "Context.>>": "(Context.generic -> Context.generic) -> unit"} \\ @{define_ML ML_Thms.bind_thms: "string * thm list -> unit"} \\ @{define_ML ML_Thms.bind_thm: "string * thm -> unit"} \\ \end{mldecls} \<^descr> \<^ML>\Context.the_generic_context ()\ refers to the theory context of the ML toplevel --- at compile time. ML code needs to take care to refer to \<^ML>\Context.the_generic_context ()\ correctly. Recall that evaluation of a function body is delayed until actual run-time. \<^descr> \<^ML>\Context.>>\~\f\ applies context transformation \f\ to the implicit context of the ML toplevel. \<^descr> \<^ML>\ML_Thms.bind_thms\~\(name, thms)\ stores a list of theorems produced in ML both in the (global) theory context and the ML toplevel, associating it with the provided name. \<^descr> \<^ML>\ML_Thms.bind_thm\ is similar to \<^ML>\ML_Thms.bind_thms\ but refers to a singleton fact. It is important to note that the above functions are really restricted to the compile time, even though the ML compiler is invoked at run-time. The majority of ML code either uses static antiquotations (\secref{sec:ML-antiq}) or refers to the theory or proof context at run-time, by explicit functional abstraction. \ subsection \Antiquotations \label{sec:ML-antiq}\ text \ A very important consequence of embedding ML into Isar is the concept of \<^emph>\ML antiquotation\. The standard token language of ML is augmented by special syntactic entities of the following form: \<^rail>\ @{syntax_def antiquote}: '@{' name args '}' \ Here @{syntax name} and @{syntax args} are outer syntax categories, as - defined in @{cite "isabelle-isar-ref"}. + defined in \<^cite>\"isabelle-isar-ref"\. \<^medskip> A regular antiquotation \@{name args}\ processes its arguments by the usual means of the Isar source language, and produces corresponding ML source text, either as literal \<^emph>\inline\ text (e.g.\ \@{term t}\) or abstract \<^emph>\value\ (e.g. \@{thm th}\). This pre-compilation scheme allows to refer to formal entities in a robust manner, with proper static scoping and with some degree of logical checking of small portions of the code. \ subsection \Printing ML values\ text \ The ML compiler knows about the structure of values according to their static type, and can print them in the manner of its toplevel, although the details are non-portable. The antiquotations @{ML_antiquotation_def "make_string"} and @{ML_antiquotation_def "print"} provide a quasi-portable way to refer to this potential capability of the underlying ML system in generic Isabelle/ML sources. This is occasionally useful for diagnostic or demonstration purposes. Note that production-quality tools require proper user-level error messages, avoiding raw ML values in the output. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "make_string"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "print"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation make_string} ; @@{ML_antiquotation print} embedded? \ \<^descr> \@{make_string}\ inlines a function to print arbitrary values similar to the ML toplevel. The result is compiler dependent and may fall back on "?" in certain situations. The value of configuration option @{attribute_ref ML_print_depth} determines further details of output. \<^descr> \@{print f}\ uses the ML function \f: string -> unit\ to output the result of \@{make_string}\ above, together with the source position of the antiquotation. The default output function is \<^ML>\writeln\. \ text %mlex \ The following artificial examples show how to produce adhoc output of ML values for debugging purposes. \ ML_val \ val x = 42; val y = true; writeln (\<^make_string> {x = x, y = y}); \<^print> {x = x, y = y}; \<^print>\tracing\ {x = x, y = y}; \ section \Canonical argument order \label{sec:canonical-argument-order}\ text \ Standard ML is a language in the tradition of \\\-calculus and \<^emph>\higher-order functional programming\, similar to OCaml, Haskell, or Isabelle/Pure and HOL as logical languages. Getting acquainted with the native style of representing functions in that setting can save a lot of extra boiler-plate of redundant shuffling of arguments, auxiliary abstractions etc. Functions are usually \<^emph>\curried\: the idea of turning arguments of type \\\<^sub>i\ (for \i \ {1, \ n}\) into a result of type \\\ is represented by the iterated function space \\\<^sub>1 \ \ \ \\<^sub>n \ \\. This is isomorphic to the well-known encoding via tuples \\\<^sub>1 \ \ \ \\<^sub>n \ \\, but the curried version fits more smoothly into the basic calculus.\<^footnote>\The difference is even more significant in HOL, because the redundant tuple structure needs to be accommodated extraneous proof steps.\ Currying gives some flexibility due to \<^emph>\partial application\. A function \f: \\<^sub>1 \ \\<^sub>2 \ \\ can be applied to \x: \\<^sub>1\ and the remaining \(f x): \\<^sub>2 \ \\ passed to another function etc. How well this works in practice depends on the order of arguments. In the worst case, arguments are arranged erratically, and using a function in a certain situation always requires some glue code. Thus we would get exponentially many opportunities to decorate the code with meaningless permutations of arguments. This can be avoided by \<^emph>\canonical argument order\, which observes certain standard patterns and minimizes adhoc permutations in their application. In Isabelle/ML, large portions of text can be written without auxiliary operations like \swap: \ \ \ \ \ \ \\ or \C: (\ \ \ \ \) \ (\ \ \ \ \)\ (the latter is not present in the Isabelle/ML library). \<^medskip> The main idea is that arguments that vary less are moved further to the left than those that vary more. Two particularly important categories of functions are \<^emph>\selectors\ and \<^emph>\updates\. The subsequent scheme is based on a hypothetical set-like container of type \\\ that manages elements of type \\\. Both the names and types of the associated operations are canonical for Isabelle/ML. \begin{center} \begin{tabular}{ll} kind & canonical name and type \\\hline selector & \member: \ \ \ \ bool\ \\ update & \insert: \ \ \ \ \\ \\ \end{tabular} \end{center} Given a container \B: \\, the partially applied \member B\ is a predicate over elements \\ \ bool\, and thus represents the intended denotation directly. It is customary to pass the abstract predicate to further operations, not the concrete container. The argument order makes it easy to use other combinators: \forall (member B) list\ will check a list of elements for membership in \B\ etc. Often the explicit \list\ is pointless and can be contracted to \forall (member B)\ to get directly a predicate again. In contrast, an update operation varies the container, so it moves to the right: \insert a\ is a function \\ \ \\ to insert a value \a\. These can be composed naturally as \insert c \ insert b \ insert a\. The slightly awkward inversion of the composition order is due to conventional mathematical notation, which can be easily amended as explained below. \ subsection \Forward application and composition\ text \ Regular function application and infix notation works best for relatively deeply structured expressions, e.g.\ \h (f x y + g z)\. The important special case of \<^emph>\linear transformation\ applies a cascade of functions \f\<^sub>n (\ (f\<^sub>1 x))\. This becomes hard to read and maintain if the functions are themselves given as complex expressions. The notation can be significantly improved by introducing \<^emph>\forward\ versions of application and composition as follows: \<^medskip> \begin{tabular}{lll} \x |> f\ & \\\ & \f x\ \\ \(f #> g) x\ & \\\ & \x |> f |> g\ \\ \end{tabular} \<^medskip> This enables to write conveniently \x |> f\<^sub>1 |> \ |> f\<^sub>n\ or \f\<^sub>1 #> \ #> f\<^sub>n\ for its functional abstraction over \x\. \<^medskip> There is an additional set of combinators to accommodate multiple results (via pairs) that are passed on as multiple arguments (via currying). \<^medskip> \begin{tabular}{lll} \(x, y) |-> f\ & \\\ & \f x y\ \\ \(f #-> g) x\ & \\\ & \x |> f |-> g\ \\ \end{tabular} \<^medskip> \ text %mlref \ \begin{mldecls} @{define_ML_infix "|>" : "'a * ('a -> 'b) -> 'b"} \\ @{define_ML_infix "|->" : "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\ @{define_ML_infix "#>" : "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\ @{define_ML_infix "#->" : "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\ \end{mldecls} \ subsection \Canonical iteration\ text \ As explained above, a function \f: \ \ \ \ \\ can be understood as update on a configuration of type \\\, parameterized by an argument of type \\\. Given \a: \\ the partial application \(f a): \ \ \\ operates homogeneously on \\\. This can be iterated naturally over a list of parameters \[a\<^sub>1, \, a\<^sub>n]\ as \f a\<^sub>1 #> \ #> f a\<^sub>n\. The latter expression is again a function \\ \ \\. It can be applied to an initial configuration \b: \\ to start the iteration over the given list of arguments: each \a\ in \a\<^sub>1, \, a\<^sub>n\ is applied consecutively by updating a cumulative configuration. The \fold\ combinator in Isabelle/ML lifts a function \f\ as above to its iterated version over a list of arguments. Lifting can be repeated, e.g.\ \(fold \ fold) f\ iterates over a list of lists as expected. The variant \fold_rev\ works inside-out over the list of arguments, such that \fold_rev f \ fold f \ rev\ holds. The \fold_map\ combinator essentially performs \fold\ and \map\ simultaneously: each application of \f\ produces an updated configuration together with a side-result; the iteration collects all such side-results as a separate list. \ text %mlref \ \begin{mldecls} @{define_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\ @{define_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\ @{define_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\ \end{mldecls} \<^descr> \<^ML>\fold\~\f\ lifts the parametrized update function \f\ to a list of parameters. \<^descr> \<^ML>\fold_rev\~\f\ is similar to \<^ML>\fold\~\f\, but works inside-out, as if the list would be reversed. \<^descr> \<^ML>\fold_map\~\f\ lifts the parametrized update function \f\ (with side-result) to a list of parameters and cumulative side-results. \begin{warn} The literature on functional programming provides a confusing multitude of combinators called \foldl\, \foldr\ etc. SML97 provides its own variations as \<^ML>\List.foldl\ and \<^ML>\List.foldr\, while the classic Isabelle library also has the historic \<^ML>\Library.foldl\ and \<^ML>\Library.foldr\. To avoid unnecessary complication, all these historical versions should be ignored, and the canonical \<^ML>\fold\ (or \<^ML>\fold_rev\) used exclusively. \end{warn} \ text %mlex \ The following example shows how to fill a text buffer incrementally by adding strings, either individually or from a given list. \ ML_val \ val s = Buffer.empty |> Buffer.add "digits: " |> fold (Buffer.add o string_of_int) (0 upto 9) |> Buffer.content; \<^assert> (s = "digits: 0123456789"); \ text \ Note how \<^ML>\fold (Buffer.add o string_of_int)\ above saves an extra \<^ML>\map\ over the given list. This kind of peephole optimization reduces both the code size and the tree structures in memory (``deforestation''), but it requires some practice to read and write fluently. \<^medskip> The next example elaborates the idea of canonical iteration, demonstrating fast accumulation of tree content using a text buffer. \ ML \ datatype tree = Text of string | Elem of string * tree list; fun slow_content (Text txt) = txt | slow_content (Elem (name, ts)) = "<" ^ name ^ ">" ^ implode (map slow_content ts) ^ "" fun add_content (Text txt) = Buffer.add txt | add_content (Elem (name, ts)) = Buffer.add ("<" ^ name ^ ">") #> fold add_content ts #> Buffer.add (""); fun fast_content tree = Buffer.empty |> add_content tree |> Buffer.content; \ text \ The slowness of \<^ML>\slow_content\ is due to the \<^ML>\implode\ of the recursive results, because it copies previously produced strings again and again. The incremental \<^ML>\add_content\ avoids this by operating on a buffer that is passed through in a linear fashion. Using \<^ML_text>\#>\ and contraction over the actual buffer argument saves some additional boiler-plate. Of course, the two \<^ML>\Buffer.add\ invocations with concatenated strings could have been split into smaller parts, but this would have obfuscated the source without making a big difference in performance. Here we have done some peephole-optimization for the sake of readability. Another benefit of \<^ML>\add_content\ is its ``open'' form as a function on buffers that can be continued in further linear transformations, folding etc. Thus it is more compositional than the naive \<^ML>\slow_content\. As realistic example, compare the old-style \<^ML>\Term.maxidx_of_term: term -> int\ with the newer \<^ML>\Term.maxidx_term: term -> int -> int\ in Isabelle/Pure. Note that \<^ML>\fast_content\ above is only defined as example. In many practical situations, it is customary to provide the incremental \<^ML>\add_content\ only and leave the initialization and termination to the concrete application to the user. \ section \Message output channels \label{sec:message-channels}\ text \ Isabelle provides output channels for different kinds of messages: regular output, high-volume tracing information, warnings, and errors. Depending on the user interface involved, these messages may appear in different text styles or colours. The standard output for batch sessions prefixes each line of warnings by \<^verbatim>\###\ and errors by \<^verbatim>\***\, but leaves anything else unchanged. The message body may contain further markup and - formatting, which is routinely used in the Prover IDE @{cite - "isabelle-jedit"}. + formatting, which is routinely used in the Prover IDE \<^cite>\"isabelle-jedit"\. Messages are associated with the transaction context of the running Isar command. This enables the front-end to manage commands and resulting messages together. For example, after deleting a command from a given theory document version, the corresponding message output can be retracted from the display. \ text %mlref \ \begin{mldecls} @{define_ML writeln: "string -> unit"} \\ @{define_ML tracing: "string -> unit"} \\ @{define_ML warning: "string -> unit"} \\ @{define_ML error: "string -> 'a"} % FIXME Output.error_message (!?) \\ \end{mldecls} \<^descr> \<^ML>\writeln\~\text\ outputs \text\ as regular message. This is the primary message output operation of Isabelle and should be used by default. \<^descr> \<^ML>\tracing\~\text\ outputs \text\ as special tracing message, indicating potential high-volume output to the front-end (hundreds or thousands of messages issued by a single command). The idea is to allow the user-interface to downgrade the quality of message display to achieve higher throughput. Note that the user might have to take special actions to see tracing output, e.g.\ switch to a different output window. So this channel should not be used for regular output. \<^descr> \<^ML>\warning\~\text\ outputs \text\ as warning, which typically means some extra emphasis on the front-end side (color highlighting, icons, etc.). \<^descr> \<^ML>\error\~\text\ raises exception \<^ML>\ERROR\~\text\ and thus lets the Isar toplevel print \text\ on the error channel, which typically means some extra emphasis on the front-end side (color highlighting, icons, etc.). This assumes that the exception is not handled before the command terminates. Handling exception \<^ML>\ERROR\~\text\ is a perfectly legal alternative: it means that the error is absorbed without any message output. \begin{warn} The actual error channel is accessed via \<^ML>\Output.error_message\, but this is normally not used directly in user code. \end{warn} \begin{warn} Regular Isabelle/ML code should output messages exclusively by the official channels. Using raw I/O on \<^emph>\stdout\ or \<^emph>\stderr\ instead (e.g.\ via \<^ML>\TextIO.output\) is apt to cause problems in the presence of parallel and asynchronous processing of Isabelle theories. Such raw output might be displayed by the front-end in some system console log, with a low chance that the user will ever see it. Moreover, as a genuine side-effect on global process channels, there is no proper way to retract output when Isar command transactions are reset by the system. \end{warn} \begin{warn} The message channels should be used in a message-oriented manner. This means that multi-line output that logically belongs together is issued by a single invocation of \<^ML>\writeln\ etc.\ with the functional concatenation of all message constituents. \end{warn} \ text %mlex \ The following example demonstrates a multi-line warning. Note that in some situations the user sees only the first line, so the most important point should be made first. \ ML_command \ warning (cat_lines ["Beware the Jabberwock, my son!", "The jaws that bite, the claws that catch!", "Beware the Jubjub Bird, and shun", "The frumious Bandersnatch!"]); \ text \ \<^medskip> An alternative is to make a paragraph of freely-floating words as follows. \ ML_command \ warning (Pretty.string_of (Pretty.para "Beware the Jabberwock, my son! \ \The jaws that bite, the claws that catch! \ \Beware the Jubjub Bird, and shun \ \The frumious Bandersnatch!")) \ text \ This has advantages with variable window / popup sizes, but might make it harder to search for message content systematically, e.g.\ by other tools or by humans expecting the ``verse'' of a formal message in a fixed layout. \ section \Exceptions \label{sec:exceptions}\ text \ The Standard ML semantics of strict functional evaluation together with exceptions is rather well defined, but some delicate points need to be observed to avoid that ML programs go wrong despite static type-checking. Exceptions in Isabelle/ML are subsequently categorized as follows. \ paragraph \Regular user errors.\ text \ These are meant to provide informative feedback about malformed input etc. The \<^emph>\error\ function raises the corresponding \<^ML>\ERROR\ exception, with a plain text message as argument. \<^ML>\ERROR\ exceptions can be handled internally, in order to be ignored, turned into other exceptions, or cascaded by appending messages. If the corresponding Isabelle/Isar command terminates with an \<^ML>\ERROR\ exception state, the system will print the result on the error channel (see \secref{sec:message-channels}). It is considered bad style to refer to internal function names or values in ML source notation in user error messages. Do not use \@{make_string}\ nor \@{here}\! Grammatical correctness of error messages can be improved by \<^emph>\omitting\ final punctuation: messages are often concatenated or put into a larger context (e.g.\ augmented with source position). Note that punctuation after formal entities (types, terms, theorems) is particularly prone to user confusion. \ paragraph \Program failures.\ text \ There is a handful of standard exceptions that indicate general failure situations, or failures of core operations on logical entities (types, terms, theorems, theories, see \chref{ch:logic}). These exceptions indicate a genuine breakdown of the program, so the main purpose is to determine quickly what has happened where. Traditionally, the (short) exception message would include the name of an ML function, although this is no longer necessary, because the ML runtime system attaches detailed source position stemming from the corresponding \<^ML_text>\raise\ keyword. \<^medskip> User modules can always introduce their own custom exceptions locally, e.g.\ to organize internal failures robustly without overlapping with existing exceptions. Exceptions that are exposed in module signatures require extra care, though, and should \<^emph>\not\ be introduced by default. Surprise by users of a module can be often minimized by using plain user errors instead. \ paragraph \Interrupts.\ text \ These indicate arbitrary system events: both the ML runtime system and the Isabelle/ML infrastructure signal various exceptional situations by raising the special \<^ML>\Exn.Interrupt\ exception in user code. This is the one and only way that physical events can intrude an Isabelle/ML program. Such an interrupt can mean out-of-memory, stack overflow, timeout, internal signaling of threads, or a POSIX process signal. An Isabelle/ML program that intercepts interrupts becomes dependent on physical effects of the environment. Even worse, exception handling patterns that are too general by accident, e.g.\ by misspelled exception constructors, will cover interrupts unintentionally and thus render the program semantics ill-defined. Note that the Interrupt exception dates back to the original SML90 language definition. It was excluded from the SML97 version to avoid its malign impact on ML program semantics, but without providing a viable alternative. Isabelle/ML recovers physical interruptibility (which is an indispensable tool to implement managed evaluation of command transactions), but requires user code to be strictly transparent wrt.\ interrupts. \begin{warn} Isabelle/ML user code needs to terminate promptly on interruption, without guessing at its meaning to the system infrastructure. Temporary handling of interrupts for cleanup of global resources etc.\ needs to be followed immediately by re-raising of the original exception. \end{warn} \ text %mlref \ \begin{mldecls} @{define_ML try: "('a -> 'b) -> 'a -> 'b option"} \\ @{define_ML can: "('a -> 'b) -> 'a -> bool"} \\ @{define_ML_exception ERROR of string} \\ @{define_ML_exception Fail of string} \\ @{define_ML Exn.is_interrupt: "exn -> bool"} \\ @{define_ML Exn.reraise: "exn -> 'a"} \\ @{define_ML Runtime.exn_trace: "(unit -> 'a) -> 'a"} \\ \end{mldecls} \<^rail>\ (@@{ML_antiquotation try} | @@{ML_antiquotation can}) embedded \ \<^descr> \<^ML>\try\~\f x\ makes the partiality of evaluating \f x\ explicit via the option datatype. Interrupts are \<^emph>\not\ handled here, i.e.\ this form serves as safe replacement for the \<^emph>\unsafe\ version \<^ML_text>\(SOME\~\f x\~\<^ML_text>\handle _ => NONE)\ that is occasionally seen in books about SML97, but not in Isabelle/ML. \<^descr> \<^ML>\can\ is similar to \<^ML>\try\ with more abstract result. \<^descr> \<^ML>\ERROR\~\msg\ represents user errors; this exception is normally raised indirectly via the \<^ML>\error\ function (see \secref{sec:message-channels}). \<^descr> \<^ML>\Fail\~\msg\ represents general program failures. \<^descr> \<^ML>\Exn.is_interrupt\ identifies interrupts robustly, without mentioning concrete exception constructors in user code. Handled interrupts need to be re-raised promptly! \<^descr> \<^ML>\Exn.reraise\~\exn\ raises exception \exn\ while preserving its implicit position information (if possible, depending on the ML platform). \<^descr> \<^ML>\Runtime.exn_trace\~\<^ML_text>\(fn () =>\~\e\\<^ML_text>\)\ evaluates expression \e\ while printing a full trace of its stack of nested exceptions (if possible, depending on the ML platform). Inserting \<^ML>\Runtime.exn_trace\ into ML code temporarily is useful for debugging, but not suitable for production code. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "try"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "can"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "assert"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "undefined"} & : & \ML_antiquotation\ \\ \end{matharray} \<^descr> \@{try}\ and \{can}\ are similar to the corresponding functions, but the argument is taken directly as ML expression: functional abstraction and application is done implicitly. \<^descr> \@{assert}\ inlines a function \<^ML_type>\bool -> unit\ that raises \<^ML>\Fail\ if the argument is \<^ML>\false\. Due to inlining the source position of failed assertions is included in the error output. \<^descr> \@{undefined}\ inlines \<^verbatim>\raise\~\<^ML>\Match\, i.e.\ the ML program behaves as in some function application of an undefined case. \ text %mlex \ We define a total version of division: any failures are swept under the carpet and mapped to a default value. Thus division-by-zero becomes 0, but there could be other exceptions like overflow that produce the same result (for unbounded integers this does not happen). \ ML \ fun div_total x y = \<^try>\x div y\ |> the_default 0; \<^assert> (div_total 1 0 = 0); \ text \ The ML function \<^ML>\undefined\ is defined in \<^file>\~~/src/Pure/library.ML\ as follows: \ ML \fun undefined _ = raise Match\ text \ \<^medskip> The following variant uses the antiquotation @{ML_antiquotation undefined} instead: \ ML \fun undefined _ = \<^undefined>\ text \ \<^medskip> Here is the same, using control-symbol notation for the antiquotation, with special rendering of \<^verbatim>\\<^undefined>\: \ ML \fun undefined _ = \<^undefined>\ text \ \<^medskip> Semantically, all forms are equivalent to a function definition without any clauses, but that is syntactically not allowed in ML. \ section \Strings of symbols \label{sec:symbols}\ text \ A \<^emph>\symbol\ constitutes the smallest textual unit in Isabelle/ML --- raw ML characters are normally not encountered at all. Isabelle strings consist of a sequence of symbols, represented as a packed string or an exploded list of strings. Each symbol is in itself a small string, which has either one of the following forms: \<^enum> a single ASCII character ``\c\'', for example ``\<^verbatim>\a\'', \<^enum> a codepoint according to UTF-8 (non-ASCII byte sequence), \<^enum> a regular symbol ``\<^verbatim>\\\'', for example ``\<^verbatim>\\\'', \<^enum> a control symbol ``\<^verbatim>\\<^ident>\'', for example ``\<^verbatim>\\<^bold>\'', The \ident\ syntax for symbol names is \letter (letter | digit)\<^sup>*\, where \letter = A..Za..z\ and \digit = 0..9\. There are infinitely many regular symbols and control symbols, but a fixed collection of standard symbols is treated specifically. For example, ``\<^verbatim>\\\'' is classified as a letter, which means it may occur within regular Isabelle identifiers. The character set underlying Isabelle symbols is 7-bit ASCII, but 8-bit character sequences are passed-through unchanged. Unicode/UCS data in UTF-8 encoding is processed in a non-strict fashion, such that well-formed code sequences are recognized accordingly. Unicode provides its own collection of mathematical symbols, but within the core Isabelle/ML world there is no link to the standard collection of Isabelle regular symbols. \<^medskip> Output of Isabelle symbols depends on the print mode. For example, the standard {\LaTeX} setup of the Isabelle document preparation system would present ``\<^verbatim>\\\'' as \\\, and ``\<^verbatim>\\<^bold>\\'' as \\<^bold>\\. On-screen rendering usually works by mapping a finite subset of Isabelle symbols to suitable Unicode characters. \ text %mlref \ \begin{mldecls} @{define_ML_type Symbol.symbol = string} \\ @{define_ML Symbol.explode: "string -> Symbol.symbol list"} \\ @{define_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\ @{define_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\ @{define_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\ @{define_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type "Symbol.sym"} \\ @{define_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Symbol.symbol\ represents individual Isabelle symbols. \<^descr> \<^ML>\Symbol.explode\~\str\ produces a symbol list from the packed form. This function supersedes \<^ML>\String.explode\ for virtually all purposes of manipulating text in Isabelle!\<^footnote>\The runtime overhead for exploded strings is mainly that of the list structure: individual symbols that happen to be a singleton string do not require extra memory in Poly/ML.\ \<^descr> \<^ML>\Symbol.is_letter\, \<^ML>\Symbol.is_digit\, \<^ML>\Symbol.is_quasi\, \<^ML>\Symbol.is_blank\ classify standard symbols - according to fixed syntactic conventions of Isabelle, cf.\ @{cite - "isabelle-isar-ref"}. + according to fixed syntactic conventions of Isabelle, cf.\ \<^cite>\"isabelle-isar-ref"\. \<^descr> Type \<^ML_type>\Symbol.sym\ is a concrete datatype that represents the different kinds of symbols explicitly, with constructors \<^ML>\Symbol.Char\, \<^ML>\Symbol.UTF8\, \<^ML>\Symbol.Sym\, \<^ML>\Symbol.Control\, \<^ML>\Symbol.Malformed\. \<^descr> \<^ML>\Symbol.decode\ converts the string representation of a symbol into the datatype version. \ paragraph \Historical note.\ text \ In the original SML90 standard the primitive ML type \<^ML_type>\char\ did not exists, and \<^ML_text>\explode: string -> string list\ produced a list of singleton strings like \<^ML>\raw_explode: string -> string list\ in Isabelle/ML today. When SML97 came out, Isabelle did not adopt its somewhat anachronistic 8-bit or 16-bit characters, but the idea of exploding a string into a list of small strings was extended to ``symbols'' as explained above. Thus Isabelle sources can refer to an infinite store of user-defined symbols, without having to worry about the multitude of Unicode encodings that have emerged over the years. \ section \Basic data types\ text \ The basis library proposal of SML97 needs to be treated with caution. Many of its operations simply do not fit with important Isabelle/ML conventions (like ``canonical argument order'', see \secref{sec:canonical-argument-order}), others cause problems with the parallel evaluation model of Isabelle/ML (such as \<^ML>\TextIO.print\ or \<^ML>\OS.Process.system\). Subsequently we give a brief overview of important operations on basic ML data types. \ subsection \Characters\ text %mlref \ \begin{mldecls} @{define_ML_type char} \\ \end{mldecls} \<^descr> Type \<^ML_type>\char\ is \<^emph>\not\ used. The smallest textual unit in Isabelle is represented as a ``symbol'' (see \secref{sec:symbols}). \ subsection \Strings\ text %mlref \ \begin{mldecls} @{define_ML_type string} \\ \end{mldecls} \<^descr> Type \<^ML_type>\string\ represents immutable vectors of 8-bit characters. There are operations in SML to convert back and forth to actual byte vectors, which are seldom used. This historically important raw text representation is used for Isabelle-specific purposes with the following implicit substructures packed into the string content: \<^enum> sequence of Isabelle symbols (see also \secref{sec:symbols}), with \<^ML>\Symbol.explode\ as key operation; - \<^enum> XML tree structure via YXML (see also @{cite "isabelle-system"}), with + \<^enum> XML tree structure via YXML (see also \<^cite>\"isabelle-system"\), with \<^ML>\YXML.parse_body\ as key operation. Note that Isabelle/ML string literals may refer Isabelle symbols like ``\<^verbatim>\\\'' natively, \<^emph>\without\ escaping the backslash. This is a consequence of Isabelle treating all source text as strings of symbols, instead of raw characters. \begin{warn} The regular \<^verbatim>\64_32\ platform of Poly/ML has a size limit of 64\,MB for \<^ML_type>\string\ values. This is usually sufficient for text applications, with a little bit of YXML markup. Very large XML trees or binary blobs are better stored as scalable byte strings, see type \<^ML_type>\Bytes.T\ and corresponding operations in \<^file>\~~/src/Pure/General/bytes.ML\. \end{warn} \ text %mlex \ The subsequent example illustrates the difference of physical addressing of bytes versus logical addressing of symbols in Isabelle strings. \ ML_val \ val s = "\"; \<^assert> (length (Symbol.explode s) = 1); \<^assert> (size s = 4); \ text \ Note that in Unicode renderings of the symbol \\\, variations of encodings like UTF-8 or UTF-16 pose delicate questions about the multi-byte representations of its codepoint, which is outside of the 16-bit address space of the original Unicode standard from the 1990-ies. In Isabelle/ML it is just ``\<^verbatim>\\\'' literally, using plain ASCII characters beyond any doubts. \ subsection \Integers\ text %mlref \ \begin{mldecls} @{define_ML_type int} \\ \end{mldecls} \<^descr> Type \<^ML_type>\int\ represents regular mathematical integers, which are \<^emph>\unbounded\. Overflow is treated properly, but should never happen in practice.\<^footnote>\The size limit for integer bit patterns in memory is 64\,MB for the regular \<^verbatim>\64_32\ platform, and much higher for native \<^verbatim>\64\ architecture.\ Structure \<^ML_structure>\IntInf\ of SML97 is obsolete and superseded by \<^ML_structure>\Int\. Structure \<^ML_structure>\Integer\ in \<^file>\~~/src/Pure/General/integer.ML\ provides some additional operations. \ subsection \Rational numbers\ text %mlref \ \begin{mldecls} @{define_ML_type Rat.rat} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Rat.rat\ represents rational numbers, based on the unbounded integers of Poly/ML. Literal rationals may be written with special antiquotation syntax \<^verbatim>\@\\int\\<^verbatim>\/\\nat\ or \<^verbatim>\@\\int\ (without any white space). For example \<^verbatim>\@~1/4\ or \<^verbatim>\@10\. The ML toplevel pretty printer uses the same format. Standard operations are provided via ad-hoc overloading of \<^verbatim>\+\, \<^verbatim>\-\, \<^verbatim>\*\, \<^verbatim>\/\, etc. \ subsection \Time\ text %mlref \ \begin{mldecls} @{define_ML_type Time.time} \\ @{define_ML seconds: "real -> Time.time"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Time.time\ represents time abstractly according to the SML97 basis library definition. This is adequate for internal ML operations, but awkward in concrete time specifications. \<^descr> \<^ML>\seconds\~\s\ turns the concrete scalar \s\ (measured in seconds) into an abstract time value. Floating point numbers are easy to use as configuration options in the context (see \secref{sec:config-options}) or system options that are maintained externally. \ subsection \Options\ text %mlref \ \begin{mldecls} @{define_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\ @{define_ML is_some: "'a option -> bool"} \\ @{define_ML is_none: "'a option -> bool"} \\ @{define_ML the: "'a option -> 'a"} \\ @{define_ML these: "'a list option -> 'a list"} \\ @{define_ML the_list: "'a option -> 'a list"} \\ @{define_ML the_default: "'a -> 'a option -> 'a"} \\ \end{mldecls} \ text \ Apart from \<^ML>\Option.map\ most other operations defined in structure \<^ML_structure>\Option\ are alien to Isabelle/ML and never used. The operations shown above are defined in \<^file>\~~/src/Pure/General/basics.ML\. \ subsection \Lists\ text \ Lists are ubiquitous in ML as simple and light-weight ``collections'' for many everyday programming tasks. Isabelle/ML provides important additions and improvements over operations that are predefined in the SML97 library. \ text %mlref \ \begin{mldecls} @{define_ML cons: "'a -> 'a list -> 'a list"} \\ @{define_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\ @{define_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\ @{define_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\ @{define_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\ \end{mldecls} \<^descr> \<^ML>\cons\~\x xs\ evaluates to \x :: xs\. Tupled infix operators are a historical accident in Standard ML. The curried \<^ML>\cons\ amends this, but it should be only used when partial application is required. \<^descr> \<^ML>\member\, \<^ML>\insert\, \<^ML>\remove\, \<^ML>\update\ treat lists as a set-like container that maintains the order of elements. See \<^file>\~~/src/Pure/library.ML\ for the full specifications (written in ML). There are some further derived operations like \<^ML>\union\ or \<^ML>\inter\. Note that \<^ML>\insert\ is conservative about elements that are already a \<^ML>\member\ of the list, while \<^ML>\update\ ensures that the latest entry is always put in front. The latter discipline is often more appropriate in declarations of context data (\secref{sec:context-data}) that are issued by the user in Isar source: later declarations take precedence over earlier ones. \ text %mlex \ Using canonical \<^ML>\fold\ together with \<^ML>\cons\ (or similar standard operations) alternates the orientation of data. The is quite natural and should not be altered forcible by inserting extra applications of \<^ML>\rev\. The alternative \<^ML>\fold_rev\ can be used in the few situations, where alternation should be prevented. \ ML_val \ val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10]; val list1 = fold cons items []; \<^assert> (list1 = rev items); val list2 = fold_rev cons items []; \<^assert> (list2 = items); \ text \ The subsequent example demonstrates how to \<^emph>\merge\ two lists in a natural way. \ ML_val \ fun merge_lists eq (xs, ys) = fold_rev (insert eq) ys xs; \ text \ Here the first list is treated conservatively: only the new elements from the second list are inserted. The inside-out order of insertion via \<^ML>\fold_rev\ attempts to preserve the order of elements in the result. This way of merging lists is typical for context data (\secref{sec:context-data}). See also \<^ML>\merge\ as defined in \<^file>\~~/src/Pure/library.ML\. \ subsection \Association lists\ text \ The operations for association lists interpret a concrete list of pairs as a finite function from keys to values. Redundant representations with multiple occurrences of the same key are implicitly normalized: lookup and update only take the first occurrence into account. \ text \ \begin{mldecls} @{define_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\ @{define_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\ @{define_ML AList.update: "('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list"} \\ \end{mldecls} \<^descr> \<^ML>\AList.lookup\, \<^ML>\AList.defined\, \<^ML>\AList.update\ implement the main ``framework operations'' for mappings in Isabelle/ML, following standard conventions for their names and types. Note that a function called \<^verbatim>\lookup\ is obliged to express its partiality via an explicit option element. There is no choice to raise an exception, without changing the name to something like \the_element\ or \get\. The \defined\ operation is essentially a contraction of \<^ML>\is_some\ and \<^verbatim>\lookup\, but this is sufficiently frequent to justify its independent existence. This also gives the implementation some opportunity for peep-hole optimization. Association lists are adequate as simple implementation of finite mappings in many practical situations. A more advanced table structure is defined in \<^file>\~~/src/Pure/General/table.ML\; that version scales easily to thousands or millions of elements. \ subsection \Unsynchronized references\ text %mlref \ \begin{mldecls} @{define_ML_type 'a "Unsynchronized.ref"} \\ @{define_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\ @{define_ML "!": "'a Unsynchronized.ref -> 'a"} \\ @{define_ML_infix ":=" : "'a Unsynchronized.ref * 'a -> unit"} \\ \end{mldecls} \ text \ Due to ubiquitous parallelism in Isabelle/ML (see also \secref{sec:multi-threading}), the mutable reference cells of Standard ML are notorious for causing problems. In a highly parallel system, both correctness \<^emph>\and\ performance are easily degraded when using mutable data. The unwieldy name of \<^ML>\Unsynchronized.ref\ for the constructor for references in Isabelle/ML emphasizes the inconveniences caused by mutability. Existing operations \<^ML>\!\ and \<^ML_infix>\:=\ are unchanged, but should be used with special precautions, say in a strictly local situation that is guaranteed to be restricted to sequential evaluation --- now and in the future. \begin{warn} Never \<^ML_text>\open Unsynchronized\, not even in a local scope! Pretending that mutable state is no problem is a very bad idea. \end{warn} \ section \Thread-safe programming \label{sec:multi-threading}\ text \ Multi-threaded execution has become an everyday reality in Isabelle since Poly/ML 5.2.1 and Isabelle2008. Isabelle/ML provides implicit and explicit parallelism by default, and there is no way for user-space tools to ``opt out''. ML programs that are purely functional, output messages only via the official channels (\secref{sec:message-channels}), and do not intercept interrupts (\secref{sec:exceptions}) can participate in the multi-threaded environment immediately without further ado. More ambitious tools with more fine-grained interaction with the environment need to observe the principles explained below. \ subsection \Multi-threading with shared memory\ text \ Multiple threads help to organize advanced operations of the system, such as real-time conditions on command transactions, sub-components with explicit communication, general asynchronous interaction etc. Moreover, parallel evaluation is a prerequisite to make adequate use of the CPU resources that are available on multi-core systems.\<^footnote>\Multi-core computing does not mean that there are ``spare cycles'' to be wasted. It means that the continued exponential speedup of CPU performance due to ``Moore's Law'' follows different rules: clock frequency has reached its peak around 2005, and applications need to be parallelized in order to avoid a perceived loss of - performance. See also @{cite "Sutter:2005"}.\ + performance. See also \<^cite>\"Sutter:2005"\.\ Isabelle/Isar exploits the inherent structure of theories and proofs to support \<^emph>\implicit parallelism\ to a large extent. LCF-style theorem proving - provides almost ideal conditions for that, see also @{cite "Wenzel:2009"}. + provides almost ideal conditions for that, see also \<^cite>\"Wenzel:2009"\. This means, significant parts of theory and proof checking is parallelized by default. In Isabelle2013, a maximum speedup-factor of 3.5 on 4 cores and - 6.5 on 8 cores can be expected @{cite "Wenzel:2013:ITP"}. + 6.5 on 8 cores can be expected \<^cite>\"Wenzel:2013:ITP"\. \<^medskip> ML threads lack the memory protection of separate processes, and operate concurrently on shared heap memory. This has the advantage that results of independent computations are directly available to other threads: abstract values can be passed without copying or awkward serialization that is typically required for separate processes. To make shared-memory multi-threading work robustly and efficiently, some programming guidelines need to be observed. While the ML system is responsible to maintain basic integrity of the representation of ML values in memory, the application programmer needs to ensure that multi-threaded execution does not break the intended semantics. \begin{warn} To participate in implicit parallelism, tools need to be thread-safe. A single ill-behaved tool can affect the stability and performance of the whole system. \end{warn} Apart from observing the principles of thread-safeness passively, advanced tools may also exploit parallelism actively, e.g.\ by using library functions for parallel list operations (\secref{sec:parlist}). \begin{warn} Parallel computing resources are managed centrally by the Isabelle/ML infrastructure. User programs should not fork their own ML threads to perform heavy computations. \end{warn} \ subsection \Critical shared resources\ text \ Thread-safeness is mainly concerned about concurrent read/write access to shared resources, which are outside the purely functional world of ML. This covers the following in particular. \<^item> Global references (or arrays), i.e.\ mutable memory cells that persist over several invocations of associated operations.\<^footnote>\This is independent of the visibility of such mutable values in the toplevel scope.\ \<^item> Global state of the running Isabelle/ML process, i.e.\ raw I/O channels, environment variables, current working directory. \<^item> Writable resources in the file-system that are shared among different threads or external processes. Isabelle/ML provides various mechanisms to avoid critical shared resources in most situations. As last resort there are some mechanisms for explicit synchronization. The following guidelines help to make Isabelle/ML programs work smoothly in a concurrent environment. \<^item> Avoid global references altogether. Isabelle/Isar maintains a uniform context that incorporates arbitrary data declared by user programs (\secref{sec:context-data}). This context is passed as plain value and user tools can get/map their own data in a purely functional manner. Configuration options within the context (\secref{sec:config-options}) provide simple drop-in replacements for historic reference variables. \<^item> Keep components with local state information re-entrant. Instead of poking initial values into (private) global references, a new state record can be created on each invocation, and passed through any auxiliary functions of the component. The state record contain mutable references in special situations, without requiring any synchronization, as long as each invocation gets its own copy and the tool itself is single-threaded. \<^item> Avoid raw output on \stdout\ or \stderr\. The Poly/ML library is thread-safe for each individual output operation, but the ordering of parallel invocations is arbitrary. This means raw output will appear on some system console with unpredictable interleaving of atomic chunks. Note that this does not affect regular message output channels (\secref{sec:message-channels}). An official message id is associated with the command transaction from where it originates, independently of other transactions. This means each running Isar command has effectively its own set of message channels, and interleaving can only happen when commands use parallelism internally (and only at message boundaries). \<^item> Treat environment variables and the current working directory of the running process as read-only. \<^item> Restrict writing to the file-system to unique temporary files. Isabelle already provides a temporary directory that is unique for the running process, and there is a centralized source of unique serial numbers in Isabelle/ML. Thus temporary files that are passed to to some external process will be always disjoint, and thus thread-safe. \ text %mlref \ \begin{mldecls} @{define_ML File.tmp_path: "Path.T -> Path.T"} \\ @{define_ML serial_string: "unit -> string"} \\ \end{mldecls} \<^descr> \<^ML>\File.tmp_path\~\path\ relocates the base component of \path\ into the unique temporary directory of the running Isabelle/ML process. \<^descr> \<^ML>\serial_string\~\()\ creates a new serial number that is unique over the runtime of the Isabelle/ML process. \ text %mlex \ The following example shows how to create unique temporary file names. \ ML_val \ val tmp1 = File.tmp_path (Path.basic ("foo" ^ serial_string ())); val tmp2 = File.tmp_path (Path.basic ("foo" ^ serial_string ())); \<^assert> (tmp1 <> tmp2); \ subsection \Explicit synchronization\ text \ Isabelle/ML provides explicit synchronization for mutable variables over immutable data, which may be updated atomically and exclusively. This addresses the rare situations where mutable shared resources are really required. Synchronization in Isabelle/ML is based on primitives of Poly/ML, which have been adapted to the specific assumptions of the concurrent Isabelle environment. User code should not break this abstraction, but stay within the confines of concurrent Isabelle/ML. A \<^emph>\synchronized variable\ is an explicit state component associated with mechanisms for locking and signaling. There are operations to await a condition, change the state, and signal the change to all other waiting threads. Synchronized access to the state variable is \<^emph>\not\ re-entrant: direct or indirect nesting within the same thread will cause a deadlock! \ text %mlref \ \begin{mldecls} @{define_ML_type 'a "Synchronized.var"} \\ @{define_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\ @{define_ML Synchronized.guarded_access: "'a Synchronized.var -> ('a -> ('b * 'a) option) -> 'b"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\'a Synchronized.var\ represents synchronized variables with state of type \<^ML_type>\'a\. \<^descr> \<^ML>\Synchronized.var\~\name x\ creates a synchronized variable that is initialized with value \x\. The \name\ is used for tracing. \<^descr> \<^ML>\Synchronized.guarded_access\~\var f\ lets the function \f\ operate within a critical section on the state \x\ as follows: if \f x\ produces \<^ML>\NONE\, it continues to wait on the internal condition variable, expecting that some other thread will eventually change the content in a suitable manner; if \f x\ produces \<^ML>\SOME\~\(y, x')\ it is satisfied and assigns the new state value \x'\, broadcasts a signal to all waiting threads on the associated condition variable, and returns the result \y\. There are some further variants of the \<^ML>\Synchronized.guarded_access\ combinator, see \<^file>\~~/src/Pure/Concurrent/synchronized.ML\ for details. \ text %mlex \ The following example implements a counter that produces positive integers that are unique over the runtime of the Isabelle process: \ ML_val \ local val counter = Synchronized.var "counter" 0; in fun next () = Synchronized.guarded_access counter (fn i => let val j = i + 1 in SOME (j, j) end); end; val a = next (); val b = next (); \<^assert> (a <> b); \ text \ \<^medskip> See \<^file>\~~/src/Pure/Concurrent/mailbox.ML\ how to implement a mailbox as synchronized variable over a purely functional list. \ section \Managed evaluation\ text \ Execution of Standard ML follows the model of strict functional evaluation with optional exceptions. Evaluation happens whenever some function is applied to (sufficiently many) arguments. The result is either an explicit value or an implicit exception. \<^emph>\Managed evaluation\ in Isabelle/ML organizes expressions and results to control certain physical side-conditions, to say more specifically when and how evaluation happens. For example, the Isabelle/ML library supports lazy evaluation with memoing, parallel evaluation via futures, asynchronous evaluation via promises, evaluation with time limit etc. \<^medskip> An \<^emph>\unevaluated expression\ is represented either as unit abstraction \<^verbatim>\fn () => a\ of type \<^verbatim>\unit -> 'a\ or as regular function \<^verbatim>\fn a => b\ of type \<^verbatim>\'a -> 'b\. Both forms occur routinely, and special care is required to tell them apart --- the static type-system of SML is only of limited help here. The first form is more intuitive: some combinator \<^verbatim>\(unit -> 'a) -> 'a\ applies the given function to \<^verbatim>\()\ to initiate the postponed evaluation process. The second form is more flexible: some combinator \<^verbatim>\('a -> 'b) -> 'a -> 'b\ acts like a modified form of function application; several such combinators may be cascaded to modify a given function, before it is ultimately applied to some argument. \<^medskip> \<^emph>\Reified results\ make the disjoint sum of regular values versions exceptional situations explicit as ML datatype: \'a result = Res of 'a | Exn of exn\. This is typically used for administrative purposes, to store the overall outcome of an evaluation process. \<^emph>\Parallel exceptions\ aggregate reified results, such that multiple exceptions are digested as a collection in canonical form that identifies exceptions according to their original occurrence. This is particular important for parallel evaluation via futures \secref{sec:futures}, which are organized as acyclic graph of evaluations that depend on other evaluations: exceptions stemming from shared sub-graphs are exposed exactly once and in the order of their original occurrence (e.g.\ when printed at the toplevel). Interrupt counts as neutral element here: it is treated as minimal information about some canceled evaluation process, and is absorbed by the presence of regular program exceptions. \ text %mlref \ \begin{mldecls} @{define_ML_type 'a "Exn.result"} \\ @{define_ML Exn.capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\ @{define_ML Exn.interruptible_capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\ @{define_ML Exn.release: "'a Exn.result -> 'a"} \\ @{define_ML Par_Exn.release_all: "'a Exn.result list -> 'a list"} \\ @{define_ML Par_Exn.release_first: "'a Exn.result list -> 'a list"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\'a Exn.result\ represents the disjoint sum of ML results explicitly, with constructor \<^ML>\Exn.Res\ for regular values and \<^ML>\Exn.Exn\ for exceptions. \<^descr> \<^ML>\Exn.capture\~\f x\ manages the evaluation of \f x\ such that exceptions are made explicit as \<^ML>\Exn.Exn\. Note that this includes physical interrupts (see also \secref{sec:exceptions}), so the same precautions apply to user code: interrupts must not be absorbed accidentally! \<^descr> \<^ML>\Exn.interruptible_capture\ is similar to \<^ML>\Exn.capture\, but interrupts are immediately re-raised as required for user code. \<^descr> \<^ML>\Exn.release\~\result\ releases the original runtime result, exposing its regular value or raising the reified exception. \<^descr> \<^ML>\Par_Exn.release_all\~\results\ combines results that were produced independently (e.g.\ by parallel evaluation). If all results are regular values, that list is returned. Otherwise, the collection of all exceptions is raised, wrapped-up as collective parallel exception. Note that the latter prevents access to individual exceptions by conventional \<^verbatim>\handle\ of ML. \<^descr> \<^ML>\Par_Exn.release_first\ is similar to \<^ML>\Par_Exn.release_all\, but only the first (meaningful) exception that has occurred in the original evaluation process is raised again, the others are ignored. That single exception may get handled by conventional means in ML. \ subsection \Parallel skeletons \label{sec:parlist}\ text \ Algorithmic skeletons are combinators that operate on lists in parallel, in the manner of well-known \map\, \exists\, \forall\ etc. Management of futures (\secref{sec:futures}) and their results as reified exceptions is wrapped up into simple programming interfaces that resemble the sequential versions. What remains is the application-specific problem to present expressions with suitable \<^emph>\granularity\: each list element corresponds to one evaluation task. If the granularity is too coarse, the available CPUs are not saturated. If it is too fine-grained, CPU cycles are wasted due to the overhead of organizing parallel processing. In the worst case, parallel performance will be less than the sequential counterpart! \ text %mlref \ \begin{mldecls} @{define_ML Par_List.map: "('a -> 'b) -> 'a list -> 'b list"} \\ @{define_ML Par_List.get_some: "('a -> 'b option) -> 'a list -> 'b option"} \\ \end{mldecls} \<^descr> \<^ML>\Par_List.map\~\f [x\<^sub>1, \, x\<^sub>n]\ is like \<^ML>\map\~\f [x\<^sub>1, \, x\<^sub>n]\, but the evaluation of \f x\<^sub>i\ for \i = 1, \, n\ is performed in parallel. An exception in any \f x\<^sub>i\ cancels the overall evaluation process. The final result is produced via \<^ML>\Par_Exn.release_first\ as explained above, which means the first program exception that happened to occur in the parallel evaluation is propagated, and all other failures are ignored. \<^descr> \<^ML>\Par_List.get_some\~\f [x\<^sub>1, \, x\<^sub>n]\ produces some \f x\<^sub>i\ that is of the form \SOME y\<^sub>i\, if that exists, otherwise \NONE\. Thus it is similar to \<^ML>\Library.get_first\, but subject to a non-deterministic parallel choice process. The first successful result cancels the overall evaluation process; other exceptions are propagated as for \<^ML>\Par_List.map\. This generic parallel choice combinator is the basis for derived forms, such as \<^ML>\Par_List.find_some\, \<^ML>\Par_List.exists\, \<^ML>\Par_List.forall\. \ text %mlex \ Subsequently, the Ackermann function is evaluated in parallel for some ranges of arguments. \ ML_val \ fun ackermann 0 n = n + 1 | ackermann m 0 = ackermann (m - 1) 1 | ackermann m n = ackermann (m - 1) (ackermann m (n - 1)); Par_List.map (ackermann 2) (500 upto 1000); Par_List.map (ackermann 3) (5 upto 10); \ subsection \Lazy evaluation\ text \ Classic lazy evaluation works via the \lazy\~/ \force\ pair of operations: \lazy\ to wrap an unevaluated expression, and \force\ to evaluate it once and store its result persistently. Later invocations of \force\ retrieve the stored result without another evaluation. Isabelle/ML refines this idea to accommodate the aspects of multi-threading, synchronous program exceptions and asynchronous interrupts. The first thread that invokes \force\ on an unfinished lazy value changes its state into a \<^emph>\promise\ of the eventual result and starts evaluating it. Any other threads that \force\ the same lazy value in the meantime need to wait for it to finish, by producing a regular result or program exception. If the evaluation attempt is interrupted, this event is propagated to all waiting threads and the lazy value is reset to its original state. This means a lazy value is completely evaluated at most once, in a thread-safe manner. There might be multiple interrupted evaluation attempts, and multiple receivers of intermediate interrupt events. Interrupts are \<^emph>\not\ made persistent: later evaluation attempts start again from the original expression. \ text %mlref \ \begin{mldecls} @{define_ML_type 'a "lazy"} \\ @{define_ML Lazy.lazy: "(unit -> 'a) -> 'a lazy"} \\ @{define_ML Lazy.value: "'a -> 'a lazy"} \\ @{define_ML Lazy.force: "'a lazy -> 'a"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\'a lazy\ represents lazy values over type \<^verbatim>\'a\. \<^descr> \<^ML>\Lazy.lazy\~\(fn () => e)\ wraps the unevaluated expression \e\ as unfinished lazy value. \<^descr> \<^ML>\Lazy.value\~\a\ wraps the value \a\ as finished lazy value. When forced, it returns \a\ without any further evaluation. There is very low overhead for this proforma wrapping of strict values as lazy values. \<^descr> \<^ML>\Lazy.force\~\x\ produces the result of the lazy value in a thread-safe manner as explained above. Thus it may cause the current thread to wait on a pending evaluation attempt by another thread. \ subsection \Futures \label{sec:futures}\ text \ Futures help to organize parallel execution in a value-oriented manner, with \fork\~/ \join\ as the main pair of operations, and some further variants; - see also @{cite "Wenzel:2009" and "Wenzel:2013:ITP"}. Unlike lazy values, + see also \<^cite>\"Wenzel:2009" and "Wenzel:2013:ITP"\. Unlike lazy values, futures are evaluated strictly and spontaneously on separate worker threads. Futures may be canceled, which leads to interrupts on running evaluation attempts, and forces structurally related futures to fail for all time; already finished futures remain unchanged. Exceptions between related futures are propagated as well, and turned into parallel exceptions (see above). Technically, a future is a single-assignment variable together with a \<^emph>\task\ that serves administrative purposes, notably within the \<^emph>\task queue\ where new futures are registered for eventual evaluation and the worker threads retrieve their work. The pool of worker threads is limited, in correlation with the number of physical cores on the machine. Note that allocation of runtime resources may be distorted either if workers yield CPU time (e.g.\ via system sleep or wait operations), or if non-worker threads contend for significant runtime resources independently. There is a limited number of replacement worker threads that get activated in certain explicit wait conditions, after a timeout. \<^medskip> Each future task belongs to some \<^emph>\task group\, which represents the hierarchic structure of related tasks, together with the exception status a that point. By default, the task group of a newly created future is a new sub-group of the presently running one, but it is also possible to indicate different group layouts under program control. Cancellation of futures actually refers to the corresponding task group and all its sub-groups. Thus interrupts are propagated down the group hierarchy. Regular program exceptions are treated likewise: failure of the evaluation of some future task affects its own group and all sub-groups. Given a particular task group, its \<^emph>\group status\ cumulates all relevant exceptions according to its position within the group hierarchy. Interrupted tasks that lack regular result information, will pick up parallel exceptions from the cumulative group status. \<^medskip> A \<^emph>\passive future\ or \<^emph>\promise\ is a future with slightly different evaluation policies: there is only a single-assignment variable and some expression to evaluate for the \<^emph>\failed\ case (e.g.\ to clean up resources when canceled). A regular result is produced by external means, using a separate \<^emph>\fulfill\ operation. Promises are managed in the same task queue, so regular futures may depend on them. This allows a form of reactive programming, where some promises are used as minimal elements (or guards) within the future dependency graph: when these promises are fulfilled the evaluation of subsequent futures starts spontaneously, according to their own inter-dependencies. \ text %mlref \ \begin{mldecls} @{define_ML_type 'a "future"} \\ @{define_ML Future.fork: "(unit -> 'a) -> 'a future"} \\ @{define_ML Future.forks: "Future.params -> (unit -> 'a) list -> 'a future list"} \\ @{define_ML Future.join: "'a future -> 'a"} \\ @{define_ML Future.joins: "'a future list -> 'a list"} \\ @{define_ML Future.value: "'a -> 'a future"} \\ @{define_ML Future.map: "('a -> 'b) -> 'a future -> 'b future"} \\ @{define_ML Future.cancel: "'a future -> unit"} \\ @{define_ML Future.cancel_group: "Future.group -> unit"} \\[0.5ex] @{define_ML Future.promise: "(unit -> unit) -> 'a future"} \\ @{define_ML Future.fulfill: "'a future -> 'a -> unit"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\'a future\ represents future values over type \<^verbatim>\'a\. \<^descr> \<^ML>\Future.fork\~\(fn () => e)\ registers the unevaluated expression \e\ as unfinished future value, to be evaluated eventually on the parallel worker-thread farm. This is a shorthand for \<^ML>\Future.forks\ below, with default parameters and a single expression. \<^descr> \<^ML>\Future.forks\~\params exprs\ is the general interface to fork several futures simultaneously. The \params\ consist of the following fields: \<^item> \name : string\ (default \<^ML>\""\) specifies a common name for the tasks of the forked futures, which serves diagnostic purposes. \<^item> \group : Future.group option\ (default \<^ML>\NONE\) specifies an optional task group for the forked futures. \<^ML>\NONE\ means that a new sub-group of the current worker-thread task context is created. If this is not a worker thread, the group will be a new root in the group hierarchy. \<^item> \deps : Future.task list\ (default \<^ML>\[]\) specifies dependencies on other future tasks, i.e.\ the adjacency relation in the global task queue. Dependencies on already finished tasks are ignored. \<^item> \pri : int\ (default \<^ML>\0\) specifies a priority within the task queue. Typically there is only little deviation from the default priority \<^ML>\0\. As a rule of thumb, \<^ML>\~1\ means ``low priority" and \<^ML>\1\ means ``high priority''. Note that the task priority only affects the position in the queue, not the thread priority. When a worker thread picks up a task for processing, it runs with the normal thread priority to the end (or until canceled). Higher priority tasks that are queued later need to wait until this (or another) worker thread becomes free again. \<^item> \interrupts : bool\ (default \<^ML>\true\) tells whether the worker thread that processes the corresponding task is initially put into interruptible state. This state may change again while running, by modifying the thread attributes. With interrupts disabled, a running future task cannot be canceled. It is the responsibility of the programmer that this special state is retained only briefly. \<^descr> \<^ML>\Future.join\~\x\ retrieves the value of an already finished future, which may lead to an exception, according to the result of its previous evaluation. For an unfinished future there are several cases depending on the role of the current thread and the status of the future. A non-worker thread waits passively until the future is eventually evaluated. A worker thread temporarily changes its task context and takes over the responsibility to evaluate the future expression on the spot. The latter is done in a thread-safe manner: other threads that intend to join the same future need to wait until the ongoing evaluation is finished. Note that excessive use of dynamic dependencies of futures by adhoc joining may lead to bad utilization of CPU cores, due to threads waiting on other threads to finish required futures. The future task farm has a limited amount of replacement threads that continue working on unrelated tasks after some timeout. Whenever possible, static dependencies of futures should be specified explicitly when forked (see \deps\ above). Thus the evaluation can work from the bottom up, without join conflicts and wait states. \<^descr> \<^ML>\Future.joins\~\xs\ joins the given list of futures simultaneously, which is more efficient than \<^ML>\map Future.join\~\xs\. Based on the dependency graph of tasks, the current thread takes over the responsibility to evaluate future expressions that are required for the main result, working from the bottom up. Waiting on future results that are presently evaluated on other threads only happens as last resort, when no other unfinished futures are left over. \<^descr> \<^ML>\Future.value\~\a\ wraps the value \a\ as finished future value, bypassing the worker-thread farm. When joined, it returns \a\ without any further evaluation. There is very low overhead for this proforma wrapping of strict values as futures. \<^descr> \<^ML>\Future.map\~\f x\ is a fast-path implementation of \<^ML>\Future.fork\~\(fn () => f (\\<^ML>\Future.join\~\x))\, which avoids the full overhead of the task queue and worker-thread farm as far as possible. The function \f\ is supposed to be some trivial post-processing or projection of the future result. \<^descr> \<^ML>\Future.cancel\~\x\ cancels the task group of the given future, using \<^ML>\Future.cancel_group\ below. \<^descr> \<^ML>\Future.cancel_group\~\group\ cancels all tasks of the given task group for all time. Threads that are presently processing a task of the given group are interrupted: it may take some time until they are actually terminated. Tasks that are queued but not yet processed are dequeued and forced into interrupted state. Since the task group is itself invalidated, any further attempt to fork a future that belongs to it will yield a canceled result as well. \<^descr> \<^ML>\Future.promise\~\abort\ registers a passive future with the given \abort\ operation: it is invoked when the future task group is canceled. \<^descr> \<^ML>\Future.fulfill\~\x a\ finishes the passive future \x\ by the given value \a\. If the promise has already been canceled, the attempt to fulfill it causes an exception. \ end diff --git a/src/Doc/Implementation/Prelim.thy b/src/Doc/Implementation/Prelim.thy --- a/src/Doc/Implementation/Prelim.thy +++ b/src/Doc/Implementation/Prelim.thy @@ -1,973 +1,973 @@ (*:maxLineLen=78:*) theory Prelim imports Base begin chapter \Preliminaries\ section \Contexts \label{sec:context}\ text \ A logical context represents the background that is required for formulating statements and composing proofs. It acts as a medium to produce formal content, depending on earlier material (declarations, results etc.). For example, derivations within the Isabelle/Pure logic can be described as a judgment \\ \\<^sub>\ \\, which means that a proposition \\\ is derivable from hypotheses \\\ within the theory \\\. There are logical reasons for keeping \\\ and \\\ separate: theories can be liberal about supporting type constructors and schematic polymorphism of constants and axioms, while the inner calculus of \\ \ \\ is strictly limited to Simple Type Theory (with fixed type variables in the assumptions). \<^medskip> Contexts and derivations are linked by the following key principles: \<^item> Transfer: monotonicity of derivations admits results to be transferred into a \<^emph>\larger\ context, i.e.\ \\ \\<^sub>\ \\ implies \\' \\<^sub>\\<^sub>' \\ for contexts \\' \ \\ and \\' \ \\. \<^item> Export: discharge of hypotheses admits results to be exported into a \<^emph>\smaller\ context, i.e.\ \\' \\<^sub>\ \\ implies \\ \\<^sub>\ \ \ \\ where \\' \ \\ and \\ = \' - \\. Note that \\\ remains unchanged here, only the \\\ part is affected. \<^medskip> By modeling the main characteristics of the primitive \\\ and \\\ above, and abstracting over any particular logical content, we arrive at the fundamental notions of \<^emph>\theory context\ and \<^emph>\proof context\ in Isabelle/Isar. These implement a certain policy to manage arbitrary \<^emph>\context data\. There is a strongly-typed mechanism to declare new kinds of data at compile time. The internal bootstrap process of Isabelle/Pure eventually reaches a stage where certain data slots provide the logical content of \\\ and \\\ sketched above, but this does not stop there! Various additional data slots support all kinds of mechanisms that are not necessarily part of the core logic. For example, there would be data for canonical introduction and elimination rules for arbitrary operators (depending on the object-logic and application), which enables users to perform standard proof steps implicitly - (cf.\ the \rule\ method @{cite "isabelle-isar-ref"}). + (cf.\ the \rule\ method \<^cite>\"isabelle-isar-ref"\). \<^medskip> Thus Isabelle/Isar is able to bring forth more and more concepts successively. In particular, an object-logic like Isabelle/HOL continues the Isabelle/Pure setup by adding specific components for automated reasoning (classical reasoner, tableau prover, structured induction etc.) and derived specification mechanisms (inductive predicates, recursive functions etc.). All of this is ultimately based on the generic data management by theory and proof contexts introduced here. \ subsection \Theory context \label{sec:context-theory}\ text \ A \<^emph>\theory\ is a data container with explicit name and unique identifier. Theories are related by a (nominal) sub-theory relation, which corresponds to the dependency graph of the original construction; each theory is derived from a certain sub-graph of ancestor theories. To this end, the system maintains a set of symbolic ``identification stamps'' within each theory. The \begin\ operation starts a new theory by importing several parent theories (with merged contents) and entering a special mode of nameless incremental updates, until the final \end\ operation is performed. \<^medskip> The example in \figref{fig:ex-theory} below shows a theory graph derived from \Pure\, with theory \Length\ importing \Nat\ and \List\. The body of \Length\ consists of a sequence of updates, resulting in locally a linear sub-theory relation for each intermediate step. \begin{figure}[htb] \begin{center} \begin{tabular}{rcccl} & & \Pure\ \\ & & \\\ \\ & & \FOL\ \\ & $\swarrow$ & & $\searrow$ & \\ \Nat\ & & & & \List\ \\ & $\searrow$ & & $\swarrow$ \\ & & \Length\ \\ & & \multicolumn{3}{l}{~~@{keyword "begin"}} \\ & & $\vdots$~~ \\ & & \multicolumn{3}{l}{~~@{command "end"}} \\ \end{tabular} \caption{A theory definition depending on ancestors}\label{fig:ex-theory} \end{center} \end{figure} \<^medskip> Derived formal entities may retain a reference to the background theory in order to indicate the formal context from which they were produced. This provides an immutable certificate of the background theory. \ text %mlref \ \begin{mldecls} @{define_ML_type theory} \\ @{define_ML Context.eq_thy: "theory * theory -> bool"} \\ @{define_ML Context.subthy: "theory * theory -> bool"} \\ @{define_ML Theory.begin_theory: "string * Position.T -> theory list -> theory"} \\ @{define_ML Theory.parents_of: "theory -> theory list"} \\ @{define_ML Theory.ancestors_of: "theory -> theory list"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\theory\ represents theory contexts. \<^descr> \<^ML>\Context.eq_thy\~\(thy\<^sub>1, thy\<^sub>2)\ check strict identity of two theories. \<^descr> \<^ML>\Context.subthy\~\(thy\<^sub>1, thy\<^sub>2)\ compares theories according to the intrinsic graph structure of the construction. This sub-theory relation is a nominal approximation of inclusion (\\\) of the corresponding content (according to the semantics of the ML modules that implement the data). \<^descr> \<^ML>\Theory.begin_theory\~\name parents\ constructs a new theory based on the given parents. This ML function is normally not invoked directly. \<^descr> \<^ML>\Theory.parents_of\~\thy\ returns the direct ancestors of \thy\. \<^descr> \<^ML>\Theory.ancestors_of\~\thy\ returns all ancestors of \thy\ (not including \thy\ itself). \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "theory"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "theory_context"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation theory} embedded? ; @@{ML_antiquotation theory_context} embedded \ \<^descr> \@{theory}\ refers to the background theory of the current context --- as abstract value. \<^descr> \@{theory A}\ refers to an explicitly named ancestor theory \A\ of the background theory of the current context --- as abstract value. \<^descr> \@{theory_context A}\ is similar to \@{theory A}\, but presents the result as initial \<^ML_type>\Proof.context\ (see also \<^ML>\Proof_Context.init_global\). \ subsection \Proof context \label{sec:context-proof}\ text \ A proof context is a container for pure data that refers to the theory from which it is derived. The \init\ operation creates a proof context from a given theory. There is an explicit \transfer\ operation to force resynchronization with updates to the background theory -- this is rarely required in practice. Entities derived in a proof context need to record logical requirements explicitly, since there is no separate context identification or symbolic inclusion as for theories. For example, hypotheses used in primitive derivations (cf.\ \secref{sec:thms}) are recorded separately within the sequent \\ \ \\, just to make double sure. Results could still leak into an alien proof context due to programming errors, but Isabelle/Isar includes some extra validity checks in critical positions, notably at the end of a sub-proof. Proof contexts may be manipulated arbitrarily, although the common discipline is to follow block structure as a mental model: a given context is extended consecutively, and results are exported back into the original context. Note that an Isar proof state models block-structured reasoning explicitly, using a stack of proof contexts internally. For various technical reasons, the background theory of an Isar proof state must not be changed while the proof is still under construction! \ text %mlref \ \begin{mldecls} @{define_ML_type Proof.context} \\ @{define_ML Proof_Context.init_global: "theory -> Proof.context"} \\ @{define_ML Proof_Context.theory_of: "Proof.context -> theory"} \\ @{define_ML Proof_Context.transfer: "theory -> Proof.context -> Proof.context"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Proof.context\ represents proof contexts. \<^descr> \<^ML>\Proof_Context.init_global\~\thy\ produces a proof context derived from \thy\, initializing all data. \<^descr> \<^ML>\Proof_Context.theory_of\~\ctxt\ selects the background theory from \ctxt\. \<^descr> \<^ML>\Proof_Context.transfer\~\thy ctxt\ promotes the background theory of \ctxt\ to the super theory \thy\. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "context"} & : & \ML_antiquotation\ \\ \end{matharray} \<^descr> \@{context}\ refers to \<^emph>\the\ context at compile-time --- as abstract value. Independently of (local) theory or proof mode, this always produces a meaningful result. This is probably the most common antiquotation in interactive experimentation with ML inside Isar. \ subsection \Generic contexts \label{sec:generic-context}\ text \ A generic context is the disjoint sum of either a theory or proof context. Occasionally, this enables uniform treatment of generic context data, typically extra-logical information. Operations on generic contexts include the usual injections, partial selections, and combinators for lifting operations on either component of the disjoint sum. Moreover, there are total operations \theory_of\ and \proof_of\ to convert a generic context into either kind: a theory can always be selected from the sum, while a proof context might have to be constructed by an ad-hoc \init\ operation, which incurs a small runtime overhead. \ text %mlref \ \begin{mldecls} @{define_ML_type Context.generic} \\ @{define_ML Context.theory_of: "Context.generic -> theory"} \\ @{define_ML Context.proof_of: "Context.generic -> Proof.context"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Context.generic\ is the direct sum of \<^ML_type>\theory\ and \<^ML_type>\Proof.context\, with the datatype constructors \<^ML>\Context.Theory\ and \<^ML>\Context.Proof\. \<^descr> \<^ML>\Context.theory_of\~\context\ always produces a theory from the generic \context\, using \<^ML>\Proof_Context.theory_of\ as required. \<^descr> \<^ML>\Context.proof_of\~\context\ always produces a proof context from the generic \context\, using \<^ML>\Proof_Context.init_global\ as required (note that this re-initializes the context data with each invocation). \ subsection \Context data \label{sec:context-data}\ text \ The main purpose of theory and proof contexts is to manage arbitrary (pure) data. New data types can be declared incrementally at compile time. There are separate declaration mechanisms for any of the three kinds of contexts: theory, proof, generic. \ paragraph \Theory data\ text \declarations need to implement the following ML signature: \<^medskip> \begin{tabular}{ll} \\ T\ & representing type \\ \\ empty: T\ & empty default value \\ \\ extend: T \ T\ & obsolete (identity function) \\ \\ merge: T \ T \ T\ & merge data \\ \end{tabular} \<^medskip> The \empty\ value acts as initial default for \<^emph>\any\ theory that does not declare actual data content; \extend\ is obsolete: it needs to be the identity function. The \merge\ operation needs to join the data from two theories in a conservative manner. The standard scheme for \merge (data\<^sub>1, data\<^sub>2)\ inserts those parts of \data\<^sub>2\ into \data\<^sub>1\ that are not yet present, while keeping the general order of things. The \<^ML>\Library.merge\ function on plain lists may serve as canonical template. Particularly note that shared parts of the data must not be duplicated by naive concatenation, or a theory graph that resembles a chain of diamonds would cause an exponential blowup! Sometimes, the data consists of a single item that cannot be ``merged'' in a sensible manner. Then the standard scheme degenerates to the projection to \data\<^sub>1\, ignoring \data\<^sub>2\ outright. \ paragraph \Proof context data\ text \declarations need to implement the following ML signature: \<^medskip> \begin{tabular}{ll} \\ T\ & representing type \\ \\ init: theory \ T\ & produce initial value \\ \end{tabular} \<^medskip> The \init\ operation is supposed to produce a pure value from the given background theory and should be somehow ``immediate''. Whenever a proof context is initialized, which happens frequently, the the system invokes the \init\ operation of \<^emph>\all\ theory data slots ever declared. This also means that one needs to be economic about the total number of proof data declarations in the system, i.e.\ each ML module should declare at most one, sometimes two data slots for its internal use. Repeated data declarations to simulate a record type should be avoided! \ paragraph \Generic data\ text \ provides a hybrid interface for both theory and proof data. The \init\ operation for proof contexts is predefined to select the current data value from the background theory. \<^bigskip> Any of the above data declarations over type \T\ result in an ML structure with the following signature: \<^medskip> \begin{tabular}{ll} \get: context \ T\ \\ \put: T \ context \ context\ \\ \map: (T \ T) \ context \ context\ \\ \end{tabular} \<^medskip> These other operations provide exclusive access for the particular kind of context (theory, proof, or generic context). This interface observes the ML discipline for types and scopes: there is no other way to access the corresponding data slot of a context. By keeping these operations private, an Isabelle/ML module may maintain abstract values authentically. \ text %mlref \ \begin{mldecls} @{define_ML_functor Theory_Data} \\ @{define_ML_functor Proof_Data} \\ @{define_ML_functor Generic_Data} \\ \end{mldecls} \<^descr> \<^ML_functor>\Theory_Data\\(spec)\ declares data for type \<^ML_type>\theory\ according to the specification provided as argument structure. The resulting structure provides data init and access operations as described above. \<^descr> \<^ML_functor>\Proof_Data\\(spec)\ is analogous to \<^ML_functor>\Theory_Data\ for type \<^ML_type>\Proof.context\. \<^descr> \<^ML_functor>\Generic_Data\\(spec)\ is analogous to \<^ML_functor>\Theory_Data\ for type \<^ML_type>\Context.generic\. \ text %mlex \ The following artificial example demonstrates theory data: we maintain a set of terms that are supposed to be wellformed wrt.\ the enclosing theory. The public interface is as follows: \ ML \ signature WELLFORMED_TERMS = sig val get: theory -> term list val add: term -> theory -> theory end; \ text \ The implementation uses private theory data internally, and only exposes an operation that involves explicit argument checking wrt.\ the given theory. \ ML \ structure Wellformed_Terms: WELLFORMED_TERMS = struct structure Terms = Theory_Data ( type T = term Ord_List.T; val empty = []; fun merge (ts1, ts2) = Ord_List.union Term_Ord.fast_term_ord ts1 ts2; ); val get = Terms.get; fun add raw_t thy = let val t = Sign.cert_term thy raw_t; in Terms.map (Ord_List.insert Term_Ord.fast_term_ord t) thy end; end; \ text \ Type \<^ML_type>\term Ord_List.T\ is used for reasonably efficient representation of a set of terms: all operations are linear in the number of stored elements. Here we assume that users of this module do not care about the declaration order, since that data structure forces its own arrangement of elements. Observe how the \<^ML_text>\merge\ operation joins the data slots of the two constituents: \<^ML>\Ord_List.union\ prevents duplication of common data from different branches, thus avoiding the danger of exponential blowup. Plain list append etc.\ must never be used for theory data merges! \<^medskip> Our intended invariant is achieved as follows: \<^enum> \<^ML>\Wellformed_Terms.add\ only admits terms that have passed the \<^ML>\Sign.cert_term\ check of the given theory at that point. \<^enum> Wellformedness in the sense of \<^ML>\Sign.cert_term\ is monotonic wrt.\ the sub-theory relation. So our data can move upwards in the hierarchy (via extension or merges), and maintain wellformedness without further checks. Note that all basic operations of the inference kernel (which includes \<^ML>\Sign.cert_term\) observe this monotonicity principle, but other user-space tools don't. For example, fully-featured type-inference via \<^ML>\Syntax.check_term\ (cf.\ \secref{sec:term-check}) is not necessarily monotonic wrt.\ the background theory, since constraints of term constants can be modified by later declarations, for example. In most cases, user-space context data does not have to take such invariants too seriously. The situation is different in the implementation of the inference kernel itself, which uses the very same data mechanisms for types, constants, axioms etc. \ subsection \Configuration options \label{sec:config-options}\ text \ A \<^emph>\configuration option\ is a named optional value of some basic type (Boolean, integer, string) that is stored in the context. It is a simple application of general context data (\secref{sec:context-data}) that is sufficiently common to justify customized setup, which includes some concrete declarations for end-users using existing notation for attributes (cf.\ \secref{sec:attributes}). For example, the predefined configuration option @{attribute show_types} controls output of explicit type constraints for variables in printed terms (cf.\ \secref{sec:read-print}). Its value can be modified within Isar text like this: \ experiment begin declare [[show_types = false]] \ \declaration within (local) theory context\ notepad begin note [[show_types = true]] \ \declaration within proof (forward mode)\ term x have "x = x" using [[show_types = false]] \ \declaration within proof (backward mode)\ .. end end text \ Configuration options that are not set explicitly hold a default value that can depend on the application context. This allows to retrieve the value from another slot within the context, or fall back on a global preference mechanism, for example. The operations to declare configuration options and get/map their values are modeled as direct replacements for historic global references, only that the context is made explicit. This allows easy configuration of tools, without relying on the execution order as required for old-style mutable references. \ text %mlref \ \begin{mldecls} @{define_ML Config.get: "Proof.context -> 'a Config.T -> 'a"} \\ @{define_ML Config.map: "'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context"} \\ @{define_ML Attrib.setup_config_bool: "binding -> (Context.generic -> bool) -> bool Config.T"} \\ @{define_ML Attrib.setup_config_int: "binding -> (Context.generic -> int) -> int Config.T"} \\ @{define_ML Attrib.setup_config_real: "binding -> (Context.generic -> real) -> real Config.T"} \\ @{define_ML Attrib.setup_config_string: "binding -> (Context.generic -> string) -> string Config.T"} \\ \end{mldecls} \<^descr> \<^ML>\Config.get\~\ctxt config\ gets the value of \config\ in the given context. \<^descr> \<^ML>\Config.map\~\config f ctxt\ updates the context by updating the value of \config\. \<^descr> \config =\~\<^ML>\Attrib.setup_config_bool\~\name default\ creates a named configuration option of type \<^ML_type>\bool\, with the given \default\ depending on the application context. The resulting \config\ can be used to get/map its value in a given context. There is an implicit update of the background theory that registers the option as attribute with some concrete syntax. \<^descr> \<^ML>\Attrib.config_int\, \<^ML>\Attrib.config_real\, and \<^ML>\Attrib.config_string\ work like \<^ML>\Attrib.config_bool\, but for types \<^ML_type>\int\ and \<^ML_type>\string\, respectively. \ text %mlex \ The following example shows how to declare and use a Boolean configuration option called \my_flag\ with constant default value \<^ML>\false\. \ ML \ val my_flag = Attrib.setup_config_bool \<^binding>\my_flag\ (K false) \ text \ Now the user can refer to @{attribute my_flag} in declarations, while ML tools can retrieve the current value from the context via \<^ML>\Config.get\. \ ML_val \\<^assert> (Config.get \<^context> my_flag = false)\ declare [[my_flag = true]] ML_val \\<^assert> (Config.get \<^context> my_flag = true)\ notepad begin { note [[my_flag = false]] ML_val \\<^assert> (Config.get \<^context> my_flag = false)\ } ML_val \\<^assert> (Config.get \<^context> my_flag = true)\ end text \ Here is another example involving ML type \<^ML_type>\real\ (floating-point numbers). \ ML \ val airspeed_velocity = Attrib.setup_config_real \<^binding>\airspeed_velocity\ (K 0.0) \ declare [[airspeed_velocity = 10]] declare [[airspeed_velocity = 9.9]] section \Names \label{sec:names}\ text \ In principle, a name is just a string, but there are various conventions for representing additional structure. For example, ``\Foo.bar.baz\'' is considered as a long name consisting of qualifier \Foo.bar\ and base name \baz\. The individual constituents of a name may have further substructure, e.g.\ the string ``\<^verbatim>\\\'' encodes as a single symbol (\secref{sec:symbols}). \<^medskip> Subsequently, we shall introduce specific categories of names. Roughly speaking these correspond to logical entities as follows: \<^item> Basic names (\secref{sec:basic-name}): free and bound variables. \<^item> Indexed names (\secref{sec:indexname}): schematic variables. \<^item> Long names (\secref{sec:long-name}): constants of any kind (type constructors, term constants, other concepts defined in user space). Such entities are typically managed via name spaces (\secref{sec:name-space}). \ subsection \Basic names \label{sec:basic-name}\ text \ A \<^emph>\basic name\ essentially consists of a single Isabelle identifier. There are conventions to mark separate classes of basic names, by attaching a suffix of underscores: one underscore means \<^emph>\internal name\, two underscores means \<^emph>\Skolem name\, three underscores means \<^emph>\internal Skolem name\. For example, the basic name \foo\ has the internal version \foo_\, with Skolem versions \foo__\ and \foo___\, respectively. These special versions provide copies of the basic name space, apart from anything that normally appears in the user text. For example, system generated variables in Isar proof contexts are usually marked as internal, which prevents mysterious names like \xaa\ to appear in human-readable text. \<^medskip> Manipulating binding scopes often requires on-the-fly renamings. A \<^emph>\name context\ contains a collection of already used names. The \declare\ operation adds names to the context. The \invents\ operation derives a number of fresh names from a given starting point. For example, the first three names derived from \a\ are \a\, \b\, \c\. The \variants\ operation produces fresh names by incrementing tentative names as base-26 numbers (with digits \a..z\) until all clashes are resolved. For example, name \foo\ results in variants \fooa\, \foob\, \fooc\, \dots, \fooaa\, \fooab\ etc.; each renaming step picks the next unused variant from this sequence. \ text %mlref \ \begin{mldecls} @{define_ML Name.internal: "string -> string"} \\ @{define_ML Name.skolem: "string -> string"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type Name.context} \\ @{define_ML Name.context: Name.context} \\ @{define_ML Name.declare: "string -> Name.context -> Name.context"} \\ @{define_ML Name.invent: "Name.context -> string -> int -> string list"} \\ @{define_ML Name.variant: "string -> Name.context -> string * Name.context"} \\ \end{mldecls} \begin{mldecls} @{define_ML Variable.names_of: "Proof.context -> Name.context"} \\ \end{mldecls} \<^descr> \<^ML>\Name.internal\~\name\ produces an internal name by adding one underscore. \<^descr> \<^ML>\Name.skolem\~\name\ produces a Skolem name by adding two underscores. \<^descr> Type \<^ML_type>\Name.context\ represents the context of already used names; the initial value is \<^ML>\Name.context\. \<^descr> \<^ML>\Name.declare\~\name\ enters a used name into the context. \<^descr> \<^ML>\Name.invent\~\context name n\ produces \n\ fresh names derived from \name\. \<^descr> \<^ML>\Name.variant\~\name context\ produces a fresh variant of \name\; the result is declared to the context. \<^descr> \<^ML>\Variable.names_of\~\ctxt\ retrieves the context of declared type and term variable names. Projecting a proof context down to a primitive name context is occasionally useful when invoking lower-level operations. Regular management of ``fresh variables'' is done by suitable operations of structure \<^ML_structure>\Variable\, which is also able to provide an official status of ``locally fixed variable'' within the logical environment (cf.\ \secref{sec:variables}). \ text %mlex \ The following simple examples demonstrate how to produce fresh names from the initial \<^ML>\Name.context\. \ ML_val \ val list1 = Name.invent Name.context "a" 5; \<^assert> (list1 = ["a", "b", "c", "d", "e"]); val list2 = #1 (fold_map Name.variant ["x", "x", "a", "a", "'a", "'a"] Name.context); \<^assert> (list2 = ["x", "xa", "a", "aa", "'a", "'aa"]); \ text \ \<^medskip> The same works relatively to the formal context as follows.\ experiment fixes a b c :: 'a begin ML_val \ val names = Variable.names_of \<^context>; val list1 = Name.invent names "a" 5; \<^assert> (list1 = ["d", "e", "f", "g", "h"]); val list2 = #1 (fold_map Name.variant ["x", "x", "a", "a", "'a", "'a"] names); \<^assert> (list2 = ["x", "xa", "aa", "ab", "'aa", "'ab"]); \ end subsection \Indexed names \label{sec:indexname}\ text \ An \<^emph>\indexed name\ (or \indexname\) is a pair of a basic name and a natural number. This representation allows efficient renaming by incrementing the second component only. The canonical way to rename two collections of indexnames apart from each other is this: determine the maximum index \maxidx\ of the first collection, then increment all indexes of the second collection by \maxidx + 1\; the maximum index of an empty collection is \-1\. Occasionally, basic names are injected into the same pair type of indexed names: then \(x, -1)\ is used to encode the basic name \x\. \<^medskip> Isabelle syntax observes the following rules for representing an indexname \(x, i)\ as a packed string: \<^item> \?x\ if \x\ does not end with a digit and \i = 0\, \<^item> \?xi\ if \x\ does not end with a digit, \<^item> \?x.i\ otherwise. Indexnames may acquire large index numbers after several maxidx shifts have been applied. Results are usually normalized towards \0\ at certain checkpoints, notably at the end of a proof. This works by producing variants of the corresponding basic name components. For example, the collection \?x1, ?x7, ?x42\ becomes \?x, ?xa, ?xb\. \ text %mlref \ \begin{mldecls} @{define_ML_type indexname = "string * int"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\indexname\ represents indexed names. This is an abbreviation for \<^ML_type>\string * int\. The second component is usually non-negative, except for situations where \(x, -1)\ is used to inject basic names into this type. Other negative indexes should not be used. \ subsection \Long names \label{sec:long-name}\ text \ A \<^emph>\long name\ consists of a sequence of non-empty name components. The packed representation uses a dot as separator, as in ``\A.b.c\''. The last component is called \<^emph>\base name\, the remaining prefix is called \<^emph>\qualifier\ (which may be empty). The qualifier can be understood as the access path to the named entity while passing through some nested block-structure, although our free-form long names do not really enforce any strict discipline. For example, an item named ``\A.b.c\'' may be understood as a local entity \c\, within a local structure \b\, within a global structure \A\. In practice, long names usually represent 1--3 levels of qualification. User ML code should not make any assumptions about the particular structure of long names! The empty name is commonly used as an indication of unnamed entities, or entities that are not entered into the corresponding name space, whenever this makes any sense. The basic operations on long names map empty names again to empty names. \ text %mlref \ \begin{mldecls} @{define_ML Long_Name.base_name: "string -> string"} \\ @{define_ML Long_Name.qualifier: "string -> string"} \\ @{define_ML Long_Name.append: "string -> string -> string"} \\ @{define_ML Long_Name.implode: "string list -> string"} \\ @{define_ML Long_Name.explode: "string -> string list"} \\ \end{mldecls} \<^descr> \<^ML>\Long_Name.base_name\~\name\ returns the base name of a long name. \<^descr> \<^ML>\Long_Name.qualifier\~\name\ returns the qualifier of a long name. \<^descr> \<^ML>\Long_Name.append\~\name\<^sub>1 name\<^sub>2\ appends two long names. \<^descr> \<^ML>\Long_Name.implode\~\names\ and \<^ML>\Long_Name.explode\~\name\ convert between the packed string representation and the explicit list form of long names. \ subsection \Name spaces \label{sec:name-space}\ text \ A \name space\ manages a collection of long names, together with a mapping between partially qualified external names and fully qualified internal names (in both directions). Note that the corresponding \intern\ and \extern\ operations are mostly used for parsing and printing only! The \declare\ operation augments a name space according to the accesses determined by a given binding, and a naming policy from the context. \<^medskip> A \binding\ specifies details about the prospective long name of a newly introduced formal entity. It consists of a base name, prefixes for qualification (separate ones for system infrastructure and user-space mechanisms), a slot for the original source position, and some additional flags. \<^medskip> A \naming\ provides some additional details for producing a long name from a binding. Normally, the naming is implicit in the theory or proof context. The \full\ operation (and its variants for different context types) produces a fully qualified internal name to be entered into a name space. The main equation of this ``chemical reaction'' when binding new entities in a context is as follows: \<^medskip> \begin{tabular}{l} \binding + naming \ long name + name space accesses\ \end{tabular} \<^bigskip> As a general principle, there is a separate name space for each kind of formal entity, e.g.\ fact, logical constant, type constructor, type class. It is usually clear from the occurrence in concrete syntax (or from the scope) which kind of entity a name refers to. For example, the very same name \c\ may be used uniformly for a constant, type constructor, and type class. There are common schemes to name derived entities systematically according to the name of the main logical entity involved, e.g.\ fact \c.intro\ for a canonical introduction rule related to constant \c\. This technique of mapping names from one space into another requires some care in order to avoid conflicts. In particular, theorem names derived from a type constructor or type class should get an additional suffix in addition to the usual qualification. This leads to the following conventions for derived names: \<^medskip> \begin{tabular}{ll} logical entity & fact name \\\hline constant \c\ & \c.intro\ \\ type \c\ & \c_type.intro\ \\ class \c\ & \c_class.intro\ \\ \end{tabular} \ text %mlref \ \begin{mldecls} @{define_ML_type binding} \\ @{define_ML Binding.empty: binding} \\ @{define_ML Binding.name: "string -> binding"} \\ @{define_ML Binding.qualify: "bool -> string -> binding -> binding"} \\ @{define_ML Binding.prefix: "bool -> string -> binding -> binding"} \\ @{define_ML Binding.concealed: "binding -> binding"} \\ @{define_ML Binding.print: "binding -> string"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type Name_Space.naming} \\ @{define_ML Name_Space.global_naming: Name_Space.naming} \\ @{define_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\ @{define_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type Name_Space.T} \\ @{define_ML Name_Space.empty: "string -> Name_Space.T"} \\ @{define_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\ @{define_ML Name_Space.declare: "Context.generic -> bool -> binding -> Name_Space.T -> string * Name_Space.T"} \\ @{define_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\ @{define_ML Name_Space.extern: "Proof.context -> Name_Space.T -> string -> string"} \\ @{define_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"} \end{mldecls} \<^descr> Type \<^ML_type>\binding\ represents the abstract concept of name bindings. \<^descr> \<^ML>\Binding.empty\ is the empty binding. \<^descr> \<^ML>\Binding.name\~\name\ produces a binding with base name \name\. Note that this lacks proper source position information; see also the ML antiquotation @{ML_antiquotation binding}. \<^descr> \<^ML>\Binding.qualify\~\mandatory name binding\ prefixes qualifier \name\ to \binding\. The \mandatory\ flag tells if this name component always needs to be given in name space accesses --- this is mostly \false\ in practice. Note that this part of qualification is typically used in derived specification mechanisms. \<^descr> \<^ML>\Binding.prefix\ is similar to \<^ML>\Binding.qualify\, but affects the system prefix. This part of extra qualification is typically used in the infrastructure for modular specifications, notably ``local theory targets'' (see also \chref{ch:local-theory}). \<^descr> \<^ML>\Binding.concealed\~\binding\ indicates that the binding shall refer to an entity that serves foundational purposes only. This flag helps to mark implementation details of specification mechanism etc. Other tools should not depend on the particulars of concealed entities (cf.\ \<^ML>\Name_Space.is_concealed\). \<^descr> \<^ML>\Binding.print\~\binding\ produces a string representation for human-readable output, together with some formal markup that might get used in GUI front-ends, for example. \<^descr> Type \<^ML_type>\Name_Space.naming\ represents the abstract concept of a naming policy. \<^descr> \<^ML>\Name_Space.global_naming\ is the default naming policy: it is global and lacks any path prefix. In a regular theory context this is augmented by a path prefix consisting of the theory name. \<^descr> \<^ML>\Name_Space.add_path\~\path naming\ augments the naming policy by extending its path component. \<^descr> \<^ML>\Name_Space.full_name\~\naming binding\ turns a name binding (usually a basic name) into the fully qualified internal name, according to the given naming policy. \<^descr> Type \<^ML_type>\Name_Space.T\ represents name spaces. \<^descr> \<^ML>\Name_Space.empty\~\kind\ and \<^ML>\Name_Space.merge\~\(space\<^sub>1, space\<^sub>2)\ are the canonical operations for maintaining name spaces according to theory data management (\secref{sec:context-data}); \kind\ is a formal comment to characterize the purpose of a name space. \<^descr> \<^ML>\Name_Space.declare\~\context strict binding space\ enters a name binding as fully qualified internal name into the name space, using the naming of the context. \<^descr> \<^ML>\Name_Space.intern\~\space name\ internalizes a (partially qualified) external name. This operation is mostly for parsing! Note that fully qualified names stemming from declarations are produced via \<^ML>\Name_Space.full_name\ and \<^ML>\Name_Space.declare\ (or their derivatives for \<^ML_type>\theory\ and \<^ML_type>\Proof.context\). \<^descr> \<^ML>\Name_Space.extern\~\ctxt space name\ externalizes a (fully qualified) internal name. This operation is mostly for printing! User code should not rely on the precise result too much. \<^descr> \<^ML>\Name_Space.is_concealed\~\space name\ indicates whether \name\ refers to a strictly private entity that other tools are supposed to ignore! \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "binding"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation binding} embedded \ \<^descr> \@{binding name}\ produces a binding with base name \name\ and the source position taken from the concrete syntax of this antiquotation. In many situations this is more appropriate than the more basic \<^ML>\Binding.name\ function. \ text %mlex \ The following example yields the source position of some concrete binding inlined into the text: \ ML_val \Binding.pos_of \<^binding>\here\\ text \ \<^medskip> That position can be also printed in a message as follows: \ ML_command \writeln ("Look here" ^ Position.here (Binding.pos_of \<^binding>\here\))\ text \ This illustrates a key virtue of formalized bindings as opposed to raw specifications of base names: the system can use this additional information for feedback given to the user (error messages etc.). \<^medskip> The following example refers to its source position directly, which is occasionally useful for experimentation and diagnostic purposes: \ ML_command \warning ("Look here" ^ Position.here \<^here>)\ end 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,483 +1,483 @@ (*: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 -> ((ctyp TVars.table * cterm 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"} \\ @{define_ML Assumption.export_term: "Proof.context -> Proof.context -> term -> term"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Assumption.export\ represents export rules, as a pair of functions \<^ML_type>\bool -> cterm list -> (thm -> thm) * (term -> term)\. The \<^ML_type>\bool\ argument 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 \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. \<^descr> \<^ML>\Assumption.export_term\~\inner outer t\ exports term \t\ from the \inner\ context back into the \outer\ one. This is analogous to \<^ML>\Assumption.export\, but only takes syntactical aspects of the context into account (such as locally specified variables as seen in @{command define} or @{command obtain}). \ 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 corresponding Isar + parameters. See also \<^cite>\"isabelle-isar-ref"\ for the corresponding Isar proof command @{command obtain}. 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/Implementation/Syntax.thy b/src/Doc/Implementation/Syntax.thy --- a/src/Doc/Implementation/Syntax.thy +++ b/src/Doc/Implementation/Syntax.thy @@ -1,259 +1,257 @@ (*:maxLineLen=78:*) theory Syntax imports Base begin chapter \Concrete syntax and type-checking\ text \ Pure \\\-calculus as introduced in \chref{ch:logic} is an adequate foundation for logical languages --- in the tradition of \<^emph>\higher-order abstract syntax\ --- but end-users require additional means for reading and printing of terms and types. This important add-on outside the logical core is called \<^emph>\inner syntax\ in Isabelle jargon, as opposed to the \<^emph>\outer - syntax\ of the theory and proof language @{cite "isabelle-isar-ref"}. + syntax\ of the theory and proof language \<^cite>\"isabelle-isar-ref"\. - For example, according to @{cite church40} quantifiers are represented as + For example, according to \<^cite>\church40\ quantifiers are represented as higher-order constants \All :: ('a \ bool) \ bool\ such that \All (\x::'a. B x)\ faithfully represents the idea that is displayed in Isabelle as \\x::'a. B x\ via @{keyword "binder"} notation. Moreover, type-inference in the style - of Hindley-Milner @{cite hindleymilner} (and extensions) enables users to + of Hindley-Milner \<^cite>\hindleymilner\ (and extensions) enables users to write \\x. B x\ concisely, when the type \'a\ is already clear from the context.\<^footnote>\Type-inference taken to the extreme can easily confuse users. Beginners often stumble over unexpectedly general types inferred by the system.\ \<^medskip> The main inner syntax operations are \<^emph>\read\ for parsing together with type-checking, and \<^emph>\pretty\ for formatted output. See also \secref{sec:read-print}. Furthermore, the input and output syntax layers are sub-divided into separate phases for \<^emph>\concrete syntax\ versus \<^emph>\abstract syntax\, see also \secref{sec:parse-unparse} and \secref{sec:term-check}, respectively. This results in the following decomposition of the main operations: \<^item> \read = parse; check\ \<^item> \pretty = uncheck; unparse\ For example, some specification package might thus intercept syntax processing at a well-defined stage after \parse\, to a augment the resulting pre-term before full type-reconstruction is performed by \check\. Note that the formal status of bound variables, versus free variables, versus constants must not be changed between these phases. \<^medskip> In general, \check\ and \uncheck\ operate simultaneously on a list of terms. This is particular important for type-checking, to reconstruct types for several terms of the same context and scope. In contrast, \parse\ and \unparse\ operate separately on single terms. There are analogous operations to read and print types, with the same sub-division into phases. \ section \Reading and pretty printing \label{sec:read-print}\ text \ Read and print operations are roughly dual to each other, such that for the user \s' = pretty (read s)\ looks similar to the original source text \s\, but the details depend on many side-conditions. There are also explicit options to control the removal of type information in the output. The default configuration routinely looses information, so \t' = read (pretty t)\ might fail, or produce a differently typed term, or a completely different term in the face of syntactic overloading. \ text %mlref \ \begin{mldecls} @{define_ML Syntax.read_typs: "Proof.context -> string list -> typ list"} \\ @{define_ML Syntax.read_terms: "Proof.context -> string list -> term list"} \\ @{define_ML Syntax.read_props: "Proof.context -> string list -> term list"} \\[0.5ex] @{define_ML Syntax.read_typ: "Proof.context -> string -> typ"} \\ @{define_ML Syntax.read_term: "Proof.context -> string -> term"} \\ @{define_ML Syntax.read_prop: "Proof.context -> string -> term"} \\[0.5ex] @{define_ML Syntax.pretty_typ: "Proof.context -> typ -> Pretty.T"} \\ @{define_ML Syntax.pretty_term: "Proof.context -> term -> Pretty.T"} \\ @{define_ML Syntax.string_of_typ: "Proof.context -> typ -> string"} \\ @{define_ML Syntax.string_of_term: "Proof.context -> term -> string"} \\ \end{mldecls} \<^descr> \<^ML>\Syntax.read_typs\~\ctxt strs\ parses and checks a simultaneous list of source strings as types of the logic. \<^descr> \<^ML>\Syntax.read_terms\~\ctxt strs\ parses and checks a simultaneous list of source strings as terms of the logic. Type-reconstruction puts all parsed terms into the same scope: types of free variables ultimately need to coincide. If particular type-constraints are required for some of the arguments, the read operations needs to be split into its parse and check phases. Then it is possible to use \<^ML>\Type.constraint\ on the intermediate pre-terms (\secref{sec:term-check}). \<^descr> \<^ML>\Syntax.read_props\~\ctxt strs\ parses and checks a simultaneous list of source strings as terms of the logic, with an implicit type-constraint for each argument to enforce type \<^typ>\prop\; this also affects the inner syntax for parsing. The remaining type-reconstruction works as for \<^ML>\Syntax.read_terms\. \<^descr> \<^ML>\Syntax.read_typ\, \<^ML>\Syntax.read_term\, \<^ML>\Syntax.read_prop\ are like the simultaneous versions, but operate on a single argument only. This convenient shorthand is adequate in situations where a single item in its own scope is processed. Do not use \<^ML>\map o Syntax.read_term\ where \<^ML>\Syntax.read_terms\ is actually intended! \<^descr> \<^ML>\Syntax.pretty_typ\~\ctxt T\ and \<^ML>\Syntax.pretty_term\~\ctxt t\ uncheck and pretty-print the given type or term, respectively. Although the uncheck phase acts on a simultaneous list as well, this is rarely used in practice, so only the singleton case is provided as combined pretty operation. There is no distinction of term vs.\ proposition. \<^descr> \<^ML>\Syntax.string_of_typ\ and \<^ML>\Syntax.string_of_term\ are convenient compositions of \<^ML>\Syntax.pretty_typ\ and \<^ML>\Syntax.pretty_term\ with \<^ML>\Pretty.string_of\ for output. The result may be concatenated with other strings, as long as there is no further formatting and line-breaking involved. \<^ML>\Syntax.read_term\, \<^ML>\Syntax.read_prop\, and \<^ML>\Syntax.string_of_term\ are the most important operations in practice. \<^medskip> Note that the string values that are passed in and out are annotated by the - system, to carry further markup that is relevant for the Prover IDE @{cite - "isabelle-jedit"}. User code should neither compose its own input strings, + system, to carry further markup that is relevant for the Prover IDE \<^cite>\"isabelle-jedit"\. User code should neither compose its own input strings, nor try to analyze the output strings. Conceptually this is an abstract datatype, encoded as concrete string for historical reasons. The standard way to provide the required position markup for input works via the outer syntax parser wrapper \<^ML>\Parse.inner_syntax\, which is already part of \<^ML>\Parse.typ\, \<^ML>\Parse.term\, \<^ML>\Parse.prop\. So a string obtained from one of the latter may be directly passed to the corresponding read operation: this yields PIDE markup of the input and precise positions for warning and error messages. \ section \Parsing and unparsing \label{sec:parse-unparse}\ text \ Parsing and unparsing converts between actual source text and a certain \<^emph>\pre-term\ format, where all bindings and scopes are already resolved faithfully. Thus the names of free variables or constants are determined in the sense of the logical context, but type information might be still missing. Pre-terms support an explicit language of \<^emph>\type constraints\ that may be augmented by user code to guide the later \<^emph>\check\ phase. Actual parsing is based on traditional lexical analysis and Earley parsing for arbitrary context-free grammars. The user can specify the grammar declaratively via mixfix annotations. Moreover, there are \<^emph>\syntax translations\ that can be augmented by the user, either declaratively via @{command translations} or programmatically via @{command - parse_translation}, @{command print_translation} @{cite - "isabelle-isar-ref"}. The final scope-resolution is performed by the system, + parse_translation}, @{command print_translation} \<^cite>\"isabelle-isar-ref"\. The final scope-resolution is performed by the system, according to name spaces for types, term variables and constants determined by the context. \ text %mlref \ \begin{mldecls} @{define_ML Syntax.parse_typ: "Proof.context -> string -> typ"} \\ @{define_ML Syntax.parse_term: "Proof.context -> string -> term"} \\ @{define_ML Syntax.parse_prop: "Proof.context -> string -> term"} \\[0.5ex] @{define_ML Syntax.unparse_typ: "Proof.context -> typ -> Pretty.T"} \\ @{define_ML Syntax.unparse_term: "Proof.context -> term -> Pretty.T"} \\ \end{mldecls} \<^descr> \<^ML>\Syntax.parse_typ\~\ctxt str\ parses a source string as pre-type that is ready to be used with subsequent check operations. \<^descr> \<^ML>\Syntax.parse_term\~\ctxt str\ parses a source string as pre-term that is ready to be used with subsequent check operations. \<^descr> \<^ML>\Syntax.parse_prop\~\ctxt str\ parses a source string as pre-term that is ready to be used with subsequent check operations. The inner syntax category is \<^typ>\prop\ and a suitable type-constraint is included to ensure that this information is observed in subsequent type reconstruction. \<^descr> \<^ML>\Syntax.unparse_typ\~\ctxt T\ unparses a type after uncheck operations, to turn it into a pretty tree. \<^descr> \<^ML>\Syntax.unparse_term\~\ctxt T\ unparses a term after uncheck operations, to turn it into a pretty tree. There is no distinction for propositions here. These operations always operate on a single item; use the combinator \<^ML>\map\ to apply them to a list. \ section \Checking and unchecking \label{sec:term-check}\ text \ These operations define the transition from pre-terms and fully-annotated terms in the sense of the logical core (\chref{ch:logic}). The \<^emph>\check\ phase is meant to subsume a variety of mechanisms in the manner of ``type-inference'' or ``type-reconstruction'' or ``type-improvement'', not just type-checking in the narrow sense. The \<^emph>\uncheck\ phase is roughly dual, it prunes type-information before pretty printing. A typical add-on for the check/uncheck syntax layer is the @{command - abbreviation} mechanism @{cite "isabelle-isar-ref"}. Here the user specifies + abbreviation} mechanism \<^cite>\"isabelle-isar-ref"\. Here the user specifies syntactic definitions that are managed by the system as polymorphic \let\ bindings. These are expanded during the \check\ phase, and contracted during the \uncheck\ phase, without affecting the type-assignment of the given terms. \<^medskip> The precise meaning of type checking depends on the context --- additional check/uncheck modules might be defined in user space. For example, the @{command class} command defines a context where \check\ treats certain type instances of overloaded constants according to the ``dictionary construction'' of its logical foundation. This involves ``type improvement'' (specialization of slightly too general types) and replacement - by certain locale parameters. See also @{cite "Haftmann-Wenzel:2009"}. + by certain locale parameters. See also \<^cite>\"Haftmann-Wenzel:2009"\. \ text %mlref \ \begin{mldecls} @{define_ML Syntax.check_typs: "Proof.context -> typ list -> typ list"} \\ @{define_ML Syntax.check_terms: "Proof.context -> term list -> term list"} \\ @{define_ML Syntax.check_props: "Proof.context -> term list -> term list"} \\[0.5ex] @{define_ML Syntax.uncheck_typs: "Proof.context -> typ list -> typ list"} \\ @{define_ML Syntax.uncheck_terms: "Proof.context -> term list -> term list"} \\ \end{mldecls} \<^descr> \<^ML>\Syntax.check_typs\~\ctxt Ts\ checks a simultaneous list of pre-types as types of the logic. Typically, this involves normalization of type synonyms. \<^descr> \<^ML>\Syntax.check_terms\~\ctxt ts\ checks a simultaneous list of pre-terms as terms of the logic. Typically, this involves type-inference and normalization term abbreviations. The types within the given terms are treated in the same way as for \<^ML>\Syntax.check_typs\. Applications sometimes need to check several types and terms together. The standard approach uses \<^ML>\Logic.mk_type\ to embed the language of types into that of terms; all arguments are appended into one list of terms that is checked; afterwards the type arguments are recovered with \<^ML>\Logic.dest_type\. \<^descr> \<^ML>\Syntax.check_props\~\ctxt ts\ checks a simultaneous list of pre-terms as terms of the logic, such that all terms are constrained by type \<^typ>\prop\. The remaining check operation works as \<^ML>\Syntax.check_terms\ above. \<^descr> \<^ML>\Syntax.uncheck_typs\~\ctxt Ts\ unchecks a simultaneous list of types of the logic, in preparation of pretty printing. \<^descr> \<^ML>\Syntax.uncheck_terms\~\ctxt ts\ unchecks a simultaneous list of terms of the logic, in preparation of pretty printing. There is no distinction for propositions here. These operations always operate simultaneously on a list; use the combinator \<^ML>\singleton\ to apply them to a single item. \ end diff --git a/src/Doc/Isar_Ref/Document_Preparation.thy b/src/Doc/Isar_Ref/Document_Preparation.thy --- a/src/Doc/Isar_Ref/Document_Preparation.thy +++ b/src/Doc/Isar_Ref/Document_Preparation.thy @@ -1,736 +1,735 @@ (*:maxLineLen=78:*) theory Document_Preparation imports Main Base begin chapter \Document preparation \label{ch:document-prep}\ text \ Isabelle/Isar provides a simple document preparation system based on {PDF-\LaTeX}, with support for hyperlinks and bookmarks within that format. This allows to produce papers, books, theses etc.\ from Isabelle theory sources. {\LaTeX} output is generated while processing a \<^emph>\session\ in batch mode, as - explained in the \<^emph>\The Isabelle System Manual\ @{cite "isabelle-system"}. + explained in the \<^emph>\The Isabelle System Manual\ \<^cite>\"isabelle-system"\. The main Isabelle tools to get started with document preparation are @{tool_ref mkroot} and @{tool_ref build}. - The classic Isabelle/HOL tutorial @{cite "isabelle-hol-book"} also explains + The classic Isabelle/HOL tutorial \<^cite>\"isabelle-hol-book"\ also explains some aspects of theory presentation. \ section \Markup commands \label{sec:markup}\ text \ \begin{matharray}{rcl} @{command_def "chapter"} & : & \any \ any\ \\ @{command_def "section"} & : & \any \ any\ \\ @{command_def "subsection"} & : & \any \ any\ \\ @{command_def "subsubsection"} & : & \any \ any\ \\ @{command_def "paragraph"} & : & \any \ any\ \\ @{command_def "subparagraph"} & : & \any \ any\ \\ @{command_def "text"} & : & \any \ any\ \\ @{command_def "txt"} & : & \any \ any\ \\ @{command_def "text_raw"} & : & \any \ any\ \\ \end{matharray} Markup commands provide a structured way to insert text into the document generated from a theory. Each markup command takes a single @{syntax text} argument, which is passed as argument to a corresponding {\LaTeX} macro. The default macros provided by \<^file>\~~/lib/texinputs/isabelle.sty\ can be redefined according to the needs of the underlying document and {\LaTeX} styles. Note that formal comments (\secref{sec:comments}) are similar to markup commands, but have a different status within Isabelle/Isar syntax. \<^rail>\ (@@{command chapter} | @@{command section} | @@{command subsection} | @@{command subsubsection} | @@{command paragraph} | @@{command subparagraph}) @{syntax text} ';'? | (@@{command text} | @@{command txt} | @@{command text_raw}) @{syntax text} \ \<^descr> @{command chapter}, @{command section}, @{command subsection} etc.\ mark section headings within the theory source. This works in any context, even before the initial @{command theory} command. The corresponding {\LaTeX} macros are \<^verbatim>\\isamarkupchapter\, \<^verbatim>\\isamarkupsection\, \<^verbatim>\\isamarkupsubsection\ etc.\ \<^descr> @{command text} and @{command txt} specify paragraphs of plain text. This corresponds to a {\LaTeX} environment \<^verbatim>\\begin{isamarkuptext}\ \\\ \<^verbatim>\\end{isamarkuptext}\ etc. \<^descr> @{command text_raw} is similar to @{command text}, but without any surrounding markup environment. This allows to inject arbitrary {\LaTeX} source into the generated document. All text passed to any of the above markup commands may refer to formal entities via \<^emph>\document antiquotations\, see also \secref{sec:antiq}. These are interpreted in the present theory or proof context. \<^medskip> The proof markup commands closely resemble those for theory specifications, but have a different formal status and produce different {\LaTeX} macros. \ section \Document antiquotations \label{sec:antiq}\ text \ \begin{matharray}{rcl} @{antiquotation_def "theory"} & : & \antiquotation\ \\ @{antiquotation_def "thm"} & : & \antiquotation\ \\ @{antiquotation_def "lemma"} & : & \antiquotation\ \\ @{antiquotation_def "prop"} & : & \antiquotation\ \\ @{antiquotation_def "term"} & : & \antiquotation\ \\ @{antiquotation_def term_type} & : & \antiquotation\ \\ @{antiquotation_def typeof} & : & \antiquotation\ \\ @{antiquotation_def const} & : & \antiquotation\ \\ @{antiquotation_def abbrev} & : & \antiquotation\ \\ @{antiquotation_def typ} & : & \antiquotation\ \\ @{antiquotation_def type} & : & \antiquotation\ \\ @{antiquotation_def class} & : & \antiquotation\ \\ @{antiquotation_def locale} & : & \antiquotation\ \\ @{antiquotation_def bundle} & : & \antiquotation\ \\ @{antiquotation_def "text"} & : & \antiquotation\ \\ @{antiquotation_def goals} & : & \antiquotation\ \\ @{antiquotation_def subgoals} & : & \antiquotation\ \\ @{antiquotation_def prf} & : & \antiquotation\ \\ @{antiquotation_def full_prf} & : & \antiquotation\ \\ @{antiquotation_def ML_text} & : & \antiquotation\ \\ @{antiquotation_def ML} & : & \antiquotation\ \\ @{antiquotation_def ML_def} & : & \antiquotation\ \\ @{antiquotation_def ML_ref} & : & \antiquotation\ \\ @{antiquotation_def ML_infix} & : & \antiquotation\ \\ @{antiquotation_def ML_infix_def} & : & \antiquotation\ \\ @{antiquotation_def ML_infix_ref} & : & \antiquotation\ \\ @{antiquotation_def ML_type} & : & \antiquotation\ \\ @{antiquotation_def ML_type_def} & : & \antiquotation\ \\ @{antiquotation_def ML_type_ref} & : & \antiquotation\ \\ @{antiquotation_def ML_structure} & : & \antiquotation\ \\ @{antiquotation_def ML_structure_def} & : & \antiquotation\ \\ @{antiquotation_def ML_structure_ref} & : & \antiquotation\ \\ @{antiquotation_def ML_functor} & : & \antiquotation\ \\ @{antiquotation_def ML_functor_def} & : & \antiquotation\ \\ @{antiquotation_def ML_functor_ref} & : & \antiquotation\ \\ @{antiquotation_def emph} & : & \antiquotation\ \\ @{antiquotation_def bold} & : & \antiquotation\ \\ @{antiquotation_def verbatim} & : & \antiquotation\ \\ @{antiquotation_def bash_function} & : & \antiquotation\ \\ @{antiquotation_def system_option} & : & \antiquotation\ \\ @{antiquotation_def session} & : & \antiquotation\ \\ @{antiquotation_def "file"} & : & \antiquotation\ \\ @{antiquotation_def "url"} & : & \antiquotation\ \\ @{antiquotation_def "cite"} & : & \antiquotation\ \\ @{command_def "print_antiquotations"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} The overall content of an Isabelle/Isar theory may alternate between formal and informal text. The main body consists of formal specification and proof commands, interspersed with markup commands (\secref{sec:markup}) or document comments (\secref{sec:comments}). The argument of markup commands quotes informal text to be printed in the resulting document, but may again refer to formal entities via \<^emph>\document antiquotations\. For example, embedding \<^verbatim>\@{term [show_types] "f x = a + x"}\ within a text block makes \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} appear in the final {\LaTeX} document. Antiquotations usually spare the author tedious typing of logical entities in full detail. Even more importantly, some degree of consistency-checking between the main body of formal text and its informal explanation is achieved, since terms and types appearing in antiquotations are checked within the current theory or proof context. \<^medskip> Antiquotations are in general written as \<^verbatim>\@{\\name\~\<^verbatim>\[\\options\\<^verbatim>\]\~\arguments\\<^verbatim>\}\. The short form \<^verbatim>\\\\<^verbatim>\<^\\name\\<^verbatim>\>\\\argument_content\\ (without surrounding \<^verbatim>\@{\\\\\<^verbatim>\}\) works for a single argument that is a cartouche. A cartouche without special decoration is equivalent to \<^verbatim>\\<^cartouche>\\\argument_content\\, which is equivalent to \<^verbatim>\@{cartouche\~\\argument_content\\\<^verbatim>\}\. The special name @{antiquotation_def cartouche} is defined in the context: Isabelle/Pure introduces that as an alias to @{antiquotation_ref text} (see below). Consequently, \\foo_bar + baz \ bazar\\ prints literal quasi-formal text (unchecked). A control symbol \<^verbatim>\\\\<^verbatim>\<^\\name\\<^verbatim>\>\ within the body text, but without a subsequent cartouche, is equivalent to \<^verbatim>\@{\\name\\<^verbatim>\}\. \begingroup \def\isasymcontrolstart{\isatt{\isacharbackslash\isacharless\isacharcircum}} \<^rail>\ @{syntax_def antiquotation}: '@{' antiquotation_body '}' | '\' @{syntax_ref name} '>' @{syntax_ref cartouche} | @{syntax_ref cartouche} ; options: '[' (option * ',') ']' ; option: @{syntax name} | @{syntax name} '=' @{syntax name} ; \ \endgroup Note that the syntax of antiquotations may \<^emph>\not\ include source comments \<^verbatim>\(*\~\\\~\<^verbatim>\*)\ nor verbatim text \<^verbatim>\{*\~\\\~\<^verbatim>\*}\. %% FIXME less monolithic presentation, move to individual sections!? \<^rail>\ @{syntax_def antiquotation_body}: (@@{antiquotation text} | @@{antiquotation cartouche} | @@{antiquotation theory_text}) options @{syntax text} | @@{antiquotation theory} options @{syntax embedded} | @@{antiquotation thm} options styles @{syntax thms} | @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? | @@{antiquotation prop} options styles @{syntax prop} | @@{antiquotation term} options styles @{syntax term} | @@{antiquotation (HOL) value} options styles @{syntax term} | @@{antiquotation term_type} options styles @{syntax term} | @@{antiquotation typeof} options styles @{syntax term} | @@{antiquotation const} options @{syntax term} | @@{antiquotation abbrev} options @{syntax term} | @@{antiquotation typ} options @{syntax type} | @@{antiquotation type} options @{syntax embedded} | @@{antiquotation class} options @{syntax embedded} | @@{antiquotation locale} options @{syntax embedded} | @@{antiquotation bundle} options @{syntax embedded} | (@@{antiquotation command} | @@{antiquotation method} | @@{antiquotation attribute}) options @{syntax name} ; @{syntax antiquotation}: @@{antiquotation goals} options | @@{antiquotation subgoals} options | @@{antiquotation prf} options @{syntax thms} | @@{antiquotation full_prf} options @{syntax thms} | @@{antiquotation ML_text} options @{syntax text} | @@{antiquotation ML} options @{syntax text} | @@{antiquotation ML_infix} options @{syntax text} | @@{antiquotation ML_type} options @{syntax typeargs} @{syntax text} | @@{antiquotation ML_structure} options @{syntax text} | @@{antiquotation ML_functor} options @{syntax text} | @@{antiquotation emph} options @{syntax text} | @@{antiquotation bold} options @{syntax text} | @@{antiquotation verbatim} options @{syntax text} | @@{antiquotation bash_function} options @{syntax embedded} | @@{antiquotation system_option} options @{syntax embedded} | @@{antiquotation session} options @{syntax embedded} | @@{antiquotation path} options @{syntax embedded} | @@{antiquotation "file"} options @{syntax name} | @@{antiquotation dir} options @{syntax name} | @@{antiquotation url} options @{syntax embedded} | (@@{antiquotation cite} | @@{antiquotation nocite} | @@{antiquotation citet} | @@{antiquotation citep}) @{syntax embedded} ; styles: '(' (style + ',') ')' ; style: (@{syntax name} +) ; @@{command print_antiquotations} ('!'?) \ \<^descr> \@{text s}\ prints uninterpreted source text \s\, i.e.\ inner syntax. This is particularly useful to print portions of text according to the Isabelle document style, without demanding well-formedness, e.g.\ small pieces of terms that should not be parsed or type-checked yet. It is also possible to write this in the short form \\s\\ without any further decoration. \<^descr> \@{theory_text s}\ prints uninterpreted theory source text \s\, i.e.\ outer syntax with command keywords and other tokens. \<^descr> \@{theory A}\ prints the session-qualified theory name \A\, which is guaranteed to refer to a valid ancestor theory in the current context. \<^descr> \@{thm a\<^sub>1 \ a\<^sub>n}\ prints theorems \a\<^sub>1 \ a\<^sub>n\. Full fact expressions are allowed here, including attributes (\secref{sec:syn-att}). \<^descr> \@{prop \}\ prints a well-typed proposition \\\. \<^descr> \@{lemma \ by m}\ proves a well-typed proposition \\\ by method \m\ and prints the original \\\. \<^descr> \@{term t}\ prints a well-typed term \t\. \<^descr> \@{value t}\ evaluates a term \t\ and prints its result, see also @{command_ref (HOL) value}. \<^descr> \@{term_type t}\ prints a well-typed term \t\ annotated with its type. \<^descr> \@{typeof t}\ prints the type of a well-typed term \t\. \<^descr> \@{const c}\ prints a logical or syntactic constant \c\. \<^descr> \@{abbrev c x\<^sub>1 \ x\<^sub>n}\ prints a constant abbreviation \c x\<^sub>1 \ x\<^sub>n \ rhs\ as defined in the current context. \<^descr> \@{typ \}\ prints a well-formed type \\\. \<^descr> \@{type \}\ prints a (logical or syntactic) type constructor \\\. \<^descr> \@{class c}\ prints a class \c\. \<^descr> \@{locale c}\ prints a locale \c\. \<^descr> \@{bundle c}\ prints a bundle \c\. \<^descr> \@{command name}\, \@{method name}\, \@{attribute name}\ print checked entities of the Isar language. \<^descr> \@{goals}\ prints the current \<^emph>\dynamic\ goal state. This is mainly for support of tactic-emulation scripts within Isar. Presentation of goal states does not conform to the idea of human-readable proof documents! When explaining proofs in detail it is usually better to spell out the reasoning via proper Isar proof commands, instead of peeking at the internal machine configuration. \<^descr> \@{subgoals}\ is similar to \@{goals}\, but does not print the main goal. \<^descr> \@{prf a\<^sub>1 \ a\<^sub>n}\ prints the (compact) proof terms corresponding to the theorems \a\<^sub>1 \ a\<^sub>n\. Note that this requires proof terms to be switched on for the current logic session. \<^descr> \@{full_prf a\<^sub>1 \ a\<^sub>n}\ is like \@{prf a\<^sub>1 \ a\<^sub>n}\, but prints the full proof terms, i.e.\ also displays information omitted in the compact proof term, which is denoted by ``\_\'' placeholders there. \<^descr> \@{ML_text s}\ prints ML text verbatim: only the token language is checked. \<^descr> \@{ML s}\, \@{ML_infix s}\, \@{ML_type s}\, \@{ML_structure s}\, and \@{ML_functor s}\ check text \s\ as ML value, infix operator, type, exception, structure, and functor respectively. The source is printed verbatim. The variants \@{ML_def s}\ and \@{ML_ref s}\ etc. maintain the document index: ``def'' means to make a bold entry, ``ref'' means to make a regular entry. There are two forms for type constructors, with or without separate type arguments: this impacts only the index entry. For example, \@{ML_type_ref \'a list\}\ makes an entry literally for ``\'a list\'' (sorted under the letter ``a''), but \@{ML_type_ref 'a \list\}\ makes an entry for the constructor name ``\list\''. \<^descr> \@{emph s}\ prints document source recursively, with {\LaTeX} markup \<^verbatim>\\emph{\\\\\<^verbatim>\}\. \<^descr> \@{bold s}\ prints document source recursively, with {\LaTeX} markup \<^verbatim>\\textbf{\\\\\<^verbatim>\}\. \<^descr> \@{verbatim s}\ prints uninterpreted source text literally as ASCII characters, using some type-writer font style. \<^descr> \@{bash_function name}\ prints the given GNU bash function verbatim. The - name is checked wrt.\ the Isabelle system environment @{cite - "isabelle-system"}. + name is checked wrt.\ the Isabelle system environment \<^cite>\"isabelle-system"\. \<^descr> \@{system_option name}\ prints the given system option verbatim. The name is checked wrt.\ cumulative \<^verbatim>\etc/options\ of all Isabelle components, notably \<^file>\~~/etc/options\. \<^descr> \@{session name}\ prints given session name verbatim. The name is checked wrt.\ the dependencies of the current session. \<^descr> \@{path name}\ prints the file-system path name verbatim. \<^descr> \@{file name}\ is like \@{path name}\, but ensures that \name\ refers to a plain file. \<^descr> \@{dir name}\ is like \@{path name}\, but ensures that \name\ refers to a directory. \<^descr> \@{url name}\ produces markup for the given URL, which results in an active hyperlink within the text. \<^descr> \@{cite arg}\ produces the Bib{\TeX} citation macro \<^verbatim>\\cite[...]{...}\ with its optional and mandatory argument. The analogous \<^verbatim>\\nocite\, and the \<^verbatim>\\citet\ and \<^verbatim>\\citep\ variants from the \<^verbatim>\natbib\ package\<^footnote>\\<^url>\https://ctan.org/pkg/natbib\\ are supported as well. The argument syntax is uniform for all variants and is usually presented in control-symbol-cartouche form: \\<^cite>\arg\\. The formal syntax of the nested argument language is defined as follows: \<^rail>\arg: (embedded @'in')? (name + 'and') \ (@'using' name)?\ Here the embedded text is free-form {\LaTeX}, which becomes the optional argument of the \<^verbatim>\\cite\ macro. The named items are Bib{\TeX} database entries and become the mandatory argument (separated by comma). The optional part ``\<^theory_text>\using name\'' specifies an alternative {\LaTeX} macro name. The validity of Bib{\TeX} database entries is only partially checked in Isabelle/PIDE, when the corresponding \<^verbatim>\.bib\ files are open. Ultimately, the \<^verbatim>\bibtex\ program will complain about bad input in final document preparation. \<^descr> @{command "print_antiquotations"} prints all document antiquotations that are defined in the current context; the ``\!\'' option indicates extra verbosity. \ subsection \Styled antiquotations\ text \ The antiquotations \thm\, \prop\ and \term\ admit an extra \<^emph>\style\ specification to modify the printed result. A style is specified by a name with a possibly empty number of arguments; multiple styles can be sequenced with commas. The following standard styles are available: \<^descr> \lhs\ extracts the first argument of any application form with at least two arguments --- typically meta-level or object-level equality, or any other binary relation. \<^descr> \rhs\ is like \lhs\, but extracts the second argument. \<^descr> \concl\ extracts the conclusion \C\ from a rule in Horn-clause normal form \A\<^sub>1 \ \ A\<^sub>n \ C\. \<^descr> \prem\ \n\ extract premise number \n\ from from a rule in Horn-clause normal form \A\<^sub>1 \ \ A\<^sub>n \ C\. \ subsection \General options\ text \ The following options are available to tune the printed output of antiquotations. Note that many of these coincide with system and configuration options of the same names. \<^descr> @{antiquotation_option_def show_types}~\= bool\ and @{antiquotation_option_def show_sorts}~\= bool\ control printing of explicit type and sort constraints. \<^descr> @{antiquotation_option_def show_structs}~\= bool\ controls printing of implicit structures. \<^descr> @{antiquotation_option_def show_abbrevs}~\= bool\ controls folding of abbreviations. \<^descr> @{antiquotation_option_def names_long}~\= bool\ forces names of types and constants etc.\ to be printed in their fully qualified internal form. \<^descr> @{antiquotation_option_def names_short}~\= bool\ forces names of types and constants etc.\ to be printed unqualified. Note that internalizing the output again in the current context may well yield a different result. \<^descr> @{antiquotation_option_def names_unique}~\= bool\ determines whether the printed version of qualified names should be made sufficiently long to avoid overlap with names declared further back. Set to \false\ for more concise output. \<^descr> @{antiquotation_option_def eta_contract}~\= bool\ prints terms in \\\-contracted form. \<^descr> @{antiquotation_option_def display}~\= bool\ indicates if the text is to be output as multi-line ``display material'', rather than a small piece of text without line breaks (which is the default). In this mode the embedded entities are printed in the same style as the main theory text. \<^descr> @{antiquotation_option_def break}~\= bool\ controls line breaks in non-display material. \<^descr> @{antiquotation_option_def cartouche}~\= bool\ indicates if the output should be delimited as cartouche. \<^descr> @{antiquotation_option_def quotes}~\= bool\ indicates if the output should be delimited via double quotes (option @{antiquotation_option cartouche} takes precedence). Note that the Isabelle {\LaTeX} styles may suppress quotes on their own account. \<^descr> @{antiquotation_option_def mode}~\= name\ adds \name\ to the print mode to be used for presentation. Note that the standard setup for {\LaTeX} output is already present by default, with mode ``\latex\''. \<^descr> @{antiquotation_option_def margin}~\= nat\ and @{antiquotation_option_def indent}~\= nat\ change the margin or indentation for pretty printing of display material. \<^descr> @{antiquotation_option_def goals_limit}~\= nat\ determines the maximum number of subgoals to be printed (for goal-based antiquotation). \<^descr> @{antiquotation_option_def source}~\= bool\ prints the original source text of the antiquotation arguments, rather than its internal representation. Note that formal checking of @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still enabled; use the @{antiquotation "text"} antiquotation for unchecked output. Regular \term\ and \typ\ antiquotations with \source = false\ involve a full round-trip from the original source to an internalized logical entity back to a source form, according to the syntax of the current context. Thus the printed output is not under direct control of the author, it may even fluctuate a bit as the underlying theory is changed later on. In contrast, @{antiquotation_option source}~\= true\ admits direct printing of the given source text, with the desirable well-formedness check in the background, but without modification of the printed text. Cartouche delimiters of the argument are stripped for antiquotations that are internally categorized as ``embedded''. \<^descr> @{antiquotation_option_def source_cartouche} is like @{antiquotation_option source}, but cartouche delimiters are always preserved in the output. For Boolean flags, ``\name = true\'' may be abbreviated as ``\name\''. All of the above flags are disabled by default, unless changed specifically for a logic session in the corresponding \<^verbatim>\ROOT\ file. \ section \Markdown-like text structure\ text \ The markup commands @{command_ref text}, @{command_ref txt}, @{command_ref text_raw} (\secref{sec:markup}) consist of plain text. Its internal structure consists of paragraphs and (nested) lists, using special Isabelle symbols and some rules for indentation and blank lines. This quasi-visual format resembles \<^emph>\Markdown\\<^footnote>\\<^url>\http://commonmark.org\\, but the full complexity of that notation is avoided. This is a summary of the main principles of minimal Markdown in Isabelle: \<^item> List items start with the following markers \<^descr>[itemize:] \<^verbatim>\\<^item>\ \<^descr>[enumerate:] \<^verbatim>\\<^enum>\ \<^descr>[description:] \<^verbatim>\\<^descr>\ \<^item> Adjacent list items with same indentation and same marker are grouped into a single list. \<^item> Singleton blank lines separate paragraphs. \<^item> Multiple blank lines escape from the current list hierarchy. Notable differences to official Markdown: \<^item> Indentation of list items needs to match exactly. \<^item> Indentation is unlimited (official Markdown interprets four spaces as block quote). \<^item> List items always consist of paragraphs --- there is no notion of ``tight'' list. \<^item> Section headings are expressed via Isar document markup commands (\secref{sec:markup}). \<^item> URLs, font styles, other special content is expressed via antiquotations (\secref{sec:antiq}), usually with proper nesting of sub-languages via text cartouches. \ section \Document markers and command tags \label{sec:document-markers}\ text \ \emph{Document markers} are formal comments of the form \\<^marker>\marker_body\\ (using the control symbol \<^verbatim>\\<^marker>\) and may occur anywhere within the outer syntax of a command: the inner syntax of a marker body resembles that for attributes (\secref{sec:syn-att}). In contrast, \emph{Command tags} may only occur after a command keyword and are treated as special markers as explained below. \<^rail>\ @{syntax_def marker}: '\<^marker>' @{syntax cartouche} ; @{syntax_def marker_body}: (@{syntax name} @{syntax args} * ',') ; @{syntax_def tags}: tag* ; tag: '%' (@{syntax short_ident} | @{syntax string}) \ Document markers are stripped from the document output, but surrounding white-space is preserved: e.g.\ a marker at the end of a line does not affect the subsequent line break. Markers operate within the semantic presentation context of a command, and may modify it to change the overall appearance of a command span (e.g.\ by adding tags). Each document marker has its own syntax defined in the theory context; the following markers are predefined in Isabelle/Pure: \<^rail>\ (@@{document_marker_def title} | @@{document_marker_def creator} | @@{document_marker_def contributor} | @@{document_marker_def date} | @@{document_marker_def license} | @@{document_marker_def description}) @{syntax embedded} ; @@{document_marker_def tag} (scope?) @{syntax name} ; scope: '(' ('proof' | 'command') ')' \ \<^descr> \\<^marker>\title arg\\, \\<^marker>\creator arg\\, \\<^marker>\contributor arg\\, \\<^marker>\date arg\\, \\<^marker>\license arg\\, and \\<^marker>\description arg\\ produce markup in the PIDE document, without any immediate effect on typesetting. This vocabulary is taken from the Dublin Core Metadata Initiative\<^footnote>\\<^url>\https://www.dublincore.org/specifications/dublin-core/dcmi-terms\\. The argument is an uninterpreted string, except for @{document_marker description}, which consists of words that are subject to spell-checking. \<^descr> \\<^marker>\tag name\\ updates the list of command tags in the presentation context: later declarations take precedence, so \\<^marker>\tag a, tag b, tag c\\ produces a reversed list. The default tags are given by the original \<^theory_text>\keywords\ declaration of a command, and the system option @{system_option_ref document_tags}. The optional \scope\ tells how far the tagging is applied to subsequent proof structure: ``\<^theory_text>\("proof")\'' means it applies to the following proof text, and ``\<^theory_text>\(command)\'' means it only applies to the current command. The default within a proof body is ``\<^theory_text>\("proof")\'', but for toplevel goal statements it is ``\<^theory_text>\(command)\''. Thus a \tag\ marker for \<^theory_text>\theorem\, \<^theory_text>\lemma\ etc. does \emph{not} affect its proof by default. An old-style command tag \<^verbatim>\%\\name\ is treated like a document marker \\<^marker>\tag (proof) name\\: the list of command tags precedes the list of document markers. The head of the resulting tags in the presentation context is turned into {\LaTeX} environments to modify the type-setting. The following tags are pre-declared for certain classes of commands, and serve as default markup for certain kinds of commands: \<^medskip> \begin{tabular}{ll} \document\ & document markup commands \\ \theory\ & theory begin/end \\ \proof\ & all proof commands \\ \ML\ & all commands involving ML code \\ \end{tabular} \<^medskip> - The Isabelle document preparation system @{cite "isabelle-system"} allows + The Isabelle document preparation system \<^cite>\"isabelle-system"\ allows tagged command regions to be presented specifically, e.g.\ to fold proof texts, or drop parts of the text completely. For example ``\<^theory_text>\by auto\~\\<^marker>\tag invisible\\'' causes that piece of proof to be treated as \invisible\ instead of \proof\ (the default), which may be shown or hidden depending on the document setup. In contrast, ``\<^theory_text>\by auto\~\\<^marker>\tag visible\\'' forces this text to be shown invariably. Explicit tag specifications within a proof apply to all subsequent commands of the same level of nesting. For example, ``\<^theory_text>\proof\~\\<^marker>\tag invisible\ \\~\<^theory_text>\qed\'' forces the whole sub-proof to be typeset as \visible\ (unless some of its parts are tagged differently). \<^medskip> Command tags merely produce certain markup environments for type-setting. The meaning of these is determined by {\LaTeX} macros, as defined in \<^file>\~~/lib/texinputs/isabelle.sty\ or by the document author. The Isabelle document preparation tools also provide some high-level options to specify the meaning of arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding parts of the text. Logic sessions may also specify ``document versions'', where given tags are interpreted in some particular way. Again - see @{cite "isabelle-system"} for further details. + see \<^cite>\"isabelle-system"\ for further details. \ section \Railroad diagrams\ text \ \begin{matharray}{rcl} @{antiquotation_def "rail"} & : & \antiquotation\ \\ \end{matharray} \<^rail>\ 'rail' @{syntax text} \ The @{antiquotation rail} antiquotation allows to include syntax diagrams into Isabelle documents. {\LaTeX} requires the style file \<^file>\~~/lib/texinputs/railsetup.sty\, which can be used via \<^verbatim>\\usepackage{railsetup}\ in \<^verbatim>\root.tex\, for example. The rail specification language is quoted here as Isabelle @{syntax string} or text @{syntax "cartouche"}; it has its own grammar given below. \begingroup \def\isasymnewline{\isatt{\isacharbackslash\isacharless newline\isachargreater}} \<^rail>\ rule? + ';' ; rule: ((identifier | @{syntax antiquotation}) ':')? body ; body: concatenation + '|' ; concatenation: ((atom '?'?) +) (('*' | '+') atom?)? ; atom: '(' body? ')' | identifier | '@'? (string | @{syntax antiquotation}) | '\' \ \endgroup The lexical syntax of \identifier\ coincides with that of @{syntax short_ident} in regular Isabelle syntax, but \string\ uses single quotes instead of double quotes of the standard @{syntax string} category. Each \rule\ defines a formal language (with optional name), using a notation that is similar to EBNF or regular expressions with recursion. The meaning and visual appearance of these rail language elements is illustrated by the following representative examples. \<^item> Empty \<^verbatim>\()\ \<^rail>\()\ \<^item> Nonterminal \<^verbatim>\A\ \<^rail>\A\ \<^item> Nonterminal via Isabelle antiquotation \<^verbatim>\@{syntax method}\ \<^rail>\@{syntax method}\ \<^item> Terminal \<^verbatim>\'xyz'\ \<^rail>\'xyz'\ \<^item> Terminal in keyword style \<^verbatim>\@'xyz'\ \<^rail>\@'xyz'\ \<^item> Terminal via Isabelle antiquotation \<^verbatim>\@@{method rule}\ \<^rail>\@@{method rule}\ \<^item> Concatenation \<^verbatim>\A B C\ \<^rail>\A B C\ \<^item> Newline inside concatenation \<^verbatim>\A B C \ D E F\ \<^rail>\A B C \ D E F\ \<^item> Variants \<^verbatim>\A | B | C\ \<^rail>\A | B | C\ \<^item> Option \<^verbatim>\A ?\ \<^rail>\A ?\ \<^item> Repetition \<^verbatim>\A *\ \<^rail>\A *\ \<^item> Repetition with separator \<^verbatim>\A * sep\ \<^rail>\A * sep\ \<^item> Strict repetition \<^verbatim>\A +\ \<^rail>\A +\ \<^item> Strict repetition with separator \<^verbatim>\A + sep\ \<^rail>\A + sep\ \ end diff --git a/src/Doc/Isar_Ref/First_Order_Logic.thy b/src/Doc/Isar_Ref/First_Order_Logic.thy --- a/src/Doc/Isar_Ref/First_Order_Logic.thy +++ b/src/Doc/Isar_Ref/First_Order_Logic.thy @@ -1,498 +1,498 @@ (*:maxLineLen=78:*) section \Example: First-Order Logic\ theory %visible First_Order_Logic imports Base (* FIXME Pure!? *) begin text \ In order to commence a new object-logic within Isabelle/Pure we introduce abstract syntactic categories \i\ for individuals and \o\ for object-propositions. The latter is embedded into the language of Pure propositions by means of a separate judgment. \ typedecl i typedecl o judgment Trueprop :: "o \ prop" ("_" 5) text \ Note that the object-logic judgment is implicit in the syntax: writing \<^prop>\A\ produces \<^term>\Trueprop A\ internally. From the Pure perspective this means ``\<^prop>\A\ is derivable in the object-logic''. \ subsection \Equational reasoning \label{sec:framework-ex-equal}\ text \ Equality is axiomatized as a binary predicate on individuals, with reflexivity as introduction, and substitution as elimination principle. Note that the latter is particularly convenient in a framework like Isabelle, because syntactic congruences are implicitly produced by unification of \B x\ against expressions containing occurrences of \x\. \ axiomatization equal :: "i \ i \ o" (infix "=" 50) where refl [intro]: "x = x" and subst [elim]: "x = y \ B x \ B y" text \ Substitution is very powerful, but also hard to control in full generality. We derive some common symmetry~/ transitivity schemes of \<^term>\equal\ as particular consequences. \ theorem sym [sym]: assumes "x = y" shows "y = x" proof - have "x = x" .. with \x = y\ show "y = x" .. qed theorem forw_subst [trans]: assumes "y = x" and "B x" shows "B y" proof - from \y = x\ have "x = y" .. from this and \B x\ show "B y" .. qed theorem back_subst [trans]: assumes "B x" and "x = y" shows "B y" proof - from \x = y\ and \B x\ show "B y" .. qed theorem trans [trans]: assumes "x = y" and "y = z" shows "x = z" proof - from \y = z\ and \x = y\ show "x = z" .. qed subsection \Basic group theory\ text \ As an example for equational reasoning we consider some bits of group theory. The subsequent locale definition postulates group operations and axioms; we also derive some consequences of this specification. \ locale group = fixes prod :: "i \ i \ i" (infix "\" 70) and inv :: "i \ i" ("(_\)" [1000] 999) and unit :: i ("1") assumes assoc: "(x \ y) \ z = x \ (y \ z)" and left_unit: "1 \ x = x" and left_inv: "x\ \ x = 1" begin theorem right_inv: "x \ x\ = 1" proof - have "x \ x\ = 1 \ (x \ x\)" by (rule left_unit [symmetric]) also have "\ = (1 \ x) \ x\" by (rule assoc [symmetric]) also have "1 = (x\)\ \ x\" by (rule left_inv [symmetric]) also have "\ \ x = (x\)\ \ (x\ \ x)" by (rule assoc) also have "x\ \ x = 1" by (rule left_inv) also have "((x\)\ \ \) \ x\ = (x\)\ \ (1 \ x\)" by (rule assoc) also have "1 \ x\ = x\" by (rule left_unit) also have "(x\)\ \ \ = 1" by (rule left_inv) finally show "x \ x\ = 1" . qed theorem right_unit: "x \ 1 = x" proof - have "1 = x\ \ x" by (rule left_inv [symmetric]) also have "x \ \ = (x \ x\) \ x" by (rule assoc [symmetric]) also have "x \ x\ = 1" by (rule right_inv) also have "\ \ x = x" by (rule left_unit) finally show "x \ 1 = x" . qed text \ Reasoning from basic axioms is often tedious. Our proofs work by producing various instances of the given rules (potentially the symmetric form) using the pattern ``\<^theory_text>\have eq by (rule r)\'' and composing the chain of results via \<^theory_text>\also\/\<^theory_text>\finally\. These steps may involve any of the transitivity rules declared in \secref{sec:framework-ex-equal}, namely @{thm trans} in combining the first two results in @{thm right_inv} and in the final steps of both proofs, @{thm forw_subst} in the first combination of @{thm right_unit}, and @{thm back_subst} in all other calculational steps. Occasional substitutions in calculations are adequate, but should not be over-emphasized. The other extreme is to compose a chain by plain transitivity only, with replacements occurring always in topmost position. For example: \ (*<*) theorem "\A. PROP A \ PROP A" proof - assume [symmetric, defn]: "\x y. (x \ y) \ Trueprop (x = y)" fix x (*>*) have "x \ 1 = x \ (x\ \ x)" unfolding left_inv .. also have "\ = (x \ x\) \ x" unfolding assoc .. also have "\ = 1 \ x" unfolding right_inv .. also have "\ = x" unfolding left_unit .. finally have "x \ 1 = x" . (*<*) qed (*>*) text \ Here we have re-used the built-in mechanism for unfolding definitions in order to normalize each equational problem. A more realistic object-logic would include proper setup for the Simplifier (\secref{sec:simplifier}), the main automated tool for equational reasoning in Isabelle. Then ``\<^theory_text>\unfolding left_inv ..\'' would become ``\<^theory_text>\by (simp only: left_inv)\'' etc. \ end subsection \Propositional logic \label{sec:framework-ex-prop}\ text \ We axiomatize basic connectives of propositional logic: implication, disjunction, and conjunction. The associated rules are modeled after - Gentzen's system of Natural Deduction @{cite "Gentzen:1935"}. + Gentzen's system of Natural Deduction \<^cite>\"Gentzen:1935"\. \ axiomatization imp :: "o \ o \ o" (infixr "\" 25) where impI [intro]: "(A \ B) \ A \ B" and impD [dest]: "(A \ B) \ A \ B" axiomatization disj :: "o \ o \ o" (infixr "\" 30) where disjI\<^sub>1 [intro]: "A \ A \ B" and disjI\<^sub>2 [intro]: "B \ A \ B" and disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" axiomatization conj :: "o \ o \ o" (infixr "\" 35) where conjI [intro]: "A \ B \ A \ B" and conjD\<^sub>1: "A \ B \ A" and conjD\<^sub>2: "A \ B \ B" text \ The conjunctive destructions have the disadvantage that decomposing \<^prop>\A \ B\ involves an immediate decision which component should be projected. The more convenient simultaneous elimination \<^prop>\A \ B \ (A \ B \ C) \ C\ can be derived as follows: \ theorem conjE [elim]: assumes "A \ B" obtains A and B proof from \A \ B\ show A by (rule conjD\<^sub>1) from \A \ B\ show B by (rule conjD\<^sub>2) qed text \ Here is an example of swapping conjuncts with a single intermediate elimination step: \ (*<*) lemma "\A. PROP A \ PROP A" proof - fix A B (*>*) assume "A \ B" then obtain B and A .. then have "B \ A" .. (*<*) qed (*>*) text \ Note that the analogous elimination rule for disjunction ``\<^theory_text>\assumes "A \ B" obtains A \ B\'' coincides with the original axiomatization of @{thm disjE}. \<^medskip> We continue propositional logic by introducing absurdity with its characteristic elimination. Plain truth may then be defined as a proposition that is trivially true. \ axiomatization false :: o ("\") where falseE [elim]: "\ \ A" definition true :: o ("\") where "\ \ \ \ \" theorem trueI [intro]: \ unfolding true_def .. text \ \<^medskip> Now negation represents an implication towards absurdity: \ definition not :: "o \ o" ("\ _" [40] 40) where "\ A \ A \ \" theorem notI [intro]: assumes "A \ \" shows "\ A" unfolding not_def proof assume A then show \ by (rule \A \ \\) qed theorem notE [elim]: assumes "\ A" and A shows B proof - from \\ A\ have "A \ \" unfolding not_def . from \A \ \\ and \A\ have \ .. then show B .. qed subsection \Classical logic\ text \ Subsequently we state the principle of classical contradiction as a local assumption. Thus we refrain from forcing the object-logic into the classical perspective. Within that context, we may derive well-known consequences of the classical principle. \ locale classical = assumes classical: "(\ C \ C) \ C" begin theorem double_negation: assumes "\ \ C" shows C proof (rule classical) assume "\ C" with \\ \ C\ show C .. qed theorem tertium_non_datur: "C \ \ C" proof (rule double_negation) show "\ \ (C \ \ C)" proof assume "\ (C \ \ C)" have "\ C" proof assume C then have "C \ \ C" .. with \\ (C \ \ C)\ show \ .. qed then have "C \ \ C" .. with \\ (C \ \ C)\ show \ .. qed qed text \ These examples illustrate both classical reasoning and non-trivial propositional proofs in general. All three rules characterize classical logic independently, but the original rule is already the most convenient to use, because it leaves the conclusion unchanged. Note that \<^prop>\(\ C \ C) \ C\ fits again into our format for eliminations, despite the additional twist that the context refers to the main conclusion. So we may write @{thm classical} as the Isar statement ``\<^theory_text>\obtains \ thesis\''. This also explains nicely how classical reasoning really works: whatever the main \thesis\ might be, we may always assume its negation! \ end subsection \Quantifiers \label{sec:framework-ex-quant}\ text \ Representing quantifiers is easy, thanks to the higher-order nature of the underlying framework. According to the well-known technique introduced by - Church @{cite "church40"}, quantifiers are operators on predicates, which + Church \<^cite>\"church40"\, quantifiers are operators on predicates, which are syntactically represented as \\\-terms of type \<^typ>\i \ o\. Binder notation turns \All (\x. B x)\ into \\x. B x\ etc. \ axiomatization All :: "(i \ o) \ o" (binder "\" 10) where allI [intro]: "(\x. B x) \ \x. B x" and allD [dest]: "(\x. B x) \ B a" axiomatization Ex :: "(i \ o) \ o" (binder "\" 10) where exI [intro]: "B a \ (\x. B x)" and exE [elim]: "(\x. B x) \ (\x. B x \ C) \ C" text \ The statement of @{thm exE} corresponds to ``\<^theory_text>\assumes "\x. B x" obtains x where "B x"\'' in Isar. In the subsequent example we illustrate quantifier reasoning involving all four rules: \ theorem assumes "\x. \y. R x y" shows "\y. \x. R x y" proof \ \\\\ introduction\ obtain x where "\y. R x y" using \\x. \y. R x y\ .. \ \\\\ elimination\ fix y have "R x y" using \\y. R x y\ .. \ \\\\ destruction\ then show "\x. R x y" .. \ \\\\ introduction\ qed subsection \Canonical reasoning patterns\ text \ The main rules of first-order predicate logic from \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant} can now be summarized as follows, using the native Isar statement format of \secref{sec:framework-stmt}. \<^medskip> \begin{tabular}{l} \<^theory_text>\impI: assumes "A \ B" shows "A \ B"\ \\ \<^theory_text>\impD: assumes "A \ B" and A shows B\ \\[1ex] \<^theory_text>\disjI\<^sub>1: assumes A shows "A \ B"\ \\ \<^theory_text>\disjI\<^sub>2: assumes B shows "A \ B"\ \\ \<^theory_text>\disjE: assumes "A \ B" obtains A \ B\ \\[1ex] \<^theory_text>\conjI: assumes A and B shows A \ B\ \\ \<^theory_text>\conjE: assumes "A \ B" obtains A and B\ \\[1ex] \<^theory_text>\falseE: assumes \ shows A\ \\ \<^theory_text>\trueI: shows \\ \\[1ex] \<^theory_text>\notI: assumes "A \ \" shows "\ A"\ \\ \<^theory_text>\notE: assumes "\ A" and A shows B\ \\[1ex] \<^theory_text>\allI: assumes "\x. B x" shows "\x. B x"\ \\ \<^theory_text>\allE: assumes "\x. B x" shows "B a"\ \\[1ex] \<^theory_text>\exI: assumes "B a" shows "\x. B x"\ \\ \<^theory_text>\exE: assumes "\x. B x" obtains a where "B a"\ \end{tabular} \<^medskip> This essentially provides a declarative reading of Pure rules as Isar reasoning patterns: the rule statements tells how a canonical proof outline shall look like. Since the above rules have already been declared as @{attribute (Pure) intro}, @{attribute (Pure) elim}, @{attribute (Pure) dest} --- each according to its particular shape --- we can immediately write Isar proof texts as follows: \ (*<*) theorem "\A. PROP A \ PROP A" proof - (*>*) text_raw \\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "A \ B" proof assume A show B \ qed text_raw \\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "A \ B" and A \ then have B .. text_raw \\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have A \ then have "A \ B" .. have B \ then have "A \ B" .. text_raw \\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "A \ B" \ then have C proof assume A then show C \ next assume B then show C \ qed text_raw \\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have A and B \ then have "A \ B" .. text_raw \\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "A \ B" \ then obtain A and B .. text_raw \\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "\" \ then have A .. text_raw \\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "\" .. text_raw \\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "\ A" proof assume A then show "\" \ qed text_raw \\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "\ A" and A \ then have B .. text_raw \\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "\x. B x" proof fix x show "B x" \ qed text_raw \\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "\x. B x" \ then have "B a" .. text_raw \\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "\x. B x" proof show "B a" \ qed text_raw \\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\(*<*)next(*>*) have "\x. B x" \ then obtain a where "B a" .. text_raw \\end{minipage}\ (*<*) qed (*>*) text \ \<^bigskip> Of course, these proofs are merely examples. As sketched in \secref{sec:framework-subproof}, there is a fair amount of flexibility in expressing Pure deductions in Isar. Here the user is asked to express himself adequately, aiming at proof texts of literary quality. \ end %visible diff --git a/src/Doc/Isar_Ref/Framework.thy b/src/Doc/Isar_Ref/Framework.thy --- a/src/Doc/Isar_Ref/Framework.thy +++ b/src/Doc/Isar_Ref/Framework.thy @@ -1,1032 +1,1026 @@ (*:maxLineLen=78:*) theory Framework imports Main Base begin chapter \The Isabelle/Isar Framework \label{ch:isar-framework}\ text \ - Isabelle/Isar @{cite "Wenzel:1999:TPHOL" and "Wenzel-PhD" and + Isabelle/Isar \<^cite>\"Wenzel:1999:TPHOL" and "Wenzel-PhD" and "Nipkow-TYPES02" and "Wiedijk:1999:Mizar" and "Wenzel-Paulson:2006" and - "Wenzel:2006:Festschrift"} is a generic framework for developing formal + "Wenzel:2006:Festschrift"\ is a generic framework for developing formal mathematical documents with full proof checking. Definitions, statements and proofs are organized as theories. A collection of theories sources may be presented as a printed document; see also \chref{ch:document-prep}. The main concern of Isar is the design of a human-readable structured proof language, which is called the ``primary proof format'' in Isar terminology. Such a primary proof language is somewhere in the middle between the extremes of primitive proof objects and actual natural language. Thus Isar challenges the traditional way of recording informal proofs in mathematical prose, as well as the common tendency to see fully formal proofs directly as objects of some logical calculus (e.g.\ \\\-terms in a version of type theory). Technically, Isar is an interpreter of a simple block-structured language for describing the data flow of local facts and goals, interspersed with occasional invocations of proof methods. Everything is reduced to logical inferences internally, but these steps are somewhat marginal compared to the overall bookkeeping of the interpretation process. Thanks to careful design of the syntax and semantics of Isar language elements, a formal record of Isar commands may later appear as an intelligible text to the human reader. The Isar proof language has emerged from careful analysis of some inherent - virtues of the logical framework Isabelle/Pure @{cite "paulson-found" and - "paulson700"}, notably composition of higher-order natural deduction rules, - which is a generalization of Gentzen's original calculus @{cite - "Gentzen:1935"}. The approach of generic inference systems in Pure is + virtues of the logical framework Isabelle/Pure \<^cite>\"paulson-found" and + "paulson700"\, notably composition of higher-order natural deduction rules, + which is a generalization of Gentzen's original calculus \<^cite>\"Gentzen:1935"\. The approach of generic inference systems in Pure is continued by Isar towards actual proof texts. See also \figref{fig:natural-deduction} \begin{figure}[htb] \begin{center} \begin{minipage}[t]{0.9\textwidth} \textbf{Inferences:} \begin{center} \begin{tabular}{l@ {\qquad}l} \infer{\B\}{\A \ B\ & \A\} & \infer{\A \ B\}{\infer*{\B\}{\[A]\}} \\ \end{tabular} \end{center} \textbf{Isabelle/Pure:} \begin{center} \begin{tabular}{l@ {\qquad}l} \(A \ B) \ A \ B\ & \(A \ B) \ A \ B\ \end{tabular} \end{center} \textbf{Isabelle/Isar:} \begin{center} \begin{minipage}[t]{0.4\textwidth} @{theory_text [display, indent = 2] \have "A \ B" \ also have A \ finally have B .\} \end{minipage} \begin{minipage}[t]{0.4\textwidth} @{theory_text [display, indent = 2] \have "A \ B" proof assume A then show B \ qed\} \end{minipage} \end{center} \end{minipage} \end{center} \caption{Natural Deduction via inferences according to Gentzen, rules in Isabelle/Pure, and proofs in Isabelle/Isar}\label{fig:natural-deduction} \end{figure} \<^medskip> Concrete applications require another intermediate layer: an object-logic. - Isabelle/HOL @{cite "isa-tutorial"} (simply-typed set-theory) is most + Isabelle/HOL \<^cite>\"isa-tutorial"\ (simply-typed set-theory) is most commonly used; elementary examples are given in the directories \<^dir>\~~/src/Pure/Examples\ and \<^dir>\~~/src/HOL/Examples\. Some examples demonstrate how to start a fresh object-logic from Isabelle/Pure, and use Isar proofs from the very start, despite the lack of advanced proof tools at such an early stage (e.g.\ see - \<^file>\~~/src/Pure/Examples/Higher_Order_Logic.thy\). Isabelle/FOL @{cite - "isabelle-logics"} and Isabelle/ZF @{cite "isabelle-ZF"} also work, but are + \<^file>\~~/src/Pure/Examples/Higher_Order_Logic.thy\). Isabelle/FOL \<^cite>\"isabelle-logics"\ and Isabelle/ZF \<^cite>\"isabelle-ZF"\ also work, but are much less developed. In order to illustrate natural deduction in Isar, we shall subsequently refer to the background theory and library of Isabelle/HOL. This includes common notions of predicate logic, naive set-theory etc.\ using fairly standard mathematical notation. From the perspective of generic natural deduction there is nothing special about the logical connectives of HOL (\\\, \\\, \\\, \\\, etc.), only the resulting reasoning principles are relevant to the user. There are similar rules available for set-theory operators (\\\, \\\, \\\, \\\, etc.), or any other theory developed in the library (lattice theory, topology etc.). Subsequently we briefly review fragments of Isar proof texts corresponding directly to such general deduction schemes. The examples shall refer to set-theory, to minimize the danger of understanding connectives of predicate logic as something special. \<^medskip> The following deduction performs \\\-introduction, working forwards from assumptions towards the conclusion. We give both the Isar text, and depict the primitive rule involved, as determined by unification of fact and goal statements against rules that are declared in the library context. \ text_raw \\medskip\begin{minipage}{0.6\textwidth}\ (*<*) notepad begin fix x :: 'a and A B (*>*) assume "x \ A" and "x \ B" then have "x \ A \ B" .. (*<*) end (*>*) text_raw \\end{minipage}\begin{minipage}{0.4\textwidth}\ text \ \infer{\x \ A \ B\}{\x \ A\ & \x \ B\} \ text_raw \\end{minipage}\ text \ \<^medskip> Note that \<^theory_text>\assume\ augments the proof context, \<^theory_text>\then\ indicates that the current fact shall be used in the next step, and \<^theory_text>\have\ states an intermediate goal. The two dots ``\<^theory_text>\..\'' refer to a complete proof of this claim, using the indicated facts and a canonical rule from the context. We could have been more explicit here by spelling out the final proof step via the \<^theory_text>\by\ command: \ (*<*) notepad begin fix x :: 'a and A B (*>*) assume "x \ A" and "x \ B" then have "x \ A \ B" by (rule IntI) (*<*) end (*>*) text \ The format of the \\\-introduction rule represents the most basic inference, which proceeds from given premises to a conclusion, without any nested proof context involved. The next example performs backwards introduction of \\\\, the intersection of all sets within a given set. This requires a nested proof of set membership within a local context, where \A\ is an arbitrary-but-fixed member of the collection: \ text_raw \\medskip\begin{minipage}{0.6\textwidth}\ (*<*) notepad begin fix x :: 'a and \ (*>*) have "x \ \\" proof fix A assume "A \ \" show "x \ A" \ qed (*<*) end (*>*) text_raw \\end{minipage}\begin{minipage}{0.4\textwidth}\ text \ \infer{\x \ \\\}{\infer*{\x \ A\}{\[A][A \ \]\}} \ text_raw \\end{minipage}\ text \ \<^medskip> This Isar reasoning pattern again refers to the primitive rule depicted above. The system determines it in the ``\<^theory_text>\proof\'' step, which could have been spelled out more explicitly as ``\<^theory_text>\proof (rule InterI)\''. Note that the rule involves both a local parameter \A\ and an assumption \A \ \\ in the nested reasoning. Such compound rules typically demands a genuine subproof in Isar, working backwards rather than forwards as seen before. In the proof body we encounter the \<^theory_text>\fix\-\<^theory_text>\assume\-\<^theory_text>\show\ outline of nested subproofs that is typical for Isar. The final \<^theory_text>\show\ is like \<^theory_text>\have\ followed by an additional refinement of the enclosing claim, using the rule derived from the proof body. \<^medskip> The next example involves \\\\, which can be characterized as the set of all \x\ such that \\A. x \ A \ A \ \\. The elimination rule for \x \ \\\ does not mention \\\ and \\\ at all, but admits to obtain directly a local \A\ such that \x \ A\ and \A \ \\ hold. This corresponds to the following Isar proof and inference rule, respectively: \ text_raw \\medskip\begin{minipage}{0.6\textwidth}\ (*<*) notepad begin fix x :: 'a and \ C (*>*) assume "x \ \\" then have C proof fix A assume "x \ A" and "A \ \" show C \ qed (*<*) end (*>*) text_raw \\end{minipage}\begin{minipage}{0.4\textwidth}\ text \ \infer{\C\}{\x \ \\\ & \infer*{\C\~}{\[A][x \ A, A \ \]\}} \ text_raw \\end{minipage}\ text \ \<^medskip> Although the Isar proof follows the natural deduction rule closely, the text reads not as natural as anticipated. There is a double occurrence of an arbitrary conclusion \C\, which represents the final result, but is irrelevant for now. This issue arises for any elimination rule involving local parameters. Isar provides the derived language element \<^theory_text>\obtain\, which is able to perform the same elimination proof more conveniently: \ (*<*) notepad begin fix x :: 'a and \ (*>*) assume "x \ \\" then obtain A where "x \ A" and "A \ \" .. (*<*) end (*>*) text \ Here we avoid to mention the final conclusion \C\ and return to plain forward reasoning. The rule involved in the ``\<^theory_text>\..\'' proof is the same as before. \ section \The Pure framework \label{sec:framework-pure}\ text \ - The Pure logic @{cite "paulson-found" and "paulson700"} is an intuitionistic - fragment of higher-order logic @{cite "church40"}. In type-theoretic + The Pure logic \<^cite>\"paulson-found" and "paulson700"\ is an intuitionistic + fragment of higher-order logic \<^cite>\"church40"\. In type-theoretic parlance, there are three levels of \\\-calculus with corresponding arrows \\\/\\\/\\\: \<^medskip> \begin{tabular}{ll} \\ \ \\ & syntactic function space (terms depending on terms) \\ \\x. B(x)\ & universal quantification (proofs depending on terms) \\ \A \ B\ & implication (proofs depending on proofs) \\ \end{tabular} \<^medskip> Here only the types of syntactic terms, and the propositions of proof terms have been shown. The \\\-structure of proofs can be recorded as an optional - feature of the Pure inference kernel @{cite "Berghofer-Nipkow:2000:TPHOL"}, + feature of the Pure inference kernel \<^cite>\"Berghofer-Nipkow:2000:TPHOL"\, but the formal system can never depend on them due to \<^emph>\proof irrelevance\. On top of this most primitive layer of proofs, Pure implements a generic - calculus for nested natural deduction rules, similar to @{cite - "Schroeder-Heister:1984"}. Here object-logic inferences are internalized as + calculus for nested natural deduction rules, similar to \<^cite>\"Schroeder-Heister:1984"\. Here object-logic inferences are internalized as formulae over \\\ and \\\. Combining such rule statements may involve - higher-order unification @{cite "paulson-natural"}. + higher-order unification \<^cite>\"paulson-natural"\. \ subsection \Primitive inferences\ text \ Term syntax provides explicit notation for abstraction \\x :: \. b(x)\ and application \b a\, while types are usually implicit thanks to type-inference; terms of type \prop\ are called propositions. Logical statements are composed via \\x :: \. B(x)\ and \A \ B\. Primitive reasoning operates on judgments of the form \\ \ \\, with standard introduction and elimination rules for \\\ and \\\ that refer to fixed parameters \x\<^sub>1, \, x\<^sub>m\ and hypotheses \A\<^sub>1, \, A\<^sub>n\ from the context \\\; the corresponding proof terms are left implicit. The subsequent inference rules define \\ \ \\ inductively, relative to a collection of axioms from the implicit background theory: \[ \infer{\\ A\}{\A\ \mbox{~is axiom}} \qquad \infer{\A \ A\}{} \] \[ \infer{\\ \ \x. B(x)\}{\\ \ B(x)\ & \x \ \\} \qquad \infer{\\ \ B(a)\}{\\ \ \x. B(x)\} \] \[ \infer{\\ - A \ A \ B\}{\\ \ B\} \qquad \infer{\\\<^sub>1 \ \\<^sub>2 \ B\}{\\\<^sub>1 \ A \ B\ & \\\<^sub>2 \ A\} \] Furthermore, Pure provides a built-in equality \\ :: \ \ \ \ prop\ with axioms for reflexivity, substitution, extensionality, and \\\\\-conversion on \\\-terms. \<^medskip> An object-logic introduces another layer on top of Pure, e.g.\ with types \i\ for individuals and \o\ for propositions, term constants \Trueprop :: o \ prop\ as (implicit) derivability judgment and connectives like \\ :: o \ o \ o\ or \\ :: (i \ o) \ o\, and axioms for object-level rules such as \conjI: A \ B \ A \ B\ or \allI: (\x. B x) \ \x. B x\. Derived object rules are represented as theorems of Pure. After the initial object-logic setup, further axiomatizations are usually avoided: definitional principles are used instead (e.g.\ \<^theory_text>\definition\, \<^theory_text>\inductive\, \<^theory_text>\fun\, \<^theory_text>\function\). \ subsection \Reasoning with rules \label{sec:framework-resolution}\ text \ Primitive inferences mostly serve foundational purposes. The main reasoning mechanisms of Pure operate on nested natural deduction rules expressed as formulae, using \\\ to bind local parameters and \\\ to express entailment. Multiple parameters and premises are represented by repeating these connectives in a right-associative manner. Thanks to the Pure theorem \<^prop>\(A \ (\x. B x)) \ (\x. A \ B x)\ the connectives \\\ and \\\ commute. So we may assume w.l.o.g.\ that rule statements always observe the normal form where quantifiers are pulled in front of implications at each level of nesting. This means that any Pure - proposition may be presented as a \<^emph>\Hereditary Harrop Formula\ @{cite - "Miller:1991"} which is of the form \\x\<^sub>1 \ x\<^sub>m. H\<^sub>1 \ \ H\<^sub>n \ A\ for \m, n + proposition may be presented as a \<^emph>\Hereditary Harrop Formula\ \<^cite>\"Miller:1991"\ which is of the form \\x\<^sub>1 \ x\<^sub>m. H\<^sub>1 \ \ H\<^sub>n \ A\ for \m, n \ 0\, and \A\ atomic, and \H\<^sub>1, \, H\<^sub>n\ being recursively of the same format. Following the convention that outermost quantifiers are implicit, Horn clauses \A\<^sub>1 \ \ A\<^sub>n \ A\ are a special case of this. For example, the \\\-introduction rule encountered before is represented as a Pure theorem as follows: \[ \IntI:\~\<^prop>\x \ A \ x \ B \ x \ A \ B\ \] This is a plain Horn clause, since no further nesting on the left is involved. The general \\\-introduction corresponds to a Hereditary Harrop Formula with one additional level of nesting: \[ \InterI:\~\<^prop>\(\A. A \ \ \ x \ A) \ x \ \\\ \] \<^medskip> Goals are also represented as rules: \A\<^sub>1 \ \ A\<^sub>n \ C\ states that the subgoals \A\<^sub>1, \, A\<^sub>n\ entail the result \C\; for \n = 0\ the goal is finished. To allow \C\ being a rule statement itself, there is an internal protective marker \# :: prop \ prop\, which is defined as identity and hidden from the user. We initialize and finish goal states as follows: \[ \begin{array}{c@ {\qquad}c} \infer[(@{inference_def init})]{\C \ #C\}{} & \infer[(@{inference_def finish})]{\C\}{\#C\} \end{array} \] Goal states are refined in intermediate proof steps until a finished form is achieved. Here the two main reasoning principles are @{inference resolution}, for back-chaining a rule against a subgoal (replacing it by zero or more subgoals), and @{inference assumption}, for solving a subgoal (finding a short-circuit with local assumptions). Below \\<^vec>x\ stands for \x\<^sub>1, \, x\<^sub>n\ (for \n \ 0\). \[ \infer[(@{inference_def resolution})] {\(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (\<^vec>a \<^vec>x))\ \ C\\} {\begin{tabular}{rl} \rule:\ & \\<^vec>A \<^vec>a \ B \<^vec>a\ \\ \goal:\ & \(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C\ \\ \goal unifier:\ & \(\\<^vec>x. B (\<^vec>a \<^vec>x))\ = B'\\ \\ \end{tabular}} \] \<^medskip> \[ \infer[(@{inference_def assumption})]{\C\\} {\begin{tabular}{rl} \goal:\ & \(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C\ \\ \assm unifier:\ & \A\ = H\<^sub>i\\~~\mbox{for some~\H\<^sub>i\} \\ \end{tabular}} \] The following trace illustrates goal-oriented reasoning in Isabelle/Pure: {\footnotesize \<^medskip> \begin{tabular}{r@ {\quad}l} \(A \ B \ B \ A) \ #(A \ B \ B \ A)\ & \(init)\ \\ \(A \ B \ B) \ (A \ B \ A) \ #\\ & \(resolution B \ A \ B \ A)\ \\ \(A \ B \ A \ B) \ (A \ B \ A) \ #\\ & \(resolution A \ B \ B)\ \\ \(A \ B \ A) \ #\\ & \(assumption)\ \\ \(A \ B \ A \ B) \ #\\ & \(resolution A \ B \ A)\ \\ \#\\ & \(assumption)\ \\ \A \ B \ B \ A\ & \(finish)\ \\ \end{tabular} \<^medskip> } Compositions of @{inference assumption} after @{inference resolution} occurs quite often, typically in elimination steps. Traditional Isabelle tactics accommodate this by a combined @{inference_def elim_resolution} principle. In contrast, Isar uses a combined refinement rule as follows:\footnote{For simplicity and clarity, the presentation ignores \<^emph>\weak premises\ as introduced via \<^theory_text>\presume\ or \<^theory_text>\show \ when\.} {\small \[ \infer[(@{inference refinement})] {\C\\} {\begin{tabular}{rl} \subgoal:\ & \(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C\ \\ \subproof:\ & \\<^vec>G \<^vec>a \ B \<^vec>a\ \quad for schematic \\<^vec>a\ \\ \concl unifier:\ & \(\\<^vec>x. B (\<^vec>a \<^vec>x))\ = B'\\ \\ \assm unifiers:\ & \(\\<^vec>x. G\<^sub>j (\<^vec>a \<^vec>x))\ = H\<^sub>i\\ \quad for each \G\<^sub>j\ some \H\<^sub>i\ \\ \end{tabular}} \]} Here the \subproof\ rule stems from the main \<^theory_text>\fix\-\<^theory_text>\assume\-\<^theory_text>\show\ outline of Isar (cf.\ \secref{sec:framework-subproof}): each assumption indicated in the text results in a marked premise \G\ above. Consequently, \<^theory_text>\fix\-\<^theory_text>\assume\-\<^theory_text>\show\ enables to fit the result of a subproof quite robustly into a pending subgoal, while maintaining a good measure of flexibility: the subproof only needs to fit modulo unification, and its assumptions may be a proper subset of the subgoal premises (see \secref{sec:framework-subproof}). \ section \The Isar proof language \label{sec:framework-isar}\ text \ Structured proofs are presented as high-level expressions for composing entities of Pure (propositions, facts, and goals). The Isar proof language allows to organize reasoning within the underlying rule calculus of Pure, but Isar is not another logical calculus. Isar merely imposes certain structure and policies on Pure inferences. The main grammar of the Isar proof language is given in \figref{fig:isar-syntax}. \begin{figure}[htb] \begin{center} \begin{tabular}{rcl} \main\ & \=\ & \<^theory_text>\notepad begin "statement\<^sup>*" end\ \\ & \|\ & \<^theory_text>\theorem name: props if name: props for vars\ \\ & \|\ & \<^theory_text>\theorem name:\ \\ & & \quad\<^theory_text>\fixes vars\ \\ & & \quad\<^theory_text>\assumes name: props\ \\ & & \quad\<^theory_text>\shows name: props "proof"\ \\ & \|\ & \<^theory_text>\theorem name:\ \\ & & \quad\<^theory_text>\fixes vars\ \\ & & \quad\<^theory_text>\assumes name: props\ \\ & & \quad\<^theory_text>\obtains (name) clause "\<^bold>|" \ "proof"\ \\ \proof\ & \=\ & \<^theory_text>\"refinement\<^sup>* proper_proof"\ \\ \refinement\ & \=\ & \<^theory_text>\apply method\ \\ & \|\ & \<^theory_text>\supply name = thms\ \\ & \|\ & \<^theory_text>\subgoal premises name for vars "proof"\ \\ & \|\ & \<^theory_text>\using thms\ \\ & \|\ & \<^theory_text>\unfolding thms\ \\ \proper_proof\ & \=\ & \<^theory_text>\proof "method\<^sup>?" "statement\<^sup>*" qed "method\<^sup>?"\ \\ & \|\ & \<^theory_text>\by method method\~~~\|\~~~\<^theory_text>\..\~~~\|\~~~\<^theory_text>\.\~~~\|\~~~\<^theory_text>\sorry\~~~\|\~~~\<^theory_text>\done\ \\ \statement\ & \=\ & \<^theory_text>\{ "statement\<^sup>*" }\~~~\|\~~~\<^theory_text>\next\ \\ & \|\ & \<^theory_text>\note name = thms\ \\ & \|\ & \<^theory_text>\let "term" = "term"\ \\ & \|\ & \<^theory_text>\write name (mixfix)\ \\ & \|\ & \<^theory_text>\fix vars\ \\ & \|\ & \<^theory_text>\assume name: props if props for vars\ \\ & \|\ & \<^theory_text>\presume name: props if props for vars\ \\ & \|\ & \<^theory_text>\define clause\ \\ & \|\ & \<^theory_text>\case name: "case"\ \\ & \|\ & \<^theory_text>\then"\<^sup>?" goal\ \\ & \|\ & \<^theory_text>\from thms goal\ \\ & \|\ & \<^theory_text>\with thms goal\ \\ & \|\ & \<^theory_text>\also\~~~\|\~~~\<^theory_text>\finally goal\ \\ & \|\ & \<^theory_text>\moreover\~~~\|\~~~\<^theory_text>\ultimately goal\ \\ \goal\ & \=\ & \<^theory_text>\have name: props if name: props for vars "proof"\ \\ & \|\ & \<^theory_text>\show name: props if name: props for vars "proof"\ \\ & \|\ & \<^theory_text>\show name: props when name: props for vars "proof"\ \\ & \|\ & \<^theory_text>\consider (name) clause "\<^bold>|" \ "proof"\ \\ & \|\ & \<^theory_text>\obtain (name) clause "proof"\ \\ \clause\ & \=\ & \<^theory_text>\vars where name: props if props for vars\ \\ \end{tabular} \end{center} \caption{Main grammar of the Isar proof language}\label{fig:isar-syntax} \end{figure} The construction of the Isar proof language proceeds in a bottom-up fashion, as an exercise in purity and minimalism. The grammar in \appref{ap:main-grammar} describes the primitive parts of the core language (category \proof\), which is embedded into the main outer theory syntax via elements that require a proof (e.g.\ \<^theory_text>\theorem\, \<^theory_text>\lemma\, \<^theory_text>\function\, \<^theory_text>\termination\). The syntax for terms and propositions is inherited from Pure (and the object-logic). A \pattern\ is a \term\ with schematic variables, to be bound by higher-order matching. Simultaneous propositions or facts may be separated by the \<^theory_text>\and\ keyword. \<^medskip> Facts may be referenced by name or proposition. For example, the result of ``\<^theory_text>\have a: A \\'' becomes accessible both via the name \a\ and the literal proposition \\A\\. Moreover, fact expressions may involve attributes that modify either the theorem or the background context. For example, the expression ``\a [OF b]\'' refers to the composition of two facts according to the @{inference resolution} inference of \secref{sec:framework-resolution}, while ``\a [intro]\'' declares a fact as introduction rule in the context. The special fact called ``@{fact this}'' always refers to the last result, as produced by \<^theory_text>\note\, \<^theory_text>\assume\, \<^theory_text>\have\, or \<^theory_text>\show\. Since \<^theory_text>\note\ occurs frequently together with \<^theory_text>\then\, there are some abbreviations: \<^medskip> \begin{tabular}{rcl} \<^theory_text>\from a\ & \\\ & \<^theory_text>\note a then\ \\ \<^theory_text>\with a\ & \\\ & \<^theory_text>\from a and this\ \\ \end{tabular} \<^medskip> The \method\ category is essentially a parameter of the Isar language and may be populated later. The command \<^theory_text>\method_setup\ allows to define proof methods semantically in Isabelle/ML. The Eisbach language allows to define proof methods symbolically, as recursive expressions over existing methods - @{cite "Matichuk-et-al:2014"}; see also \<^dir>\~~/src/HOL/Eisbach\. + \<^cite>\"Matichuk-et-al:2014"\; see also \<^dir>\~~/src/HOL/Eisbach\. Methods use the facts indicated by \<^theory_text>\then\ or \<^theory_text>\using\, and then operate on the goal state. Some basic methods are predefined in Pure: ``@{method "-"}'' leaves the goal unchanged, ``@{method this}'' applies the facts as rules to the goal, ``@{method (Pure) "rule"}'' applies the facts to another rule and the result to the goal (both ``@{method this}'' and ``@{method (Pure) rule}'' refer to @{inference resolution} of \secref{sec:framework-resolution}). The secondary arguments to ``@{method (Pure) rule}'' may be specified explicitly as in ``\(rule a)\'', or picked from the context. In the latter case, the system first tries rules declared as @{attribute (Pure) elim} or @{attribute (Pure) dest}, followed by those declared as @{attribute (Pure) intro}. The default method for \<^theory_text>\proof\ is ``@{method standard}'' (which subsumes @{method rule} with arguments picked from the context), for \<^theory_text>\qed\ it is ``@{method "succeed"}''. Further abbreviations for terminal proof steps are ``\<^theory_text>\by method\<^sub>1 method\<^sub>2\'' for ``\<^theory_text>\proof method\<^sub>1 qed method\<^sub>2\'', and ``\<^theory_text>\..\'' for ``\<^theory_text>\by standard\, and ``\<^theory_text>\.\'' for ``\<^theory_text>\by this\''. The command ``\<^theory_text>\unfolding facts\'' operates directly on the goal by applying equalities. \<^medskip> Block structure can be indicated explicitly by ``\<^theory_text>\{ \ }\'', although the body of a subproof ``\<^theory_text>\proof \ qed\'' already provides implicit nesting. In both situations, \<^theory_text>\next\ jumps into the next section of a block, i.e.\ it acts like closing an implicit block scope and opening another one. There is no direct connection to subgoals here! The commands \<^theory_text>\fix\ and \<^theory_text>\assume\ build up a local context (see \secref{sec:framework-context}), while \<^theory_text>\show\ refines a pending subgoal by the rule resulting from a nested subproof (see \secref{sec:framework-subproof}). Further derived concepts will support calculational reasoning (see \secref{sec:framework-calc}). \ subsection \Context elements \label{sec:framework-context}\ text \ In judgments \\ \ \\ of the primitive framework, \\\ essentially acts like a proof context. Isar elaborates this idea towards a more advanced concept, with additional information for type-inference, term abbreviations, local facts, hypotheses etc. The element \<^theory_text>\fix x :: \\ declares a local parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in results exported from the context, \x\ may become anything. The \<^theory_text>\assume \inference\\ element provides a general interface to hypotheses: \<^theory_text>\assume \inference\ A\ produces \A \ A\ locally, while the included inference tells how to discharge \A\ from results \A \ B\ later on. There is no surface syntax for \\inference\\, i.e.\ it may only occur internally when derived commands are defined in ML. The default inference for \<^theory_text>\assume\ is @{inference export} as given below. The derived element \<^theory_text>\define x where "x \ a"\ is defined as \<^theory_text>\fix x assume \expand\ x \ a\, with the subsequent inference @{inference expand}. \[ \infer[(@{inference_def export})]{\\\ - A \ A \ B\}{\\\ \ B\} \qquad \infer[(@{inference_def expand})]{\\\ - (x \ a) \ B a\}{\\\ \ B x\} \] \<^medskip> - The most interesting derived context element in Isar is \<^theory_text>\obtain\ @{cite - \\S5.3\ "Wenzel-PhD"}, which supports generalized elimination steps in a + The most interesting derived context element in Isar is \<^theory_text>\obtain\ \<^cite>\\\S5.3\ in "Wenzel-PhD"\, which supports generalized elimination steps in a purely forward manner. The \<^theory_text>\obtain\ command takes a specification of parameters \\<^vec>x\ and assumptions \\<^vec>A\ to be added to the context, together with a proof of a case rule stating that this extension is conservative (i.e.\ may be removed from closed results later on): \<^medskip> \begin{tabular}{l} \<^theory_text>\\facts\ obtain "\<^vec>x" where "\<^vec>A" "\<^vec>x" \ \\ \\[0.5ex] \quad \<^theory_text>\have "case": "\thesis. (\\<^vec>x. \<^vec>A \<^vec>x \ thesis) \ thesis"\ \\ \quad \<^theory_text>\proof -\ \\ \qquad \<^theory_text>\fix thesis\ \\ \qquad \<^theory_text>\assume [intro]: "\\<^vec>x. \<^vec>A \<^vec>x \ thesis"\ \\ \qquad \<^theory_text>\show thesis using \facts\ \\ \\ \quad \<^theory_text>\qed\ \\ \quad \<^theory_text>\fix "\<^vec>x" assume \elimination "case"\ "\<^vec>A \<^vec>x"\ \\ \end{tabular} \<^medskip> \[ \infer[(@{inference elimination})]{\\ \ B\}{ \begin{tabular}{rl} \case:\ & \\ \ \thesis. (\\<^vec>x. \<^vec>A \<^vec>x \ thesis) \ thesis\ \\[0.2ex] \result:\ & \\ \ \<^vec>A \<^vec>y \ B\ \\[0.2ex] \end{tabular}} \] Here the name ``\thesis\'' is a specific convention for an arbitrary-but-fixed proposition; in the primitive natural deduction rules shown before we have occasionally used \C\. The whole statement of ``\<^theory_text>\obtain x where A x\'' can be read as a claim that \A x\ may be assumed for some arbitrary-but-fixed \x\. Also note that ``\<^theory_text>\obtain A and B\'' without parameters is similar to ``\<^theory_text>\have A and B\'', but the latter involves multiple subgoals that need to be proven separately. \<^medskip> The subsequent Isar proof texts explain all context elements introduced above using the formal proof language itself. After finishing a local proof within a block, the exported result is indicated via \<^theory_text>\note\. \ (*<*) theorem True proof (*>*) text_raw \\begin{minipage}[t]{0.45\textwidth}\ { fix x have "B x" \ } note \\x. B x\ text_raw \\end{minipage}\quad\begin{minipage}[t]{0.45\textwidth}\(*<*)next(*>*) { assume A have B \ } note \A \ B\ text_raw \\end{minipage}\\[3ex]\begin{minipage}[t]{0.45\textwidth}\(*<*)next(*>*) { define x where "x \ a" have "B x" \ } note \B a\ text_raw \\end{minipage}\quad\begin{minipage}[t]{0.45\textwidth}\(*<*)next(*>*) { obtain x where "A x" \ have B \ } note \B\ text_raw \\end{minipage}\ (*<*) qed (*>*) text \ \<^bigskip> This explains the meaning of Isar context elements without, without goal states getting in the way. \ subsection \Structured statements \label{sec:framework-stmt}\ text \ The syntax of top-level theorem statements is defined as follows: \<^medskip> \begin{tabular}{rcl} \statement\ & \\\ & \<^theory_text>\name: props and \\ \\ & \|\ & \context\<^sup>* conclusion\ \\[0.5ex] \context\ & \\\ & \<^theory_text>\fixes vars and \\ \\ & \|\ & \<^theory_text>\assumes name: props and \\ \\ \conclusion\ & \\\ & \<^theory_text>\shows name: props and \\ \\ & \|\ & \<^theory_text>\obtains vars and \ where name: props and \\ \\ & & \quad \\ \\ \\ \end{tabular} \<^medskip> A simple statement consists of named propositions. The full form admits local context elements followed by the actual conclusions, such as ``\<^theory_text>\fixes x assumes A x shows B x\''. The final result emerges as a Pure rule after discharging the context: \<^prop>\\x. A x \ B x\. The \<^theory_text>\obtains\ variant is another abbreviation defined below; unlike \<^theory_text>\obtain\ (cf.\ \secref{sec:framework-context}) there may be several ``cases'' separated by ``\\\'', each consisting of several parameters (\vars\) and several premises (\props\). This specifies multi-branch elimination rules. \<^medskip> \begin{tabular}{l} \<^theory_text>\obtains "\<^vec>x" where "\<^vec>A" "\<^vec>x" "\" \ \\ \\[0.5ex] \quad \<^theory_text>\fixes thesis\ \\ \quad \<^theory_text>\assumes [intro]: "\\<^vec>x. \<^vec>A \<^vec>x \ thesis" and \\ \\ \quad \<^theory_text>\shows thesis\ \\ \end{tabular} \<^medskip> Presenting structured statements in such an ``open'' format usually simplifies the subsequent proof, because the outer structure of the problem is already laid out directly. E.g.\ consider the following canonical patterns for \<^theory_text>\shows\ and \<^theory_text>\obtains\, respectively: \ text_raw \\begin{minipage}{0.5\textwidth}\ theorem fixes x and y assumes "A x" and "B y" shows "C x y" proof - from \A x\ and \B y\ show "C x y" \ qed text_raw \\end{minipage}\begin{minipage}{0.5\textwidth}\ theorem obtains x and y where "A x" and "B y" proof - have "A a" and "B b" \ then show thesis .. qed text_raw \\end{minipage}\ text \ \<^medskip> Here local facts \\A x\\ and \\B y\\ are referenced immediately; there is no need to decompose the logical rule structure again. In the second proof the final ``\<^theory_text>\then show thesis ..\'' involves the local rule case \\x y. A x \ B y \ thesis\ for the particular instance of terms \a\ and \b\ produced in the body. \ subsection \Structured proof refinement \label{sec:framework-subproof}\ text \ By breaking up the grammar for the Isar proof language, we may understand a proof text as a linear sequence of individual proof commands. These are interpreted as transitions of the Isar virtual machine (Isar/VM), which operates on a block-structured configuration in single steps. This allows users to write proof texts in an incremental manner, and inspect intermediate configurations for debugging. The basic idea is analogous to evaluating algebraic expressions on a stack machine: \(a + b) \ c\ then corresponds to a sequence of single transitions for each symbol \(, a, +, b, ), \, c\. In Isar the algebraic values are facts or goals, and the operations are inferences. \<^medskip> The Isar/VM state maintains a stack of nodes, each node contains the local proof context, the linguistic mode, and a pending goal (optional). The mode determines the type of transition that may be performed next, it essentially alternates between forward and backward reasoning, with an intermediate stage for chained facts (see \figref{fig:isar-vm}). \begin{figure}[htb] \begin{center} \includegraphics[width=0.8\textwidth]{isar-vm} \end{center} \caption{Isar/VM modes}\label{fig:isar-vm} \end{figure} For example, in \state\ mode Isar acts like a mathematical scratch-pad, accepting declarations like \<^theory_text>\fix\, \<^theory_text>\assume\, and claims like \<^theory_text>\have\, \<^theory_text>\show\. A goal statement changes the mode to \prove\, which means that we may now refine the problem via \<^theory_text>\unfolding\ or \<^theory_text>\proof\. Then we are again in \state\ mode of a proof body, which may issue \<^theory_text>\show\ statements to solve pending subgoals. A concluding \<^theory_text>\qed\ will return to the original \state\ mode one level upwards. The subsequent Isar/VM trace indicates block structure, linguistic mode, goal state, and inferences: \ text_raw \\begingroup\footnotesize\ (*<*)notepad begin (*>*) text_raw \\begin{minipage}[t]{0.18\textwidth}\ have "A \ B" proof assume A show B \ qed text_raw \\end{minipage}\quad \begin{minipage}[t]{0.06\textwidth} \begin\ \\ \\ \\ \begin\ \\ \end\ \\ \end\ \\ \end{minipage} \begin{minipage}[t]{0.08\textwidth} \prove\ \\ \state\ \\ \state\ \\ \prove\ \\ \state\ \\ \state\ \\ \end{minipage}\begin{minipage}[t]{0.35\textwidth} \(A \ B) \ #(A \ B)\ \\ \(A \ B) \ #(A \ B)\ \\ \\ \\ \#(A \ B)\ \\ \A \ B\ \\ \end{minipage}\begin{minipage}[t]{0.4\textwidth} \(init)\ \\ \(resolution impI)\ \\ \\ \\ \(refinement #A \ B)\ \\ \(finish)\ \\ \end{minipage}\ (*<*) end (*>*) text_raw \\endgroup\ text \ Here the @{inference refinement} inference from \secref{sec:framework-resolution} mediates composition of Isar subproofs nicely. Observe that this principle incorporates some degree of freedom in proof composition. In particular, the proof body allows parameters and assumptions to be re-ordered, or commuted according to Hereditary Harrop Form. Moreover, context elements that are not used in a subproof may be omitted altogether. For example: \ text_raw \\begin{minipage}{0.5\textwidth}\ (*<*) notepad begin (*>*) have "\x y. A x \ B y \ C x y" proof - fix x and y assume "A x" and "B y" show "C x y" \ qed text_raw \\end{minipage}\begin{minipage}{0.5\textwidth}\ (*<*) next (*>*) have "\x y. A x \ B y \ C x y" proof - fix x assume "A x" fix y assume "B y" show "C x y" \ qed text_raw \\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth}\ (*<*) next (*>*) have "\x y. A x \ B y \ C x y" proof - fix y assume "B y" fix x assume "A x" show "C x y" \ qed text_raw \\end{minipage}\begin{minipage}{0.5\textwidth}\ (*<*) next (*>*) have "\x y. A x \ B y \ C x y" proof - fix y assume "B y" fix x show "C x y" \ qed (*<*) end (*>*) text_raw \\end{minipage}\ text \ \<^medskip> Such fine-tuning of Isar text is practically important to improve readability. Contexts elements are rearranged according to the natural flow of reasoning in the body, while still observing the overall scoping rules. \<^medskip> This illustrates the basic idea of structured proof processing in Isar. The main mechanisms are based on natural deduction rule composition within the Pure framework. In particular, there are no direct operations on goal states within the proof body. Moreover, there is no hidden automated reasoning involved, just plain unification. \ subsection \Calculational reasoning \label{sec:framework-calc}\ text \ The existing Isar infrastructure is sufficiently flexible to support calculational reasoning (chains of transitivity steps) as derived concept. The generic proof elements introduced below depend on rules declared as @{attribute trans} in the context. It is left to the object-logic to provide a suitable rule collection for mixed relations of \=\, \<\, \\\, \\\, \\\ etc. Due to the flexibility of rule composition (\secref{sec:framework-resolution}), substitution of equals by equals is covered as well, even substitution of inequalities involving monotonicity - conditions; see also @{cite \\S6\ "Wenzel-PhD"} and @{cite - "Bauer-Wenzel:2001"}. + conditions; see also \<^cite>\\\S6\ in "Wenzel-PhD"\ and \<^cite>\"Bauer-Wenzel:2001"\. The generic calculational mechanism is based on the observation that rules such as \trans:\~\<^prop>\x = y \ y = z \ x = z\ proceed from the premises towards the conclusion in a deterministic fashion. Thus we may reason in forward mode, feeding intermediate results into rules selected from the context. The course of reasoning is organized by maintaining a secondary fact called ``@{fact calculation}'', apart from the primary ``@{fact this}'' already provided by the Isar primitives. In the definitions below, @{attribute OF} refers to @{inference resolution} (\secref{sec:framework-resolution}) with multiple rule arguments, and \trans\ represents to a suitable rule from the context: \begin{matharray}{rcl} \<^theory_text>\also"\<^sub>0"\ & \equiv & \<^theory_text>\note calculation = this\ \\ \<^theory_text>\also"\<^sub>n\<^sub>+\<^sub>1"\ & \equiv & \<^theory_text>\note calculation = trans [OF calculation this]\ \\[0.5ex] \<^theory_text>\finally\ & \equiv & \<^theory_text>\also from calculation\ \\ \end{matharray} The start of a calculation is determined implicitly in the text: here \<^theory_text>\also\ sets @{fact calculation} to the current result; any subsequent occurrence will update @{fact calculation} by combination with the next result and a transitivity rule. The calculational sequence is concluded via \<^theory_text>\finally\, where the final result is exposed for use in a concluding claim. Here is a canonical proof pattern, using \<^theory_text>\have\ to establish the intermediate results: \ (*<*) notepad begin fix a b c d :: 'a (*>*) have "a = b" \ also have "\ = c" \ also have "\ = d" \ finally have "a = d" . (*<*) end (*>*) text \ The term ``\\\'' (literal ellipsis) is a special abbreviation provided by the Isabelle/Isar term syntax: it statically refers to the right-hand side argument of the previous statement given in the text. Thus it happens to coincide with relevant sub-expressions in the calculational chain, but the exact correspondence is dependent on the transitivity rules being involved. \<^medskip> Symmetry rules such as \<^prop>\x = y \ y = x\ are like transitivities with only one premise. Isar maintains a separate rule collection declared via the @{attribute sym} attribute, to be used in fact expressions ``\a [symmetric]\'', or single-step proofs ``\<^theory_text>\assume "x = y" then have "y = x" ..\''. \ end diff --git a/src/Doc/Isar_Ref/Generic.thy b/src/Doc/Isar_Ref/Generic.thy --- a/src/Doc/Isar_Ref/Generic.thy +++ b/src/Doc/Isar_Ref/Generic.thy @@ -1,1827 +1,1826 @@ (*:maxLineLen=78:*) theory Generic imports Main Base begin chapter \Generic tools and packages \label{ch:gen-tools}\ section \Configuration options \label{sec:config}\ text \ Isabelle/Pure maintains a record of named configuration options within the theory or proof context, with values of type \<^ML_type>\bool\, \<^ML_type>\int\, \<^ML_type>\real\, or \<^ML_type>\string\. Tools may declare options in ML, and then refer to these values (relative to the context). Thus global reference variables are easily avoided. The user may change the value of a configuration option by means of an associated attribute of the same name. This form of context declaration works particularly well with commands such as @{command "declare"} or @{command "using"} like this: \ (*<*)experiment begin(*>*) declare [[show_main_goal = false]] notepad begin note [[show_main_goal = true]] end (*<*)end(*>*) text \ \begin{matharray}{rcll} @{command_def "print_options"} & : & \context \\ \\ \end{matharray} \<^rail>\ @@{command print_options} ('!'?) ; @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))? \ \<^descr> @{command "print_options"} prints the available configuration options, with names, types, and current values; the ``\!\'' option indicates extra verbosity. \<^descr> \name = value\ as an attribute expression modifies the named option, with the syntax of the value depending on the option's type. For \<^ML_type>\bool\ the default value is \true\. Any attempt to change a global option in a local context is ignored. \ section \Basic proof tools\ subsection \Miscellaneous methods and attributes \label{sec:misc-meth-att}\ text \ \begin{matharray}{rcl} @{method_def unfold} & : & \method\ \\ @{method_def fold} & : & \method\ \\ @{method_def insert} & : & \method\ \\[0.5ex] @{method_def erule}\\<^sup>*\ & : & \method\ \\ @{method_def drule}\\<^sup>*\ & : & \method\ \\ @{method_def frule}\\<^sup>*\ & : & \method\ \\ @{method_def intro} & : & \method\ \\ @{method_def elim} & : & \method\ \\ @{method_def fail} & : & \method\ \\ @{method_def succeed} & : & \method\ \\ @{method_def sleep} & : & \method\ \\ \end{matharray} \<^rail>\ (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thms} ; (@@{method erule} | @@{method drule} | @@{method frule}) ('(' @{syntax nat} ')')? @{syntax thms} ; (@@{method intro} | @@{method elim}) @{syntax thms}? ; @@{method sleep} @{syntax real} \ \<^descr> @{method unfold}~\a\<^sub>1 \ a\<^sub>n\ and @{method fold}~\a\<^sub>1 \ a\<^sub>n\ expand (or fold back) the given definitions throughout all goals; any chained facts provided are inserted into the goal and subject to rewriting as well. Unfolding works in two stages: first, the given equations are used directly for rewriting; second, the equations are passed through the attribute @{attribute_ref abs_def} before rewriting --- to ensure that definitions are fully expanded, regardless of the actual parameters that are provided. \<^descr> @{method insert}~\a\<^sub>1 \ a\<^sub>n\ inserts theorems as facts into all goals of the proof state. Note that current facts indicated for forward chaining are ignored. \<^descr> @{method erule}~\a\<^sub>1 \ a\<^sub>n\, @{method drule}~\a\<^sub>1 \ a\<^sub>n\, and @{method frule}~\a\<^sub>1 \ a\<^sub>n\ are similar to the basic @{method rule} method (see \secref{sec:pure-meth-att}), but apply rules by elim-resolution, - destruct-resolution, and forward-resolution, respectively @{cite - "isabelle-implementation"}. The optional natural number argument (default 0) + destruct-resolution, and forward-resolution, respectively \<^cite>\"isabelle-implementation"\. The optional natural number argument (default 0) specifies additional assumption steps to be performed here. Note that these methods are improper ones, mainly serving for experimentation and tactic script emulation. Different modes of basic rule application are usually expressed in Isar at the proof language level, rather than via implicit proof state manipulations. For example, a proper single-step elimination would be done using the plain @{method rule} method, with forward chaining of current facts. \<^descr> @{method intro} and @{method elim} repeatedly refine some goal by intro- or elim-resolution, after having inserted any chained facts. Exactly the rules given as arguments are taken into account; this allows fine-tuned decomposition of a proof problem, in contrast to common automated tools. \<^descr> @{method fail} yields an empty result sequence; it is the identity of the ``\|\'' method combinator (cf.\ \secref{sec:proof-meth}). \<^descr> @{method succeed} yields a single (unchanged) result; it is the identity of the ``\,\'' method combinator (cf.\ \secref{sec:proof-meth}). \<^descr> @{method sleep}~\s\ succeeds after a real-time delay of \s\ seconds. This is occasionally useful for demonstration and testing purposes. \begin{matharray}{rcl} @{attribute_def tagged} & : & \attribute\ \\ @{attribute_def untagged} & : & \attribute\ \\[0.5ex] @{attribute_def THEN} & : & \attribute\ \\ @{attribute_def unfolded} & : & \attribute\ \\ @{attribute_def folded} & : & \attribute\ \\ @{attribute_def abs_def} & : & \attribute\ \\[0.5ex] @{attribute_def rotated} & : & \attribute\ \\ @{attribute_def (Pure) elim_format} & : & \attribute\ \\ @{attribute_def no_vars}\\<^sup>*\ & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute tagged} @{syntax name} @{syntax name} ; @@{attribute untagged} @{syntax name} ; @@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thm} ; (@@{attribute unfolded} | @@{attribute folded}) @{syntax thms} ; @@{attribute rotated} @{syntax int}? \ \<^descr> @{attribute tagged}~\name value\ and @{attribute untagged}~\name\ add and remove \<^emph>\tags\ of some theorem. Tags may be any list of string pairs that serve as formal comment. The first string is considered the tag name, the second its value. Note that @{attribute untagged} removes any tags of the same name. \<^descr> @{attribute THEN}~\a\ composes rules by resolution; it resolves with the first premise of \a\ (an alternative position may be also specified). See - also \<^ML_infix>\RS\ in @{cite "isabelle-implementation"}. + also \<^ML_infix>\RS\ in \<^cite>\"isabelle-implementation"\. \<^descr> @{attribute unfolded}~\a\<^sub>1 \ a\<^sub>n\ and @{attribute folded}~\a\<^sub>1 \ a\<^sub>n\ expand and fold back again the given definitions throughout a rule. \<^descr> @{attribute abs_def} turns an equation of the form \<^prop>\f x y \ t\ into \<^prop>\f \ \x y. t\, which ensures that @{method simp} steps always expand it. This also works for object-logic equality. \<^descr> @{attribute rotated}~\n\ rotate the premises of a theorem by \n\ (default 1). \<^descr> @{attribute (Pure) elim_format} turns a destruction rule into elimination rule format, by resolving with the rule \<^prop>\PROP A \ (PROP A \ PROP B) \ PROP B\. Note that the Classical Reasoner (\secref{sec:classical}) provides its own version of this operation. \<^descr> @{attribute no_vars} replaces schematic variables by free ones; this is mainly for tuning output of pretty printed theorems. \ subsection \Low-level equational reasoning\ text \ \begin{matharray}{rcl} @{method_def subst} & : & \method\ \\ @{method_def hypsubst} & : & \method\ \\ @{method_def split} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method subst} ('(' 'asm' ')')? \ ('(' (@{syntax nat}+) ')')? @{syntax thm} ; @@{method split} @{syntax thms} \ These methods provide low-level facilities for equational reasoning that are intended for specialized applications only. Normally, single step calculations would be performed in a structured text (see also \secref{sec:calculation}), while the Simplifier methods provide the canonical way for automated normalization (see \secref{sec:simplifier}). \<^descr> @{method subst}~\eq\ performs a single substitution step using rule \eq\, which may be either a meta or object equality. \<^descr> @{method subst}~\(asm) eq\ substitutes in an assumption. \<^descr> @{method subst}~\(i \ j) eq\ performs several substitutions in the conclusion. The numbers \i\ to \j\ indicate the positions to substitute at. Positions are ordered from the top of the term tree moving down from left to right. For example, in \(a + b) + (c + d)\ there are three positions where commutativity of \+\ is applicable: 1 refers to \a + b\, 2 to the whole term, and 3 to \c + d\. If the positions in the list \(i \ j)\ are non-overlapping (e.g.\ \(2 3)\ in \(a + b) + (c + d)\) you may assume all substitutions are performed simultaneously. Otherwise the behaviour of \subst\ is not specified. \<^descr> @{method subst}~\(asm) (i \ j) eq\ performs the substitutions in the assumptions. The positions refer to the assumptions in order from left to right. For example, given in a goal of the form \P (a + b) \ P (c + d) \ \\, position 1 of commutativity of \+\ is the subterm \a + b\ and position 2 is the subterm \c + d\. \<^descr> @{method hypsubst} performs substitution using some assumption; this only works for equations of the form \x = t\ where \x\ is a free or bound variable. \<^descr> @{method split}~\a\<^sub>1 \ a\<^sub>n\ performs single-step case splitting using the given rules. Splitting is performed in the conclusion or some assumption of the subgoal, depending of the structure of the rule. Note that the @{method simp} method already involves repeated application of split rules as declared in the current context, using @{attribute split}, for example. \ section \The Simplifier \label{sec:simplifier}\ text \ The Simplifier performs conditional and unconditional rewriting and uses contextual information: rule declarations in the background theory or local proof context are taken into account, as well as chained facts and subgoal premises (``local assumptions''). There are several general hooks that allow to modify the simplification strategy, or incorporate other proof tools that solve sub-problems, produce rewrite rules on demand etc. The rewriting strategy is always strictly bottom up, except for congruence rules, which are applied while descending into a term. Conditions in conditional rewrite rules are solved recursively before the rewrite rule is applied. The default Simplifier setup of major object logics (HOL, HOLCF, FOL, ZF) makes the Simplifier ready for immediate use, without engaging into the internal structures. Thus it serves as general-purpose proof tool with the main focus on equational reasoning, and a bit more than that. \ subsection \Simplification methods \label{sec:simp-meth}\ text \ \begin{tabular}{rcll} @{method_def simp} & : & \method\ \\ @{method_def simp_all} & : & \method\ \\ \Pure.\@{method_def (Pure) simp} & : & \method\ \\ \Pure.\@{method_def (Pure) simp_all} & : & \method\ \\ @{attribute_def simp_depth_limit} & : & \attribute\ & default \100\ \\ \end{tabular} \<^medskip> \<^rail>\ (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * ) ; opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')' ; @{syntax_def simpmod}: ('add' | 'del' | 'flip' | 'only' | 'split' (() | '!' | 'del') | 'cong' (() | 'add' | 'del')) ':' @{syntax thms} \ \<^descr> @{method simp} invokes the Simplifier on the first subgoal, after inserting chained facts as additional goal premises; further rule declarations may be included via \(simp add: facts)\. The proof method fails if the subgoal remains unchanged after simplification. Note that the original goal premises and chained facts are subject to simplification themselves, while declarations via \add\/\del\ merely follow the policies of the object-logic to extract rewrite rules from theorems, without further simplification. This may lead to slightly different behavior in either case, which might be required precisely like that in some boundary situations to perform the intended simplification step! \<^medskip> Modifier \flip\ deletes the following theorems from the simpset and adds their symmetric version (i.e.\ lhs and rhs exchanged). No warning is shown if the original theorem was not present. \<^medskip> The \only\ modifier first removes all other rewrite rules, looper tactics (including split rules), congruence rules, and then behaves like \add\. Implicit solvers remain, which means that trivial rules like reflexivity or introduction of \True\ are available to solve the simplified subgoals, but also non-trivial tools like linear arithmetic in HOL. The latter may lead to some surprise of the meaning of ``only'' in Isabelle/HOL compared to English! \<^medskip> The \split\ modifiers add or delete rules for the Splitter (see also \secref{sec:simp-strategies} on the looper). This works only if the Simplifier method has been properly setup to include the Splitter (all major object logics such HOL, HOLCF, FOL, ZF do this already). The \!\ option causes the split rules to be used aggressively: after each application of a split rule in the conclusion, the \safe\ tactic of the classical reasoner (see \secref{sec:classical:partial}) is applied to the new goal. The net effect is that the goal is split into the different cases. This option can speed up simplification of goals with many nested conditional or case expressions significantly. There is also a separate @{method_ref split} method available for single-step case splitting. The effect of repeatedly applying \(split thms)\ can be imitated by ``\(simp only: split: thms)\''. \<^medskip> The \cong\ modifiers add or delete Simplifier congruence rules (see also \secref{sec:simp-rules}); the default is to add. \<^descr> @{method simp_all} is similar to @{method simp}, but acts on all goals, working backwards from the last to the first one as usual in Isabelle.\<^footnote>\The order is irrelevant for goals without schematic variables, so simplification might actually be performed in parallel here.\ Chained facts are inserted into all subgoals, before the simplification process starts. Further rule declarations are the same as for @{method simp}. The proof method fails if all subgoals remain unchanged after simplification. \<^descr> @{attribute simp_depth_limit} limits the number of recursive invocations of the Simplifier during conditional rewriting. By default the Simplifier methods above take local assumptions fully into account, using equational assumptions in the subsequent normalization process, or simplifying assumptions themselves. Further options allow to fine-tune the behavior of the Simplifier in this respect, corresponding to a variety of ML tactics as follows.\<^footnote>\Unlike the corresponding Isar proof methods, the ML tactics do not insist in changing the goal state.\ \begin{center} \small \begin{tabular}{|l|l|p{0.3\textwidth}|} \hline Isar method & ML tactic & behavior \\\hline \(simp (no_asm))\ & \<^ML>\simp_tac\ & assumptions are ignored completely \\\hline \(simp (no_asm_simp))\ & \<^ML>\asm_simp_tac\ & assumptions are used in the simplification of the conclusion but are not themselves simplified \\\hline \(simp (no_asm_use))\ & \<^ML>\full_simp_tac\ & assumptions are simplified but are not used in the simplification of each other or the conclusion \\\hline \(simp)\ & \<^ML>\asm_full_simp_tac\ & assumptions are used in the simplification of the conclusion and to simplify other assumptions \\\hline \(simp (asm_lr))\ & \<^ML>\asm_lr_simp_tac\ & compatibility mode: an assumption is only used for simplifying assumptions which are to the right of it \\\hline \end{tabular} \end{center} \<^medskip> In Isabelle/Pure, proof methods @{method (Pure) simp} and @{method (Pure) simp_all} only know about meta-equality \\\. Any new object-logic needs to re-define these methods via \<^ML>\Simplifier.method_setup\ in ML: Isabelle/FOL or Isabelle/HOL may serve as blue-prints. \ subsubsection \Examples\ text \ We consider basic algebraic simplifications in Isabelle/HOL. The rather trivial goal \<^prop>\0 + (x + 0) = x + 0 + 0\ looks like a good candidate to be solved by a single call of @{method simp}: \ lemma "0 + (x + 0) = x + 0 + 0" apply simp? oops text \ The above attempt \<^emph>\fails\, because \<^term>\0\ and \<^term>\(+)\ in the HOL library are declared as generic type class operations, without stating any algebraic laws yet. More specific types are required to get access to certain standard simplifications of the theory context, e.g.\ like this:\ lemma fixes x :: nat shows "0 + (x + 0) = x + 0 + 0" by simp lemma fixes x :: int shows "0 + (x + 0) = x + 0 + 0" by simp lemma fixes x :: "'a :: monoid_add" shows "0 + (x + 0) = x + 0 + 0" by simp text \ \<^medskip> In many cases, assumptions of a subgoal are also needed in the simplification process. For example: \ lemma fixes x :: nat shows "x = 0 \ x + x = 0" by simp lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" apply simp oops lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" using assms by simp text \ As seen above, local assumptions that shall contribute to simplification need to be part of the subgoal already, or indicated explicitly for use by the subsequent method invocation. Both too little or too much information can make simplification fail, for different reasons. In the next example the malicious assumption \<^prop>\\x::nat. f x = g (f (g x))\ does not contribute to solve the problem, but makes the default @{method simp} method loop: the rewrite rule \f ?x \ g (f (g ?x))\ extracted from the assumption does not terminate. The Simplifier notices certain simple forms of nontermination, but not this one. The problem can be solved nonetheless, by ignoring assumptions via special options as explained before: \ lemma "(\x::nat. f x = g (f (g x))) \ f 0 = f 0 + 0" by (simp (no_asm)) text \ The latter form is typical for long unstructured proof scripts, where the control over the goal content is limited. In structured proofs it is usually better to avoid pushing too many facts into the goal state in the first place. Assumptions in the Isar proof context do not intrude the reasoning if not used explicitly. This is illustrated for a toplevel statement and a local proof body as follows: \ lemma assumes "\x::nat. f x = g (f (g x))" shows "f 0 = f 0 + 0" by simp notepad begin assume "\x::nat. f x = g (f (g x))" have "f 0 = f 0 + 0" by simp end text \ \<^medskip> Because assumptions may simplify each other, there can be very subtle cases of nontermination. For example, the regular @{method simp} method applied to \<^prop>\P (f x) \ y = x \ f x = f y \ Q\ gives rise to the infinite reduction sequence \[ \P (f x)\ \stackrel{\f x \ f y\}{\longmapsto} \P (f y)\ \stackrel{\y \ x\}{\longmapsto} \P (f x)\ \stackrel{\f x \ f y\}{\longmapsto} \cdots \] whereas applying the same to \<^prop>\y = x \ f x = f y \ P (f x) \ Q\ terminates (without solving the goal): \ lemma "y = x \ f x = f y \ P (f x) \ Q" apply simp oops text \ See also \secref{sec:simp-trace} for options to enable Simplifier trace mode, which often helps to diagnose problems with rewrite systems. \ subsection \Declaring rules \label{sec:simp-rules}\ text \ \begin{matharray}{rcl} @{attribute_def simp} & : & \attribute\ \\ @{attribute_def split} & : & \attribute\ \\ @{attribute_def cong} & : & \attribute\ \\ @{command_def "print_simpset"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} \<^rail>\ (@@{attribute simp} | @@{attribute cong}) (() | 'add' | 'del') | @@{attribute split} (() | '!' | 'del') ; @@{command print_simpset} ('!'?) \ \<^descr> @{attribute simp} declares rewrite rules, by adding or deleting them from the simpset within the theory or proof context. Rewrite rules are theorems expressing some form of equality, for example: \Suc ?m + ?n = ?m + Suc ?n\ \\ \?P \ ?P \ ?P\ \\ \?A \ ?B \ {x. x \ ?A \ x \ ?B}\ \<^medskip> Conditional rewrites such as \?m < ?n \ ?m div ?n = 0\ are also permitted; the conditions can be arbitrary formulas. \<^medskip> Internally, all rewrite rules are translated into Pure equalities, theorems with conclusion \lhs \ rhs\. The simpset contains a function for extracting equalities from arbitrary theorems, which is usually installed when the object-logic is configured initially. For example, \\ ?x \ {}\ could be turned into \?x \ {} \ False\. Theorems that are declared as @{attribute simp} and local assumptions within a goal are treated uniformly in this respect. The Simplifier accepts the following formats for the \lhs\ term: \<^enum> First-order patterns, considering the sublanguage of application of constant operators to variable operands, without \\\-abstractions or functional variables. For example: \(?x + ?y) + ?z \ ?x + (?y + ?z)\ \\ \f (f ?x ?y) ?z \ f ?x (f ?y ?z)\ - \<^enum> Higher-order patterns in the sense of @{cite "nipkow-patterns"}. These + \<^enum> Higher-order patterns in the sense of \<^cite>\"nipkow-patterns"\. These are terms in \\\-normal form (this will always be the case unless you have done something strange) where each occurrence of an unknown is of the form \?F x\<^sub>1 \ x\<^sub>n\, where the \x\<^sub>i\ are distinct bound variables. For example, \(\x. ?P x \ ?Q x) \ (\x. ?P x) \ (\x. ?Q x)\ or its symmetric form, since the \rhs\ is also a higher-order pattern. \<^enum> Physical first-order patterns over raw \\\-term structure without \\\\\-equality; abstractions and bound variables are treated like quasi-constant term material. For example, the rule \?f ?x \ range ?f = True\ rewrites the term \g a \ range g\ to \True\, but will fail to match \g (h b) \ range (\x. g (h x))\. However, offending subterms (in our case \?f ?x\, which is not a pattern) can be replaced by adding new variables and conditions like this: \?y = ?f ?x \ ?y \ range ?f = True\ is acceptable as a conditional rewrite rule of the second category since conditions can be arbitrary terms. \<^descr> @{attribute split} declares case split rules. \<^descr> @{attribute cong} declares congruence rules to the Simplifier context. Congruence rules are equalities of the form @{text [display] "\ \ f ?x\<^sub>1 \ ?x\<^sub>n = f ?y\<^sub>1 \ ?y\<^sub>n"} This controls the simplification of the arguments of \f\. For example, some arguments can be simplified under additional assumptions: @{text [display] "?P\<^sub>1 \ ?Q\<^sub>1 \ (?Q\<^sub>1 \ ?P\<^sub>2 \ ?Q\<^sub>2) \ (?P\<^sub>1 \ ?P\<^sub>2) \ (?Q\<^sub>1 \ ?Q\<^sub>2)"} Given this rule, the Simplifier assumes \?Q\<^sub>1\ and extracts rewrite rules from it when simplifying \?P\<^sub>2\. Such local assumptions are effective for rewriting formulae such as \x = 0 \ y + x = y\. %FIXME %The local assumptions are also provided as theorems to the solver; %see \secref{sec:simp-solver} below. \<^medskip> The following congruence rule for bounded quantifiers also supplies contextual information --- about the bound variable: @{text [display] "(?A = ?B) \ (\x. x \ ?B \ ?P x \ ?Q x) \ (\x \ ?A. ?P x) \ (\x \ ?B. ?Q x)"} \<^medskip> This congruence rule for conditional expressions can supply contextual information for simplifying the arms: @{text [display] "?p = ?q \ (?q \ ?a = ?c) \ (\ ?q \ ?b = ?d) \ (if ?p then ?a else ?b) = (if ?q then ?c else ?d)"} A congruence rule can also \<^emph>\prevent\ simplification of some arguments. Here is an alternative congruence rule for conditional expressions that conforms to non-strict functional evaluation: @{text [display] "?p = ?q \ (if ?p then ?a else ?b) = (if ?q then ?a else ?b)"} Only the first argument is simplified; the others remain unchanged. This can make simplification much faster, but may require an extra case split over the condition \?q\ to prove the goal. \<^descr> @{command "print_simpset"} prints the collection of rules declared to the Simplifier, which is also known as ``simpset'' internally; the ``\!\'' option indicates extra verbosity. The implicit simpset of the theory context is propagated monotonically through the theory hierarchy: forming a new theory, the union of the simpsets of its imports are taken as starting point. Also note that definitional packages like @{command "datatype"}, @{command "primrec"}, @{command "fun"} routinely declare Simplifier rules to the target context, while plain @{command "definition"} is an exception in \<^emph>\not\ declaring anything. \<^medskip> It is up the user to manipulate the current simpset further by explicitly adding or deleting theorems as simplification rules, or installing other tools via simplification procedures (\secref{sec:simproc}). Good simpsets are hard to design. Rules that obviously simplify, like \?n + 0 \ ?n\ are good candidates for the implicit simpset, unless a special non-normalizing behavior of certain operations is intended. More specific rules (such as distributive laws, which duplicate subterms) should be added only for specific proof steps. Conversely, sometimes a rule needs to be deleted just for some part of a proof. The need of frequent additions or deletions may indicate a poorly designed simpset. \begin{warn} The union of simpsets from theory imports (as described above) is not always a good starting point for the new theory. If some ancestors have deleted simplification rules because they are no longer wanted, while others have left those rules in, then the union will contain the unwanted rules, and thus have to be deleted again in the theory body. \end{warn} \ subsection \Ordered rewriting with permutative rules\ text \ A rewrite rule is \<^emph>\permutative\ if the left-hand side and right-hand side are the equal up to renaming of variables. The most common permutative rule is commutativity: \?x + ?y = ?y + ?x\. Other examples include \(?x - ?y) - ?z = (?x - ?z) - ?y\ in arithmetic and \insert ?x (insert ?y ?A) = insert ?y (insert ?x ?A)\ for sets. Such rules are common enough to merit special attention. Because ordinary rewriting loops given such rules, the Simplifier employs a special strategy, called \<^emph>\ordered rewriting\. Permutative rules are detected and only applied if the rewriting step decreases the redex wrt.\ a given term ordering. For example, commutativity rewrites \b + a\ to \a + b\, but then stops, because the redex cannot be decreased further in the sense of the term ordering. The default is lexicographic ordering of term structure, but this could be also changed locally for special applications via @{define_ML Simplifier.set_term_ord} in Isabelle/ML. \<^medskip> Permutative rewrite rules are declared to the Simplifier just like other rewrite rules. Their special status is recognized automatically, and their application is guarded by the term ordering accordingly. \ subsubsection \Rewriting with AC operators\ text \ Ordered rewriting is particularly effective in the case of associative-commutative operators. (Associativity by itself is not permutative.) When dealing with an AC-operator \f\, keep the following points in mind: \<^item> The associative law must always be oriented from left to right, namely \f (f x y) z = f x (f y z)\. The opposite orientation, if used with commutativity, leads to looping in conjunction with the standard term order. \<^item> To complete your set of rewrite rules, you must add not just associativity (A) and commutativity (C) but also a derived rule \<^emph>\left-commutativity\ (LC): \f x (f y z) = f y (f x z)\. Ordered rewriting with the combination of A, C, and LC sorts a term lexicographically --- the rewriting engine imitates bubble-sort. \ experiment fixes f :: "'a \ 'a \ 'a" (infix "\" 60) assumes assoc: "(x \ y) \ z = x \ (y \ z)" assumes commute: "x \ y = y \ x" begin lemma left_commute: "x \ (y \ z) = y \ (x \ z)" proof - have "(x \ y) \ z = (y \ x) \ z" by (simp only: commute) then show ?thesis by (simp only: assoc) qed lemmas AC_rules = assoc commute left_commute text \ Thus the Simplifier is able to establish equalities with arbitrary permutations of subterms, by normalizing to a common standard form. For example: \ lemma "(b \ c) \ a = xxx" apply (simp only: AC_rules) txt \\<^subgoals>\ oops lemma "(b \ c) \ a = a \ (b \ c)" by (simp only: AC_rules) lemma "(b \ c) \ a = c \ (b \ a)" by (simp only: AC_rules) lemma "(b \ c) \ a = (c \ b) \ a" by (simp only: AC_rules) end text \ - Martin and Nipkow @{cite "martin-nipkow"} discuss the theory and give many + Martin and Nipkow \<^cite>\"martin-nipkow"\ discuss the theory and give many examples; other algebraic structures are amenable to ordered rewriting, such - as Boolean rings. The Boyer-Moore theorem prover @{cite bm88book} also + as Boolean rings. The Boyer-Moore theorem prover \<^cite>\bm88book\ also employs ordered rewriting. \ subsubsection \Re-orienting equalities\ text \Another application of ordered rewriting uses the derived rule @{thm [source] eq_commute}: @{thm [source = false] eq_commute} to reverse equations. This is occasionally useful to re-orient local assumptions according to the term ordering, when other built-in mechanisms of reorientation and mutual simplification fail to apply.\ subsection \Simplifier tracing and debugging \label{sec:simp-trace}\ text \ \begin{tabular}{rcll} @{attribute_def simp_trace} & : & \attribute\ & default \false\ \\ @{attribute_def simp_trace_depth_limit} & : & \attribute\ & default \1\ \\ @{attribute_def simp_debug} & : & \attribute\ & default \false\ \\ @{attribute_def simp_trace_new} & : & \attribute\ \\ @{attribute_def simp_break} & : & \attribute\ \\ \end{tabular} \<^medskip> \<^rail>\ @@{attribute simp_trace_new} ('interactive')? \ ('mode' '=' ('full' | 'normal'))? \ ('depth' '=' @{syntax nat})? ; @@{attribute simp_break} (@{syntax term}*) \ These attributes and configurations options control various aspects of Simplifier tracing and debugging. \<^descr> @{attribute simp_trace} makes the Simplifier output internal operations. This includes rewrite steps, but also bookkeeping like modifications of the simpset. \<^descr> @{attribute simp_trace_depth_limit} limits the effect of @{attribute simp_trace} to the given depth of recursive Simplifier invocations (when solving conditions of rewrite rules). \<^descr> @{attribute simp_debug} makes the Simplifier output some extra information about internal operations. This includes any attempted invocation of simplification procedures. \<^descr> @{attribute simp_trace_new} controls Simplifier tracing within - Isabelle/PIDE applications, notably Isabelle/jEdit @{cite "isabelle-jedit"}. + Isabelle/PIDE applications, notably Isabelle/jEdit \<^cite>\"isabelle-jedit"\. This provides a hierarchical representation of the rewriting steps performed by the Simplifier. Users can configure the behaviour by specifying breakpoints, verbosity and enabling or disabling the interactive mode. In normal verbosity (the default), only rule applications matching a breakpoint will be shown to the user. In full verbosity, all rule applications will be logged. Interactive mode interrupts the normal flow of the Simplifier and defers the decision how to continue to the user via some GUI dialog. \<^descr> @{attribute simp_break} declares term or theorem breakpoints for @{attribute simp_trace_new} as described above. Term breakpoints are patterns which are checked for matches on the redex of a rule application. Theorem breakpoints trigger when the corresponding theorem is applied in a rewrite step. For example: \ (*<*)experiment begin(*>*) declare conjI [simp_break] declare [[simp_break "?x \ ?y"]] (*<*)end(*>*) subsection \Simplification procedures \label{sec:simproc}\ text \ Simplification procedures are ML functions that produce proven rewrite rules on demand. They are associated with higher-order patterns that approximate the left-hand sides of equations. The Simplifier first matches the current redex against one of the LHS patterns; if this succeeds, the corresponding ML function is invoked, passing the Simplifier context and redex term. Thus rules may be specifically fashioned for particular situations, resulting in a more powerful mechanism than term rewriting by a fixed set of rules. Any successful result needs to be a (possibly conditional) rewrite rule \t \ u\ that is applicable to the current redex. The rule will be applied just as any ordinary rewrite rule. It is expected to be already in \<^emph>\internal form\, bypassing the automatic preprocessing of object-level equivalences. \begin{matharray}{rcl} @{command_def "simproc_setup"} & : & \local_theory \ local_theory\ \\ simproc & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '=' @{syntax text}; @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+) \ \<^descr> @{command "simproc_setup"} defines a named simplification procedure that is invoked by the Simplifier whenever any of the given term patterns match the current redex. The implementation, which is provided as ML source text, needs to be of type \<^ML_type>\morphism -> Proof.context -> cterm -> thm option\, where the \<^ML_type>\cterm\ represents the current redex \r\ and the result is supposed to be some proven rewrite rule \r \ r'\ (or a generalized version), or \<^ML>\NONE\ to indicate failure. The \<^ML_type>\Proof.context\ argument holds the full context of the current Simplifier invocation. The \<^ML_type>\morphism\ informs about the difference of the original compilation context wrt.\ the one of the actual application later on. Morphisms are only relevant for simprocs that are defined within a local target context, e.g.\ in a locale. \<^descr> \simproc add: name\ and \simproc del: name\ add or delete named simprocs to the current Simplifier context. The default is to add a simproc. Note that @{command "simproc_setup"} already adds the new simproc to the subsequent context. \ subsubsection \Example\ text \ The following simplification procedure for @{thm [source = false, show_types] unit_eq} in HOL performs fine-grained control over rule application, beyond higher-order pattern matching. Declaring @{thm unit_eq} as @{attribute simp} directly would make the Simplifier loop! Note that a version of this simplification procedure is already active in Isabelle/HOL. \ (*<*)experiment begin(*>*) simproc_setup unit ("x::unit") = \fn _ => fn _ => fn ct => if HOLogic.is_unit (Thm.term_of ct) then NONE else SOME (mk_meta_eq @{thm unit_eq})\ (*<*)end(*>*) text \ Since the Simplifier applies simplification procedures frequently, it is important to make the failure check in ML reasonably fast.\ subsection \Configurable Simplifier strategies \label{sec:simp-strategies}\ text \ The core term-rewriting engine of the Simplifier is normally used in combination with some add-on components that modify the strategy and allow to integrate other non-Simplifier proof tools. These may be reconfigured in ML as explained below. Even if the default strategies of object-logics like Isabelle/HOL are used unchanged, it helps to understand how the standard Simplifier strategies work.\ subsubsection \The subgoaler\ text \ \begin{mldecls} @{define_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) -> Proof.context -> Proof.context"} \\ @{define_ML Simplifier.prems_of: "Proof.context -> thm list"} \\ \end{mldecls} The subgoaler is the tactic used to solve subgoals arising out of conditional rewrite rules or congruence rules. The default should be simplification itself. In rare situations, this strategy may need to be changed. For example, if the premise of a conditional rule is an instance of its conclusion, as in \Suc ?m < ?n \ ?m < ?n\, the default strategy could loop. % FIXME !?? \<^descr> \<^ML>\Simplifier.set_subgoaler\~\tac ctxt\ sets the subgoaler of the context to \tac\. The tactic will be applied to the context of the running Simplifier instance. \<^descr> \<^ML>\Simplifier.prems_of\~\ctxt\ retrieves the current set of premises from the context. This may be non-empty only if the Simplifier has been told to utilize local assumptions in the first place (cf.\ the options in \secref{sec:simp-meth}). As an example, consider the following alternative subgoaler: \ ML_val \ fun subgoaler_tac ctxt = assume_tac ctxt ORELSE' resolve_tac ctxt (Simplifier.prems_of ctxt) ORELSE' asm_simp_tac ctxt \ text \ This tactic first tries to solve the subgoal by assumption or by resolving with with one of the premises, calling simplification only if that fails.\ subsubsection \The solver\ text \ \begin{mldecls} @{define_ML_type solver} \\ @{define_ML Simplifier.mk_solver: "string -> (Proof.context -> int -> tactic) -> solver"} \\ @{define_ML_infix setSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix addSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix setSSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix addSSolver: "Proof.context * solver -> Proof.context"} \\ \end{mldecls} A solver is a tactic that attempts to solve a subgoal after simplification. Its core functionality is to prove trivial subgoals such as \<^prop>\True\ and \t = t\, but object-logics might be more ambitious. For example, Isabelle/HOL performs a restricted version of linear arithmetic here. Solvers are packaged up in abstract type \<^ML_type>\solver\, with \<^ML>\Simplifier.mk_solver\ as the only operation to create a solver. \<^medskip> Rewriting does not instantiate unknowns. For example, rewriting alone cannot prove \a \ ?A\ since this requires instantiating \?A\. The solver, however, is an arbitrary tactic and may instantiate unknowns as it pleases. This is the only way the Simplifier can handle a conditional rewrite rule whose condition contains extra variables. When a simplification tactic is to be combined with other provers, especially with the Classical Reasoner, it is important whether it can be considered safe or not. For this reason a simpset contains two solvers: safe and unsafe. The standard simplification strategy solely uses the unsafe solver, which is appropriate in most cases. For special applications where the simplification process is not allowed to instantiate unknowns within the goal, simplification starts with the safe solver, but may still apply the ordinary unsafe one in nested simplifications for conditional rules or congruences. Note that in this way the overall tactic is not totally safe: it may instantiate unknowns that appear also in other subgoals. \<^descr> \<^ML>\Simplifier.mk_solver\~\name tac\ turns \tac\ into a solver; the \name\ is only attached as a comment and has no further significance. \<^descr> \ctxt setSSolver solver\ installs \solver\ as the safe solver of \ctxt\. \<^descr> \ctxt addSSolver solver\ adds \solver\ as an additional safe solver; it will be tried after the solvers which had already been present in \ctxt\. \<^descr> \ctxt setSolver solver\ installs \solver\ as the unsafe solver of \ctxt\. \<^descr> \ctxt addSolver solver\ adds \solver\ as an additional unsafe solver; it will be tried after the solvers which had already been present in \ctxt\. \<^medskip> The solver tactic is invoked with the context of the running Simplifier. Further operations may be used to retrieve relevant information, such as the list of local Simplifier premises via \<^ML>\Simplifier.prems_of\ --- this list may be non-empty only if the Simplifier runs in a mode that utilizes local assumptions (see also \secref{sec:simp-meth}). The solver is also presented the full goal including its assumptions in any case. Thus it can use these (e.g.\ by calling \<^ML>\assume_tac\), even if the Simplifier proper happens to ignore local premises at the moment. \<^medskip> As explained before, the subgoaler is also used to solve the premises of congruence rules. These are usually of the form \s = ?x\, where \s\ needs to be simplified and \?x\ needs to be instantiated with the result. Typically, the subgoaler will invoke the Simplifier at some point, which will eventually call the solver. For this reason, solver tactics must be prepared to solve goals of the form \t = ?x\, usually by reflexivity. In particular, reflexivity should be tried before any of the fancy automated proof tools. It may even happen that due to simplification the subgoal is no longer an equality. For example, \False \ ?Q\ could be rewritten to \\ ?Q\. To cover this case, the solver could try resolving with the theorem \\ False\ of the object-logic. \<^medskip> \begin{warn} If a premise of a congruence rule cannot be proved, then the congruence is ignored. This should only happen if the rule is \<^emph>\conditional\ --- that is, contains premises not of the form \t = ?x\. Otherwise it indicates that some congruence rule, or possibly the subgoaler or solver, is faulty. \end{warn} \ subsubsection \The looper\ text \ \begin{mldecls} @{define_ML_infix setloop: "Proof.context * (Proof.context -> int -> tactic) -> Proof.context"} \\ @{define_ML_infix addloop: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delloop: "Proof.context * string -> Proof.context"} \\ @{define_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.add_split_bang: " thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\ \end{mldecls} The looper is a list of tactics that are applied after simplification, in case the solver failed to solve the simplified goal. If the looper succeeds, the simplification process is started all over again. Each of the subgoals generated by the looper is attacked in turn, in reverse order. A typical looper is \<^emph>\case splitting\: the expansion of a conditional. Another possibility is to apply an elimination rule on the assumptions. More adventurous loopers could start an induction. \<^descr> \ctxt setloop tac\ installs \tac\ as the only looper tactic of \ctxt\. \<^descr> \ctxt addloop (name, tac)\ adds \tac\ as an additional looper tactic with name \name\, which is significant for managing the collection of loopers. The tactic will be tried after the looper tactics that had already been present in \ctxt\. \<^descr> \ctxt delloop name\ deletes the looper tactic that was associated with \name\ from \ctxt\. \<^descr> \<^ML>\Splitter.add_split\~\thm ctxt\ adds split tactic for \thm\ as additional looper tactic of \ctxt\ (overwriting previous split tactic for the same constant). \<^descr> \<^ML>\Splitter.add_split_bang\~\thm ctxt\ adds aggressive (see \S\ref{sec:simp-meth}) split tactic for \thm\ as additional looper tactic of \ctxt\ (overwriting previous split tactic for the same constant). \<^descr> \<^ML>\Splitter.del_split\~\thm ctxt\ deletes the split tactic corresponding to \thm\ from the looper tactics of \ctxt\. The splitter replaces applications of a given function; the right-hand side of the replacement can be anything. For example, here is a splitting rule for conditional expressions: @{text [display] "?P (if ?Q ?x ?y) \ (?Q \ ?P ?x) \ (\ ?Q \ ?P ?y)"} Another example is the elimination operator for Cartesian products (which happens to be called \<^const>\case_prod\ in Isabelle/HOL: @{text [display] "?P (case_prod ?f ?p) \ (\a b. ?p = (a, b) \ ?P (f a b))"} For technical reasons, there is a distinction between case splitting in the conclusion and in the premises of a subgoal. The former is done by \<^ML>\Splitter.split_tac\ with rules like @{thm [source] if_split} or @{thm [source] option.split}, which do not split the subgoal, while the latter is done by \<^ML>\Splitter.split_asm_tac\ with rules like @{thm [source] if_split_asm} or @{thm [source] option.split_asm}, which split the subgoal. The function \<^ML>\Splitter.add_split\ automatically takes care of which tactic to call, analyzing the form of the rules given as argument; it is the same operation behind \split\ attribute or method modifier syntax in the Isar source language. Case splits should be allowed only when necessary; they are expensive and hard to control. Case-splitting on if-expressions in the conclusion is usually beneficial, so it is enabled by default in Isabelle/HOL and Isabelle/FOL/ZF. \begin{warn} With \<^ML>\Splitter.split_asm_tac\ as looper component, the Simplifier may split subgoals! This might cause unexpected problems in tactic expressions that silently assume 0 or 1 subgoals after simplification. \end{warn} \ subsection \Forward simplification \label{sec:simp-forward}\ text \ \begin{matharray}{rcl} @{attribute_def simplified} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute simplified} opt? @{syntax thms}? ; opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')' \ \<^descr> @{attribute simplified}~\a\<^sub>1 \ a\<^sub>n\ causes a theorem to be simplified, either by exactly the specified rules \a\<^sub>1, \, a\<^sub>n\, or the implicit Simplifier context if no arguments are given. The result is fully simplified by default, including assumptions and conclusion; the options \no_asm\ etc.\ tune the Simplifier in the same way as the for the \simp\ method. Note that forward simplification restricts the Simplifier to its most basic operation of term rewriting; solver and looper tactics (\secref{sec:simp-strategies}) are \<^emph>\not\ involved here. The @{attribute simplified} attribute should be only rarely required under normal circumstances. \ section \The Classical Reasoner \label{sec:classical}\ subsection \Basic concepts\ text \Although Isabelle is generic, many users will be working in some extension of classical first-order logic. Isabelle/ZF is built upon theory FOL, while Isabelle/HOL conceptually contains first-order logic as a fragment. Theorem-proving in predicate logic is undecidable, but many automated strategies have been developed to assist in this task. Isabelle's classical reasoner is a generic package that accepts certain information about a logic and delivers a suite of automatic proof tools, based on rules that are classified and declared in the context. These proof procedures are slow and simplistic compared with high-end automated theorem provers, but they can save considerable time and effort in practice. They can prove theorems - such as Pelletier's @{cite pelletier86} problems 40 and 41 in a few + such as Pelletier's \<^cite>\pelletier86\ problems 40 and 41 in a few milliseconds (including full proof reconstruction):\ lemma "(\y. \x. F x y \ F x x) \ \ (\x. \y. \z. F z y \ \ F z x)" by blast lemma "(\z. \y. \x. f x y \ f x z \ \ f x x) \ \ (\z. \x. f x z)" by blast text \The proof tools are generic. They are not restricted to first-order logic, and have been heavily used in the development of the Isabelle/HOL library and applications. The tactics can be traced, and their components can be called directly; in this manner, any proof can be viewed interactively.\ subsubsection \The sequent calculus\ text \Isabelle supports natural deduction, which is easy to use for interactive proof. But natural deduction does not easily lend itself to automation, and has a bias towards intuitionism. For certain proofs in classical logic, it can not be called natural. The \<^emph>\sequent calculus\, a generalization of natural deduction, is easier to automate. A \<^bold>\sequent\ has the form \\ \ \\, where \\\ and \\\ are sets of formulae.\<^footnote>\For first-order logic, sequents can equivalently be made from lists or multisets of formulae.\ The sequent \P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n\ is \<^bold>\valid\ if \P\<^sub>1 \ \ \ P\<^sub>m\ implies \Q\<^sub>1 \ \ \ Q\<^sub>n\. Thus \P\<^sub>1, \, P\<^sub>m\ represent assumptions, each of which is true, while \Q\<^sub>1, \, Q\<^sub>n\ represent alternative goals. A sequent is \<^bold>\basic\ if its left and right sides have a common formula, as in \P, Q \ Q, R\; basic sequents are trivially valid. Sequent rules are classified as \<^bold>\right\ or \<^bold>\left\, indicating which side of the \\\ symbol they operate on. Rules that operate on the right side are analogous to natural deduction's introduction rules, and left rules are analogous to elimination rules. The sequent calculus analogue of \(\I)\ is the rule \[ \infer[\(\R)\]{\\ \ \, P \ Q\}{\P, \ \ \, Q\} \] Applying the rule backwards, this breaks down some implication on the right side of a sequent; \\\ and \\\ stand for the sets of formulae that are unaffected by the inference. The analogue of the pair \(\I1)\ and \(\I2)\ is the single rule \[ \infer[\(\R)\]{\\ \ \, P \ Q\}{\\ \ \, P, Q\} \] This breaks down some disjunction on the right side, replacing it by both disjuncts. Thus, the sequent calculus is a kind of multiple-conclusion logic. To illustrate the use of multiple formulae on the right, let us prove the classical theorem \(P \ Q) \ (Q \ P)\. Working backwards, we reduce this formula to a basic sequent: \[ \infer[\(\R)\]{\\ (P \ Q) \ (Q \ P)\} {\infer[\(\R)\]{\\ (P \ Q), (Q \ P)\} {\infer[\(\R)\]{\P \ Q, (Q \ P)\} {\P, Q \ Q, P\}}} \] This example is typical of the sequent calculus: start with the desired theorem and apply rules backwards in a fairly arbitrary manner. This yields a surprisingly effective proof procedure. Quantifiers add only few complications, since Isabelle handles - parameters and schematic variables. See @{cite \Chapter 10\ - "paulson-ml2"} for further discussion.\ + parameters and schematic variables. See \<^cite>\\Chapter 10\ in + "paulson-ml2"\ for further discussion.\ subsubsection \Simulating sequents by natural deduction\ text \Isabelle can represent sequents directly, as in the object-logic LK. But natural deduction is easier to work with, and most object-logics employ it. Fortunately, we can simulate the sequent \P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n\ by the Isabelle formula \P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>2 \ ... \ \ Q\<^sub>n \ Q\<^sub>1\ where the order of the assumptions and the choice of \Q\<^sub>1\ are arbitrary. Elim-resolution plays a key role in simulating sequent proofs. We can easily handle reasoning on the left. Elim-resolution with the rules \(\E)\, \(\E)\ and \(\E)\ achieves a similar effect as the corresponding sequent rules. For the other connectives, we use sequent-style elimination rules instead of destruction rules such as \(\E1, 2)\ and \(\E)\. But note that the rule \(\L)\ has no effect under our representation of sequents! \[ \infer[\(\L)\]{\\ P, \ \ \\}{\\ \ \, P\} \] What about reasoning on the right? Introduction rules can only affect the formula in the conclusion, namely \Q\<^sub>1\. The other right-side formulae are represented as negated assumptions, \\ Q\<^sub>2, \, \ Q\<^sub>n\. In order to operate on one of these, it must first be exchanged with \Q\<^sub>1\. Elim-resolution with the \swap\ rule has this effect: \\ P \ (\ R \ P) \ R\ To ensure that swaps occur only when necessary, each introduction rule is converted into a swapped form: it is resolved with the second premise of \(swap)\. The swapped form of \(\I)\, which might be called \(\\E)\, is @{text [display] "\ (P \ Q) \ (\ R \ P) \ (\ R \ Q) \ R"} Similarly, the swapped form of \(\I)\ is @{text [display] "\ (P \ Q) \ (\ R \ P \ Q) \ R"} Swapped introduction rules are applied using elim-resolution, which deletes the negated formula. Our representation of sequents also requires the use of ordinary introduction rules. If we had no regard for readability of intermediate goal states, we could treat the right side more uniformly by representing sequents as @{text [display] "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>1 \ \ \ \ Q\<^sub>n \ \"} \ subsubsection \Extra rules for the sequent calculus\ text \As mentioned, destruction rules such as \(\E1, 2)\ and \(\E)\ must be replaced by sequent-style elimination rules. In addition, we need rules to embody the classical equivalence between \P \ Q\ and \\ P \ Q\. The introduction rules \(\I1, 2)\ are replaced by a rule that simulates \(\R)\: @{text [display] "(\ Q \ P) \ P \ Q"} The destruction rule \(\E)\ is replaced by @{text [display] "(P \ Q) \ (\ P \ R) \ (Q \ R) \ R"} Quantifier replication also requires special rules. In classical logic, \\x. P x\ is equivalent to \\ (\x. \ P x)\; the rules \(\R)\ and \(\L)\ are dual: \[ \infer[\(\R)\]{\\ \ \, \x. P x\}{\\ \ \, \x. P x, P t\} \qquad \infer[\(\L)\]{\\x. P x, \ \ \\}{\P t, \x. P x, \ \ \\} \] Thus both kinds of quantifier may be replicated. Theorems requiring multiple uses of a universal formula are easy to invent; consider @{text [display] "(\x. P x \ P (f x)) \ P a \ P (f\<^sup>n a)"} for any \n > 1\. Natural examples of the multiple use of an existential formula are rare; a standard one is \\x. \y. P x \ P y\. Forgoing quantifier replication loses completeness, but gains decidability, since the search space becomes finite. Many useful theorems can be proved without replication, and the search generally delivers its verdict in a reasonable time. To adopt this approach, represent the sequent rules \(\R)\, \(\L)\ and \(\R)\ by \(\I)\, \(\E)\ and \(\I)\, respectively, and put \(\E)\ into elimination form: @{text [display] "\x. P x \ (P t \ Q) \ Q"} Elim-resolution with this rule will delete the universal formula after a single use. To replicate universal quantifiers, replace the rule by @{text [display] "\x. P x \ (P t \ \x. P x \ Q) \ Q"} To replicate existential quantifiers, replace \(\I)\ by @{text [display] "(\ (\x. P x) \ P t) \ \x. P x"} All introduction rules mentioned above are also useful in swapped form. Replication makes the search space infinite; we must apply the rules with care. The classical reasoner distinguishes between safe and unsafe rules, applying the latter only when there is no alternative. Depth-first search may well go down a blind alley; best-first search is better behaved in an infinite search space. However, quantifier replication is too expensive to prove any but the simplest theorems. \ subsection \Rule declarations\ text \The proof tools of the Classical Reasoner depend on collections of rules declared in the context, which are classified as introduction, elimination or destruction and as \<^emph>\safe\ or \<^emph>\unsafe\. In general, safe rules can be attempted blindly, while unsafe rules must be used with care. A safe rule must never reduce a provable goal to an unprovable set of subgoals. The rule \P \ P \ Q\ is unsafe because it reduces \P \ Q\ to \P\, which might turn out as premature choice of an unprovable subgoal. Any rule whose premises contain new unknowns is unsafe. The elimination rule \\x. P x \ (P t \ Q) \ Q\ is unsafe, since it is applied via elim-resolution, which discards the assumption \\x. P x\ and replaces it by the weaker assumption \P t\. The rule \P t \ \x. P x\ is unsafe for similar reasons. The quantifier duplication rule \\x. P x \ (P t \ \x. P x \ Q) \ Q\ is unsafe in a different sense: since it keeps the assumption \\x. P x\, it is prone to looping. In classical first-order logic, all rules are safe except those mentioned above. The safe~/ unsafe distinction is vague, and may be regarded merely as a way of giving some rules priority over others. One could argue that \(\E)\ is unsafe, because repeated application of it could generate exponentially many subgoals. Induction rules are unsafe because inductive proofs are difficult to set up automatically. Any inference that instantiates an unknown in the proof state is unsafe --- thus matching must be used, rather than unification. Even proof by assumption is unsafe if it instantiates unknowns shared with other subgoals. \begin{matharray}{rcl} @{command_def "print_claset"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def intro} & : & \attribute\ \\ @{attribute_def elim} & : & \attribute\ \\ @{attribute_def dest} & : & \attribute\ \\ @{attribute_def rule} & : & \attribute\ \\ @{attribute_def iff} & : & \attribute\ \\ @{attribute_def swapped} & : & \attribute\ \\ \end{matharray} \<^rail>\ (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? ; @@{attribute rule} 'del' ; @@{attribute iff} (((() | 'add') '?'?) | 'del') \ \<^descr> @{command "print_claset"} prints the collection of rules declared to the Classical Reasoner, i.e.\ the \<^ML_type>\claset\ within the context. \<^descr> @{attribute intro}, @{attribute elim}, and @{attribute dest} declare introduction, elimination, and destruction rules, respectively. By default, rules are considered as \<^emph>\unsafe\ (i.e.\ not applied blindly without backtracking), while ``\!\'' classifies as \<^emph>\safe\. Rule declarations marked by ``\?\'' coincide with those of Isabelle/Pure, cf.\ \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps of the @{method rule} method). The optional natural number specifies an explicit weight argument, which is ignored by the automated reasoning tools, but determines the search order of single rule steps. Introduction rules are those that can be applied using ordinary resolution. Their swapped forms are generated internally, which will be applied using elim-resolution. Elimination rules are applied using elim-resolution. Rules are sorted by the number of new subgoals they will yield; rules that generate the fewest subgoals will be tried first. Otherwise, later declarations take precedence over earlier ones. Rules already present in the context with the same classification are ignored. A warning is printed if the rule has already been added with some other classification, but the rule is added anyway as requested. \<^descr> @{attribute rule}~\del\ deletes all occurrences of a rule from the classical context, regardless of its classification as introduction~/ elimination~/ destruction and safe~/ unsafe. \<^descr> @{attribute iff} declares logical equivalences to the Simplifier and the Classical reasoner at the same time. Non-conditional rules result in a safe introduction and elimination pair; conditional ones are considered unsafe. Rules with negative conclusion are automatically inverted (using \\\-elimination internally). The ``\?\'' version of @{attribute iff} declares rules to the Isabelle/Pure context only, and omits the Simplifier declaration. \<^descr> @{attribute swapped} turns an introduction rule into an elimination, by resolving with the classical swap principle \\ P \ (\ R \ P) \ R\ in the second position. This is mainly for illustrative purposes: the Classical Reasoner already swaps rules internally as explained above. \ subsection \Structured methods\ text \ \begin{matharray}{rcl} @{method_def rule} & : & \method\ \\ @{method_def contradiction} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method rule} @{syntax thms}? \ \<^descr> @{method rule} as offered by the Classical Reasoner is a refinement over the Pure one (see \secref{sec:pure-meth-att}). Both versions work the same, but the classical version observes the classical rule context in addition to that of Isabelle/Pure. Common object logics (HOL, ZF, etc.) declare a rich collection of classical rules (even if these would qualify as intuitionistic ones), but only few declarations to the rule context of Isabelle/Pure (\secref{sec:pure-meth-att}). \<^descr> @{method contradiction} solves some goal by contradiction, deriving any result from both \\ A\ and \A\. Chained facts, which are guaranteed to participate, may appear in either order. \ subsection \Fully automated methods\ text \ \begin{matharray}{rcl} @{method_def blast} & : & \method\ \\ @{method_def auto} & : & \method\ \\ @{method_def force} & : & \method\ \\ @{method_def fast} & : & \method\ \\ @{method_def slow} & : & \method\ \\ @{method_def best} & : & \method\ \\ @{method_def fastforce} & : & \method\ \\ @{method_def slowsimp} & : & \method\ \\ @{method_def bestsimp} & : & \method\ \\ @{method_def deepen} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method blast} @{syntax nat}? (@{syntax clamod} * ) ; @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) ; @@{method force} (@{syntax clasimpmod} * ) ; (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * ) ; (@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp}) (@{syntax clasimpmod} * ) ; @@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * ) ; @{syntax_def clamod}: (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thms} ; @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | 'cong' (() | 'add' | 'del') | 'split' (() | '!' | 'del') | 'iff' (((() | 'add') '?'?) | 'del') | (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thms} \ \<^descr> @{method blast} is a separate classical tableau prover that uses the same classical rule declarations as explained before. Proof search is coded directly in ML using special data structures. A successful proof is then reconstructed using regular Isabelle inferences. It is faster and more powerful than the other classical reasoning tools, but has major limitations too. \<^item> It does not use the classical wrapper tacticals, such as the integration with the Simplifier of @{method fastforce}. \<^item> It does not perform higher-order unification, as needed by the rule @{thm [source=false] rangeI} in HOL. There are often alternatives to such rules, for example @{thm [source=false] range_eqI}. \<^item> Function variables may only be applied to parameters of the subgoal. (This restriction arises because the prover does not use higher-order unification.) If other function variables are present then the prover will fail with the message @{verbatim [display] \Function unknown's argument not a bound variable\} \<^item> Its proof strategy is more general than @{method fast} but can be slower. If @{method blast} fails or seems to be running forever, try @{method fast} and the other proof tools described below. The optional integer argument specifies a bound for the number of unsafe steps used in a proof. By default, @{method blast} starts with a bound of 0 and increases it successively to 20. In contrast, \(blast lim)\ tries to prove the goal using a search bound of \lim\. Sometimes a slow proof using @{method blast} can be made much faster by supplying the successful search bound to this proof method instead. \<^descr> @{method auto} combines classical reasoning with simplification. It is intended for situations where there are a lot of mostly trivial subgoals; it proves all the easy ones, leaving the ones it cannot prove. Occasionally, attempting to prove the hard ones may take a long time. The optional depth arguments in \(auto m n)\ refer to its builtin classical reasoning procedures: \m\ (default 4) is for @{method blast}, which is tried first, and \n\ (default 2) is for a slower but more general alternative that also takes wrappers into account. \<^descr> @{method force} is intended to prove the first subgoal completely, using many fancy proof tools and performing a rather exhaustive search. As a result, proof attempts may take rather long or diverge easily. \<^descr> @{method fast}, @{method best}, @{method slow} attempt to prove the first subgoal using sequent-style reasoning as explained before. Unlike @{method blast}, they construct proofs directly in Isabelle. There is a difference in search strategy and back-tracking: @{method fast} uses depth-first search and @{method best} uses best-first search (guided by a heuristic function: normally the total size of the proof state). Method @{method slow} is like @{method fast}, but conducts a broader search: it may, when backtracking from a failed proof attempt, undo even the step of proving a subgoal by assumption. \<^descr> @{method fastforce}, @{method slowsimp}, @{method bestsimp} are like @{method fast}, @{method slow}, @{method best}, respectively, but use the Simplifier as additional wrapper. The name @{method fastforce}, reflects the behaviour of this popular method better without requiring an understanding of its implementation. \<^descr> @{method deepen} works by exhaustive search up to a certain depth. The start depth is 4 (unless specified explicitly), and the depth is increased iteratively up to 10. Unsafe rules are modified to preserve the formula they act on, so that it be used repeatedly. This method can prove more goals than @{method fast}, but is much slower, for example if the assumptions have many universal quantifiers. Any of the above methods support additional modifiers of the context of classical (and simplifier) rules, but the ones related to the Simplifier are explicitly prefixed by \simp\ here. The semantics of these ad-hoc rule declarations is analogous to the attributes given before. Facts provided by forward chaining are inserted into the goal before commencing proof search. \ subsection \Partially automated methods\label{sec:classical:partial}\ text \These proof methods may help in situations when the fully-automated tools fail. The result is a simpler subgoal that can be tackled by other means, such as by manual instantiation of quantifiers. \begin{matharray}{rcl} @{method_def safe} & : & \method\ \\ @{method_def clarify} & : & \method\ \\ @{method_def clarsimp} & : & \method\ \\ \end{matharray} \<^rail>\ (@@{method safe} | @@{method clarify}) (@{syntax clamod} * ) ; @@{method clarsimp} (@{syntax clasimpmod} * ) \ \<^descr> @{method safe} repeatedly performs safe steps on all subgoals. It is deterministic, with at most one outcome. \<^descr> @{method clarify} performs a series of safe steps without splitting subgoals; see also @{method clarify_step}. \<^descr> @{method clarsimp} acts like @{method clarify}, but also does simplification. Note that if the Simplifier context includes a splitter for the premises, the subgoal may still be split. \ subsection \Single-step tactics\ text \ \begin{matharray}{rcl} @{method_def safe_step} & : & \method\ \\ @{method_def inst_step} & : & \method\ \\ @{method_def step} & : & \method\ \\ @{method_def slow_step} & : & \method\ \\ @{method_def clarify_step} & : & \method\ \\ \end{matharray} These are the primitive tactics behind the automated proof methods of the Classical Reasoner. By calling them yourself, you can execute these procedures one step at a time. \<^descr> @{method safe_step} performs a safe step on the first subgoal. The safe wrapper tacticals are applied to a tactic that may include proof by assumption or Modus Ponens (taking care not to instantiate unknowns), or substitution. \<^descr> @{method inst_step} is like @{method safe_step}, but allows unknowns to be instantiated. \<^descr> @{method step} is the basic step of the proof procedure, it operates on the first subgoal. The unsafe wrapper tacticals are applied to a tactic that tries @{method safe}, @{method inst_step}, or applies an unsafe rule from the context. \<^descr> @{method slow_step} resembles @{method step}, but allows backtracking between using safe rules with instantiation (@{method inst_step}) and using unsafe rules. The resulting search space is larger. \<^descr> @{method clarify_step} performs a safe step on the first subgoal; no splitting step is applied. For example, the subgoal \A \ B\ is left as a conjunction. Proof by assumption, Modus Ponens, etc., may be performed provided they do not instantiate unknowns. Assumptions of the form \x = t\ may be eliminated. The safe wrapper tactical is applied. \ subsection \Modifying the search step\ text \ \begin{mldecls} @{define_ML_type wrapper = "(int -> tactic) -> (int -> tactic)"} \\[0.5ex] @{define_ML_infix addSWrapper: "Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context"} \\ @{define_ML_infix addSbefore: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix addSafter: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] @{define_ML_infix addWrapper: "Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context"} \\ @{define_ML_infix addbefore: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix addafter: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] @{define_ML addSss: "Proof.context -> Proof.context"} \\ @{define_ML addss: "Proof.context -> Proof.context"} \\ \end{mldecls} The proof strategy of the Classical Reasoner is simple. Perform as many safe inferences as possible; or else, apply certain safe rules, allowing instantiation of unknowns; or else, apply an unsafe rule. The tactics also eliminate assumptions of the form \x = t\ by substitution if they have been set up to do so. They may perform a form of Modus Ponens: if there are assumptions \P \ Q\ and \P\, then replace \P \ Q\ by \Q\. The classical reasoning tools --- except @{method blast} --- allow to modify this basic proof strategy by applying two lists of arbitrary \<^emph>\wrapper tacticals\ to it. The first wrapper list, which is considered to contain safe wrappers only, affects @{method safe_step} and all the tactics that call it. The second one, which may contain unsafe wrappers, affects the unsafe parts of @{method step}, @{method slow_step}, and the tactics that call them. A wrapper transforms each step of the search, for example by attempting other tactics before or after the original step tactic. All members of a wrapper list are applied in turn to the respective step tactic. Initially the two wrapper lists are empty, which means no modification of the step tactics. Safe and unsafe wrappers are added to the context with the functions given below, supplying them with wrapper names. These names may be used to selectively delete wrappers. \<^descr> \ctxt addSWrapper (name, wrapper)\ adds a new wrapper, which should yield a safe tactic, to modify the existing safe step tactic. \<^descr> \ctxt addSbefore (name, tac)\ adds the given tactic as a safe wrapper, such that it is tried \<^emph>\before\ each safe step of the search. \<^descr> \ctxt addSafter (name, tac)\ adds the given tactic as a safe wrapper, such that it is tried when a safe step of the search would fail. \<^descr> \ctxt delSWrapper name\ deletes the safe wrapper with the given name. \<^descr> \ctxt addWrapper (name, wrapper)\ adds a new wrapper to modify the existing (unsafe) step tactic. \<^descr> \ctxt addbefore (name, tac)\ adds the given tactic as an unsafe wrapper, such that it its result is concatenated \<^emph>\before\ the result of each unsafe step. \<^descr> \ctxt addafter (name, tac)\ adds the given tactic as an unsafe wrapper, such that it its result is concatenated \<^emph>\after\ the result of each unsafe step. \<^descr> \ctxt delWrapper name\ deletes the unsafe wrapper with the given name. \<^descr> \addSss\ adds the simpset of the context to its classical set. The assumptions and goal will be simplified, in a rather safe way, after each safe step of the search. \<^descr> \addss\ adds the simpset of the context to its classical set. The assumptions and goal will be simplified, before the each unsafe step of the search. \ section \Object-logic setup \label{sec:object-logic}\ text \ \begin{matharray}{rcl} @{command_def "judgment"} & : & \theory \ theory\ \\ @{method_def atomize} & : & \method\ \\ @{attribute_def atomize} & : & \attribute\ \\ @{attribute_def rule_format} & : & \attribute\ \\ @{attribute_def rulify} & : & \attribute\ \\ \end{matharray} The very starting point for any Isabelle object-logic is a ``truth judgment'' that links object-level statements to the meta-logic (with its minimal language of \prop\ that covers universal quantification \\\ and implication \\\). Common object-logics are sufficiently expressive to internalize rule statements over \\\ and \\\ within their own language. This is useful in certain situations where a rule needs to be viewed as an atomic statement from the meta-level perspective, e.g.\ \\x. x \ A \ P x\ versus \\x \ A. P x\. From the following language elements, only the @{method atomize} method and @{attribute rule_format} attribute are occasionally required by end-users, the rest is for those who need to setup their own object-logic. In the latter case existing formulations of Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. Generic tools may refer to the information provided by object-logic declarations internally. \<^rail>\ @@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}? ; @@{attribute atomize} ('(' 'full' ')')? ; @@{attribute rule_format} ('(' 'noasm' ')')? \ \<^descr> @{command "judgment"}~\c :: \ (mx)\ declares constant \c\ as the truth judgment of the current object-logic. Its type \\\ should specify a coercion of the category of object-level propositions to \prop\ of the Pure meta-logic; the mixfix annotation \(mx)\ would typically just link the object language (internally of syntactic category \logic\) with that of \prop\. Only one @{command "judgment"} declaration may be given in any theory development. \<^descr> @{method atomize} (as a method) rewrites any non-atomic premises of a sub-goal, using the meta-level equations declared via @{attribute atomize} (as an attribute) beforehand. As a result, heavily nested goals become amenable to fundamental operations such as resolution (cf.\ the @{method (Pure) rule} method). Giving the ``\(full)\'' option here means to turn the whole subgoal into an object-statement (if possible), including the outermost parameters and assumptions as well. A typical collection of @{attribute atomize} rules for a particular object-logic would provide an internalization for each of the connectives of \\\, \\\, and \\\. Meta-level conjunction should be covered as well (this is particularly important for locales, see \secref{sec:locale}). \<^descr> @{attribute rule_format} rewrites a theorem by the equalities declared as @{attribute rulify} rules in the current object-logic. By default, the result is fully normalized, including assumptions and conclusions at any depth. The \(no_asm)\ option restricts the transformation to the conclusion of a rule. In common object-logics (HOL, FOL, ZF), the effect of @{attribute rule_format} is to replace (bounded) universal quantification (\\\) and implication (\\\) by the corresponding rule statements over \\\ and \\\. \ section \Tracing higher-order unification\ text \ \begin{tabular}{rcll} @{attribute_def unify_trace_simp} & : & \attribute\ & default \false\ \\ @{attribute_def unify_trace_types} & : & \attribute\ & default \false\ \\ @{attribute_def unify_trace_bound} & : & \attribute\ & default \50\ \\ @{attribute_def unify_search_bound} & : & \attribute\ & default \60\ \\ \end{tabular} \<^medskip> Higher-order unification works well in most practical situations, but sometimes needs extra care to identify problems. These tracing options may help. \<^descr> @{attribute unify_trace_simp} controls tracing of the simplification phase of higher-order unification. \<^descr> @{attribute unify_trace_types} controls warnings of incompleteness, when unification is not considering all possible instantiations of schematic type variables. \<^descr> @{attribute unify_trace_bound} determines the depth where unification starts to print tracing information once it reaches depth; 0 for full tracing. At the default value, tracing information is almost never printed in practice. \<^descr> @{attribute unify_search_bound} prevents unification from searching past the given depth. Because of this bound, higher-order unification cannot return an infinite sequence, though it can return an exponentially long one. The search rarely approaches the default value in practice. If the search is cut off, unification prints a warning ``Unification bound exceeded''. \begin{warn} Options for unification cannot be modified in a local context. Only the global theory content is taken into account. \end{warn} \ end diff --git a/src/Doc/Isar_Ref/HOL_Specific.thy b/src/Doc/Isar_Ref/HOL_Specific.thy --- a/src/Doc/Isar_Ref/HOL_Specific.thy +++ b/src/Doc/Isar_Ref/HOL_Specific.thy @@ -1,2518 +1,2510 @@ (*:maxLineLen=78:*) theory HOL_Specific imports Main "HOL-Library.Old_Datatype" "HOL-Library.Old_Recdef" "HOL-Library.Adhoc_Overloading" "HOL-Library.Dlist" "HOL-Library.FSet" Base begin chapter \Higher-Order Logic\ text \ Isabelle/HOL is based on Higher-Order Logic, a polymorphic version of Church's Simple Theory of Types. HOL can be best understood as a simply-typed version of classical set theory. The logic was first - implemented in Gordon's HOL system @{cite "mgordon-hol"}. It extends - Church's original logic @{cite "church40"} by explicit type variables (naive + implemented in Gordon's HOL system \<^cite>\"mgordon-hol"\. It extends + Church's original logic \<^cite>\"church40"\ by explicit type variables (naive polymorphism) and a sound axiomatization scheme for new types based on subsets of existing types. - Andrews's book @{cite andrews86} is a full description of the original + Andrews's book \<^cite>\andrews86\ is a full description of the original Church-style higher-order logic, with proofs of correctness and completeness wrt.\ certain set-theoretic interpretations. The particular extensions of Gordon-style HOL are explained semantically in two chapters of the 1993 HOL - book @{cite pitts93}. + book \<^cite>\pitts93\. Experience with HOL over decades has demonstrated that higher-order logic is widely applicable in many areas of mathematics and computer science. In a sense, Higher-Order Logic is simpler than First-Order Logic, because there are fewer restrictions and special cases. Note that HOL is \<^emph>\weaker\ than FOL with axioms for ZF set theory, which is traditionally considered the standard foundation of regular mathematics, but for most applications this does not matter. If you prefer ML to Lisp, you will probably prefer HOL to ZF. \<^medskip> The syntax of HOL follows \\\-calculus and functional programming. Function application is curried. To apply the function \f\ of type \\\<^sub>1 \ \\<^sub>2 \ \\<^sub>3\ to the arguments \a\ and \b\ in HOL, you simply write \f a b\ (as in ML or Haskell). There is no ``apply'' operator; the existing application of the Pure \\\-calculus is re-used. Note that in HOL \f (a, b)\ means ``\f\ applied to the pair \(a, b)\ (which is notation for \Pair a b\). The latter typically introduces extra formal efforts that can be avoided by currying functions by default. Explicit tuples are as infrequent in HOL formalizations as in good ML or Haskell programs. \<^medskip> Isabelle/HOL has a distinct feel, compared to other object-logics like Isabelle/ZF. It identifies object-level types with meta-level types, taking advantage of the default type-inference mechanism of Isabelle/Pure. HOL fully identifies object-level functions with meta-level functions, with native abstraction and application. These identifications allow Isabelle to support HOL particularly nicely, but they also mean that HOL requires some sophistication from the user. In particular, an understanding of Hindley-Milner type-inference with type-classes, which are both used extensively in the standard libraries and applications. \ chapter \Derived specification elements\ section \Inductive and coinductive definitions \label{sec:hol-inductive}\ text \ \begin{matharray}{rcl} @{command_def (HOL) "inductive"} & : & \local_theory \ local_theory\ \\ @{command_def (HOL) "inductive_set"} & : & \local_theory \ local_theory\ \\ @{command_def (HOL) "coinductive"} & : & \local_theory \ local_theory\ \\ @{command_def (HOL) "coinductive_set"} & : & \local_theory \ local_theory\ \\ @{command_def "print_inductives"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def (HOL) mono} & : & \attribute\ \\ \end{matharray} An \<^emph>\inductive definition\ specifies the least predicate or set \R\ closed under given rules: applying a rule to elements of \R\ yields a result within \R\. For example, a structural operational semantics is an inductive definition of an evaluation relation. Dually, a \<^emph>\coinductive definition\ specifies the greatest predicate or set \R\ that is consistent with given rules: every element of \R\ can be seen as arising by applying a rule to elements of \R\. An important example is using bisimulation relations to formalise equivalence of processes and infinite data structures. Both inductive and coinductive definitions are based on the Knaster-Tarski fixed-point theorem for complete lattices. The collection of introduction rules given by the user determines a functor on subsets of set-theoretic relations. The required monotonicity of the recursion scheme is proven as a prerequisite to the fixed-point definition and the resulting consequences. This works by pushing inclusion through logical connectives and any other operator that might be wrapped around recursive occurrences of the defined relation: there must be a monotonicity theorem of the form \A \ B \ \ A \ \ B\, for each premise \\ R t\ in an introduction rule. The default rule declarations of Isabelle/HOL already take care of most common situations. \<^rail>\ (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} | @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set}) @{syntax vars} @{syntax for_fixes} \ (@'where' @{syntax multi_specs})? (@'monos' @{syntax thms})? ; @@{command print_inductives} ('!'?) ; @@{attribute (HOL) mono} (() | 'add' | 'del') \ \<^descr> @{command (HOL) "inductive"} and @{command (HOL) "coinductive"} define (co)inductive predicates from the introduction rules. The propositions given as \clauses\ in the @{keyword "where"} part are either rules of the usual \\/\\ format (with arbitrary nesting), or equalities using \\\. The latter specifies extra-logical abbreviations in the sense of @{command_ref abbreviation}. Introducing abstract syntax simultaneously with the actual introduction rules is occasionally useful for complex specifications. The optional @{keyword "for"} part contains a list of parameters of the (co)inductive predicates that remain fixed throughout the definition, in contrast to arguments of the relation that may vary in each occurrence within the given \clauses\. The optional @{keyword "monos"} declaration contains additional \<^emph>\monotonicity theorems\, which are required for each operator applied to a recursive set in the introduction rules. \<^descr> @{command (HOL) "inductive_set"} and @{command (HOL) "coinductive_set"} are wrappers for to the previous commands for native HOL predicates. This allows to define (co)inductive sets, where multiple arguments are simulated via tuples. \<^descr> @{command "print_inductives"} prints (co)inductive definitions and monotonicity rules; the ``\!\'' option indicates extra verbosity. \<^descr> @{attribute (HOL) mono} declares monotonicity rules in the context. These rule are involved in the automated monotonicity proof of the above inductive and coinductive definitions. \ subsection \Derived rules\ text \ A (co)inductive definition of \R\ provides the following main theorems: \<^descr> \R.intros\ is the list of introduction rules as proven theorems, for the recursive predicates (or sets). The rules are also available individually, using the names given them in the theory file; \<^descr> \R.cases\ is the case analysis (or elimination) rule; \<^descr> \R.induct\ or \R.coinduct\ is the (co)induction rule; \<^descr> \R.simps\ is the equation unrolling the fixpoint of the predicate one step. When several predicates \R\<^sub>1, \, R\<^sub>n\ are defined simultaneously, the list of introduction rules is called \R\<^sub>1_\_R\<^sub>n.intros\, the case analysis rules are called \R\<^sub>1.cases, \, R\<^sub>n.cases\, and the list of mutual induction rules is called \R\<^sub>1_\_R\<^sub>n.inducts\. \ subsection \Monotonicity theorems\ text \ The context maintains a default set of theorems that are used in monotonicity proofs. New rules can be declared via the @{attribute (HOL) mono} attribute. See the main Isabelle/HOL sources for some examples. The general format of such monotonicity theorems is as follows: \<^item> Theorems of the form \A \ B \ \ A \ \ B\, for proving monotonicity of inductive definitions whose introduction rules have premises involving terms such as \\ R t\. \<^item> Monotonicity theorems for logical operators, which are of the general form \(\ \ \) \ \ (\ \ \) \ \ \ \\. For example, in the case of the operator \\\, the corresponding theorem is \[ \infer{\P\<^sub>1 \ P\<^sub>2 \ Q\<^sub>1 \ Q\<^sub>2\}{\P\<^sub>1 \ Q\<^sub>1\ & \P\<^sub>2 \ Q\<^sub>2\} \] \<^item> De Morgan style equations for reasoning about the ``polarity'' of expressions, e.g. \[ \<^prop>\\ \ P \ P\ \qquad\qquad \<^prop>\\ (P \ Q) \ \ P \ \ Q\ \] \<^item> Equations for reducing complex operators to more primitive ones whose monotonicity can easily be proved, e.g. \[ \<^prop>\(P \ Q) \ \ P \ Q\ \qquad\qquad \<^prop>\Ball A P \ \x. x \ A \ P x\ \] \ subsubsection \Examples\ text \The finite powerset operator can be defined inductively like this:\ (*<*)experiment begin(*>*) inductive_set Fin :: "'a set \ 'a set set" for A :: "'a set" where empty: "{} \ Fin A" | insert: "a \ A \ B \ Fin A \ insert a B \ Fin A" text \The accessible part of a relation is defined as follows:\ inductive acc :: "('a \ 'a \ bool) \ 'a \ bool" for r :: "'a \ 'a \ bool" (infix "\" 50) where acc: "(\y. y \ x \ acc r y) \ acc r x" (*<*)end(*>*) text \ Common logical connectives can be easily characterized as non-recursive inductive definitions with parameters, but without arguments. \ (*<*)experiment begin(*>*) inductive AND for A B :: bool where "A \ B \ AND A B" inductive OR for A B :: bool where "A \ OR A B" | "B \ OR A B" inductive EXISTS for B :: "'a \ bool" where "B a \ EXISTS B" (*<*)end(*>*) text \ Here the \cases\ or \induct\ rules produced by the @{command inductive} package coincide with the expected elimination rules for Natural Deduction. - Already in the original article by Gerhard Gentzen @{cite "Gentzen:1935"} + Already in the original article by Gerhard Gentzen \<^cite>\"Gentzen:1935"\ there is a hint that each connective can be characterized by its introductions, and the elimination can be constructed systematically. \ section \Recursive functions \label{sec:recursion}\ text \ \begin{matharray}{rcl} @{command_def (HOL) "primrec"} & : & \local_theory \ local_theory\ \\ @{command_def (HOL) "fun"} & : & \local_theory \ local_theory\ \\ @{command_def (HOL) "function"} & : & \local_theory \ proof(prove)\ \\ @{command_def (HOL) "termination"} & : & \local_theory \ proof(prove)\ \\ @{command_def (HOL) "fun_cases"} & : & \local_theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command (HOL) primrec} @{syntax specification} ; (@@{command (HOL) fun} | @@{command (HOL) function}) opts? @{syntax specification} ; opts: '(' (('sequential' | 'domintros') + ',') ')' ; @@{command (HOL) termination} @{syntax term}? ; @@{command (HOL) fun_cases} (@{syntax thmdecl}? @{syntax prop} + @'and') \ \<^descr> @{command (HOL) "primrec"} defines primitive recursive functions over datatypes (see also @{command_ref (HOL) datatype}). The given \equations\ specify reduction rules that are produced by instantiating the generic combinator for primitive recursion that is available for each datatype. Each equation needs to be of the form: @{text [display] "f x\<^sub>1 \ x\<^sub>m (C y\<^sub>1 \ y\<^sub>k) z\<^sub>1 \ z\<^sub>n = rhs"} such that \C\ is a datatype constructor, \rhs\ contains only the free variables on the left-hand side (or from the context), and all recursive occurrences of \f\ in \rhs\ are of the form \f \ y\<^sub>i \\ for some \i\. At most one reduction rule for each constructor can be given. The order does not matter. For missing constructors, the function is defined to return a default value, but this equation is made difficult to access for users. The reduction rules are declared as @{attribute simp} by default, which enables standard proof methods like @{method simp} and @{method auto} to normalize expressions of \f\ applied to datatype constructions, by simulating symbolic computation via rewriting. \<^descr> @{command (HOL) "function"} defines functions by general wellfounded - recursion. A detailed description with examples can be found in @{cite - "isabelle-function"}. The function is specified by a set of (possibly + recursion. A detailed description with examples can be found in \<^cite>\"isabelle-function"\. The function is specified by a set of (possibly conditional) recursive equations with arbitrary pattern matching. The command generates proof obligations for the completeness and the compatibility of patterns. The defined function is considered partial, and the resulting simplification rules (named \f.psimps\) and induction rule (named \f.pinduct\) are guarded by a generated domain predicate \f_dom\. The @{command (HOL) "termination"} command can then be used to establish that the function is total. \<^descr> @{command (HOL) "fun"} is a shorthand notation for ``@{command (HOL) "function"}~\(sequential)\'', followed by automated proof attempts regarding - pattern matching and termination. See @{cite "isabelle-function"} for + pattern matching and termination. See \<^cite>\"isabelle-function"\ for further details. \<^descr> @{command (HOL) "termination"}~\f\ commences a termination proof for the previously defined function \f\. If this is omitted, the command refers to the most recent function definition. After the proof is closed, the recursive equations and the induction principle is established. \<^descr> @{command (HOL) "fun_cases"} generates specialized elimination rules for function equations. It expects one or more function equations and produces rules that eliminate the given equalities, following the cases given in the function definition. Recursive definitions introduced by the @{command (HOL) "function"} command accommodate reasoning by induction (cf.\ @{method induct}): rule \f.induct\ refers to a specific induction rule, with parameters named according to the user-specified equations. Cases are numbered starting from 1. For @{command (HOL) "primrec"}, the induction principle coincides with structural recursion on the datatype where the recursion is carried out. The equations provided by these packages may be referred later as theorem list \f.simps\, where \f\ is the (collective) name of the functions defined. Individual equations may be named explicitly as well. The @{command (HOL) "function"} command accepts the following options. \<^descr> \sequential\ enables a preprocessor which disambiguates overlapping patterns by making them mutually disjoint. Earlier equations take precedence over later ones. This allows to give the specification in a format very similar to functional programming. Note that the resulting simplification and induction rules correspond to the transformed specification, not the one given originally. This usually means that each equation given by the user may result in several theorems. Also note that this automatic transformation only works for ML-style datatype patterns. \<^descr> \domintros\ enables the automated generation of introduction rules for the domain predicate. While mostly not needed, they can be helpful in some proofs about partial functions. \ subsubsection \Example: evaluation of expressions\ text \ Subsequently, we define mutual datatypes for arithmetic and boolean expressions, and use @{command primrec} for evaluation functions that follow the same recursive structure. \ (*<*)experiment begin(*>*) datatype 'a aexp = IF "'a bexp" "'a aexp" "'a aexp" | Sum "'a aexp" "'a aexp" | Diff "'a aexp" "'a aexp" | Var 'a | Num nat and 'a bexp = Less "'a aexp" "'a aexp" | And "'a bexp" "'a bexp" | Neg "'a bexp" text \\<^medskip> Evaluation of arithmetic and boolean expressions\ primrec evala :: "('a \ nat) \ 'a aexp \ nat" and evalb :: "('a \ nat) \ 'a bexp \ bool" where "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)" | "evala env (Sum a1 a2) = evala env a1 + evala env a2" | "evala env (Diff a1 a2) = evala env a1 - evala env a2" | "evala env (Var v) = env v" | "evala env (Num n) = n" | "evalb env (Less a1 a2) = (evala env a1 < evala env a2)" | "evalb env (And b1 b2) = (evalb env b1 \ evalb env b2)" | "evalb env (Neg b) = (\ evalb env b)" text \ Since the value of an expression depends on the value of its variables, the functions \<^const>\evala\ and \<^const>\evalb\ take an additional parameter, an \<^emph>\environment\ that maps variables to their values. \<^medskip> Substitution on expressions can be defined similarly. The mapping \f\ of type \<^typ>\'a \ 'a aexp\ given as a parameter is lifted canonically on the types \<^typ>\'a aexp\ and \<^typ>\'a bexp\, respectively. \ primrec substa :: "('a \ 'b aexp) \ 'a aexp \ 'b aexp" and substb :: "('a \ 'b aexp) \ 'a bexp \ 'b bexp" where "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)" | "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)" | "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)" | "substa f (Var v) = f v" | "substa f (Num n) = Num n" | "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)" | "substb f (And b1 b2) = And (substb f b1) (substb f b2)" | "substb f (Neg b) = Neg (substb f b)" text \ In textbooks about semantics one often finds substitution theorems, which express the relationship between substitution and evaluation. For \<^typ>\'a aexp\ and \<^typ>\'a bexp\, we can prove such a theorem by mutual induction, followed by simplification. \ lemma subst_one: "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a" "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b" by (induct a and b) simp_all lemma subst_all: "evala env (substa s a) = evala (\x. evala env (s x)) a" "evalb env (substb s b) = evalb (\x. evala env (s x)) b" by (induct a and b) simp_all (*<*)end(*>*) subsubsection \Example: a substitution function for terms\ text \Functions on datatypes with nested recursion are also defined by mutual primitive recursion.\ (*<*)experiment begin(*>*) datatype ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list" text \ A substitution function on type \<^typ>\('a, 'b) term\ can be defined as follows, by working simultaneously on \<^typ>\('a, 'b) term list\: \ primrec subst_term :: "('a \ ('a, 'b) term) \ ('a, 'b) term \ ('a, 'b) term" and subst_term_list :: "('a \ ('a, 'b) term) \ ('a, 'b) term list \ ('a, 'b) term list" where "subst_term f (Var a) = f a" | "subst_term f (App b ts) = App b (subst_term_list f ts)" | "subst_term_list f [] = []" | "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts" text \ The recursion scheme follows the structure of the unfolded definition of type \<^typ>\('a, 'b) term\. To prove properties of this substitution function, mutual induction is needed: \ lemma "subst_term (subst_term f1 \ f2) t = subst_term f1 (subst_term f2 t)" and "subst_term_list (subst_term f1 \ f2) ts = subst_term_list f1 (subst_term_list f2 ts)" by (induct t and ts rule: subst_term.induct subst_term_list.induct) simp_all (*<*)end(*>*) subsubsection \Example: a map function for infinitely branching trees\ text \Defining functions on infinitely branching datatypes by primitive recursion is just as easy.\ (*<*)experiment begin(*>*) datatype 'a tree = Atom 'a | Branch "nat \ 'a tree" primrec map_tree :: "('a \ 'b) \ 'a tree \ 'b tree" where "map_tree f (Atom a) = Atom (f a)" | "map_tree f (Branch ts) = Branch (\x. map_tree f (ts x))" text \ Note that all occurrences of functions such as \ts\ above must be applied to an argument. In particular, \<^term>\map_tree f \ ts\ is not allowed here. \<^medskip> Here is a simple composition lemma for \<^term>\map_tree\: \ lemma "map_tree g (map_tree f t) = map_tree (g \ f) t" by (induct t) simp_all (*<*)end(*>*) subsection \Proof methods related to recursive definitions\ text \ \begin{matharray}{rcl} @{method_def (HOL) pat_completeness} & : & \method\ \\ @{method_def (HOL) relation} & : & \method\ \\ @{method_def (HOL) lexicographic_order} & : & \method\ \\ @{method_def (HOL) size_change} & : & \method\ \\ @{attribute_def (HOL) termination_simp} & : & \attribute\ \\ @{method_def (HOL) induction_schema} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method (HOL) relation} @{syntax term} ; @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * ) ; @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) ) ; @@{method (HOL) induction_schema} ; orders: ( 'max' | 'min' | 'ms' ) * \ \<^descr> @{method (HOL) pat_completeness} is a specialized method to solve goals regarding the completeness of pattern matching, as required by the @{command - (HOL) "function"} package (cf.\ @{cite "isabelle-function"}). + (HOL) "function"} package (cf.\ \<^cite>\"isabelle-function"\). \<^descr> @{method (HOL) relation}~\R\ introduces a termination proof using the relation \R\. The resulting proof state will contain goals expressing that \R\ is wellfounded, and that the arguments of recursive calls decrease with respect to \R\. Usually, this method is used as the initial proof step of manual termination proofs. \<^descr> @{method (HOL) "lexicographic_order"} attempts a fully automated termination proof by searching for a lexicographic combination of size measures on the arguments of the function. The method accepts the same arguments as the @{method auto} method, which it uses internally to prove local descents. The @{syntax clasimpmod} modifiers are accepted (as for @{method auto}). In case of failure, extensive information is printed, which can help to - analyse the situation (cf.\ @{cite "isabelle-function"}). + analyse the situation (cf.\ \<^cite>\"isabelle-function"\). \<^descr> @{method (HOL) "size_change"} also works on termination goals, using a variation of the size-change principle, together with a graph decomposition - technique (see @{cite krauss_phd} for details). Three kinds of orders are + technique (see \<^cite>\krauss_phd\ for details). Three kinds of orders are used internally: \max\, \min\, and \ms\ (multiset), which is only available when the theory \Multiset\ is loaded. When no order kinds are given, they are tried in order. The search for a termination proof uses SAT solving internally. For local descent proofs, the @{syntax clasimpmod} modifiers are accepted (as for @{method auto}). \<^descr> @{attribute (HOL) termination_simp} declares extra rules for the simplifier, when invoked in termination proofs. This can be useful, e.g., for special rules involving size estimations. \<^descr> @{method (HOL) induction_schema} derives user-specified induction rules from well-founded induction and completeness of patterns. This factors out some operations that are done internally by the function package and makes them available separately. See \<^file>\~~/src/HOL/Examples/Induction_Schema.thy\ for examples. \ subsection \Functions with explicit partiality\ text \ \begin{matharray}{rcl} @{command_def (HOL) "partial_function"} & : & \local_theory \ local_theory\ \\ @{attribute_def (HOL) "partial_function_mono"} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{command (HOL) partial_function} '(' @{syntax name} ')' @{syntax specification} \ \<^descr> @{command (HOL) "partial_function"}~\(mode)\ defines recursive functions based on fixpoints in complete partial orders. No termination proof is required from the user or constructed internally. Instead, the possibility of non-termination is modelled explicitly in the result type, which contains an explicit bottom element. Pattern matching and mutual recursion are currently not supported. Thus, the specification consists of a single function described by a single recursive equation. There are no fixed syntactic restrictions on the body of the function, but the induced functional must be provably monotonic wrt.\ the underlying order. The monotonicity proof is performed internally, and the definition is rejected when it fails. The proof can be influenced by declaring hints using the @{attribute (HOL) partial_function_mono} attribute. The mandatory \mode\ argument specifies the mode of operation of the command, which directly corresponds to a complete partial order on the result type. By default, the following modes are defined: \<^descr> \option\ defines functions that map into the \<^type>\option\ type. Here, the value \<^term>\None\ is used to model a non-terminating computation. Monotonicity requires that if \<^term>\None\ is returned by a recursive call, then the overall result must also be \<^term>\None\. This is best achieved through the use of the monadic operator \<^const>\Option.bind\. \<^descr> \tailrec\ defines functions with an arbitrary result type and uses the slightly degenerated partial order where \<^term>\undefined\ is the bottom element. Now, monotonicity requires that if \<^term>\undefined\ is returned by a recursive call, then the overall result must also be \<^term>\undefined\. In practice, this is only satisfied when each recursive call is a tail call, whose result is directly returned. Thus, this mode of operation allows the definition of arbitrary tail-recursive functions. Experienced users may define new modes by instantiating the locale \<^const>\partial_function_definitions\ appropriately. \<^descr> @{attribute (HOL) partial_function_mono} declares rules for use in the internal monotonicity proofs of partial function definitions. \ subsection \Old-style recursive function definitions (TFL)\ text \ \begin{matharray}{rcl} @{command_def (HOL) "recdef"} & : & \theory \ theory)\ \\ \end{matharray} The old TFL command @{command (HOL) "recdef"} for defining recursive is mostly obsolete; @{command (HOL) "function"} or @{command (HOL) "fun"} should be used instead. \<^rail>\ @@{command (HOL) recdef} ('(' @'permissive' ')')? \ @{syntax name} @{syntax term} (@{syntax prop} +) hints? ; hints: '(' @'hints' ( recdefmod * ) ')' ; recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') (() | 'add' | 'del') ':' @{syntax thms}) | @{syntax clasimpmod} \ \<^descr> @{command (HOL) "recdef"} defines general well-founded recursive functions (using the TFL package). The ``\(permissive)\'' option tells TFL to recover from failed proof attempts, returning unfinished results. The \recdef_simp\, \recdef_cong\, and \recdef_wf\ hints refer to auxiliary rules to be used in the internal automated proof process of TFL. Additional @{syntax clasimpmod} declarations may be given to tune the context of the Simplifier (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\ \secref{sec:classical}). \<^medskip> Hints for @{command (HOL) "recdef"} may be also declared globally, using the following attributes. \begin{matharray}{rcl} @{attribute_def (HOL) recdef_simp} & : & \attribute\ \\ @{attribute_def (HOL) recdef_cong} & : & \attribute\ \\ @{attribute_def (HOL) recdef_wf} & : & \attribute\ \\ \end{matharray} \<^rail>\ (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} | @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del') \ \ section \Adhoc overloading of constants\ text \ \begin{tabular}{rcll} @{command_def "adhoc_overloading"} & : & \local_theory \ local_theory\ \\ @{command_def "no_adhoc_overloading"} & : & \local_theory \ local_theory\ \\ @{attribute_def "show_variants"} & : & \attribute\ & default \false\ \\ \end{tabular} \<^medskip> Adhoc overloading allows to overload a constant depending on its type. Typically this involves the introduction of an uninterpreted constant (used for input and output) and the addition of some variants (used internally). For examples see \<^file>\~~/src/HOL/Examples/Adhoc_Overloading_Examples.thy\ and \<^file>\~~/src/HOL/Library/Monad_Syntax.thy\. \<^rail>\ (@@{command adhoc_overloading} | @@{command no_adhoc_overloading}) (@{syntax name} (@{syntax term} + ) + @'and') \ \<^descr> @{command "adhoc_overloading"}~\c v\<^sub>1 ... v\<^sub>n\ associates variants with an existing constant. \<^descr> @{command "no_adhoc_overloading"} is similar to @{command "adhoc_overloading"}, but removes the specified variants from the present context. \<^descr> @{attribute "show_variants"} controls printing of variants of overloaded constants. If enabled, the internally used variants are printed instead of their respective overloaded constants. This is occasionally useful to check whether the system agrees with a user's expectations about derived variants. \ section \Definition by specification \label{sec:hol-specification}\ text \ \begin{matharray}{rcl} @{command_def (HOL) "specification"} & : & \theory \ proof(prove)\ \\ \end{matharray} \<^rail>\ @@{command (HOL) specification} '(' (decl +) ')' \ (@{syntax thmdecl}? @{syntax prop} +) ; decl: (@{syntax name} ':')? @{syntax term} ('(' @'overloaded' ')')? \ \<^descr> @{command (HOL) "specification"}~\decls \\ sets up a goal stating the existence of terms with the properties specified to hold for the constants given in \decls\. After finishing the proof, the theory will be augmented with definitions for the given constants, as well as with theorems stating the properties for these constants. \decl\ declares a constant to be defined by the specification given. The definition for the constant \c\ is bound to the name \c_def\ unless a theorem name is given in the declaration. Overloaded constants should be declared as such. \ section \Old-style datatypes \label{sec:hol-datatype}\ text \ \begin{matharray}{rcl} @{command_def (HOL) "old_rep_datatype"} & : & \theory \ proof(prove)\ \\ \end{matharray} \<^rail>\ @@{command (HOL) old_rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +) ; spec: @{syntax typespec_sorts} @{syntax mixfix}? '=' (cons + '|') ; cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}? \ \<^descr> @{command (HOL) "old_rep_datatype"} represents existing types as old-style datatypes. These commands are mostly obsolete; @{command (HOL) "datatype"} should be used instead. - See @{cite "isabelle-datatypes"} for more details on datatypes. Apart from + See \<^cite>\"isabelle-datatypes"\ for more details on datatypes. Apart from proper proof methods for case analysis and induction, there are also emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL) induct_tac} available, see \secref{sec:hol-induct-tac}; these admit to refer directly to the internal structure of subgoals (including internally bound parameters). \ subsubsection \Examples\ text \ We define a type of finite sequences, with slightly different names than the existing \<^typ>\'a list\ that is already in \<^theory>\Main\: \ (*<*)experiment begin(*>*) datatype 'a seq = Empty | Seq 'a "'a seq" text \We can now prove some simple lemma by structural induction:\ lemma "Seq x xs \ xs" proof (induct xs arbitrary: x) case Empty txt \This case can be proved using the simplifier: the freeness properties of the datatype are already declared as @{attribute simp} rules.\ show "Seq x Empty \ Empty" by simp next case (Seq y ys) txt \The step case is proved similarly.\ show "Seq x (Seq y ys) \ Seq y ys" using \Seq y ys \ ys\ by simp qed text \Here is a more succinct version of the same proof:\ lemma "Seq x xs \ xs" by (induct xs arbitrary: x) simp_all (*<*)end(*>*) section \Records \label{sec:hol-record}\ text \ In principle, records merely generalize the concept of tuples, where components may be addressed by labels instead of just position. The logical infrastructure of records in Isabelle/HOL is slightly more advanced, though, supporting truly extensible record schemes. This admits operations that are polymorphic with respect to record extension, yielding ``object-oriented'' - effects like (single) inheritance. See also @{cite "NaraschewskiW-TPHOLs98"} + effects like (single) inheritance. See also \<^cite>\"NaraschewskiW-TPHOLs98"\ for more details on object-oriented verification and record subtyping in HOL. \ subsection \Basic concepts\ text \ Isabelle/HOL supports both \<^emph>\fixed\ and \<^emph>\schematic\ records at the level of terms and types. The notation is as follows: \begin{center} \begin{tabular}{l|l|l} & record terms & record types \\ \hline fixed & \\x = a, y = b\\ & \\x :: A, y :: B\\ \\ schematic & \\x = a, y = b, \ = m\\ & \\x :: A, y :: B, \ :: M\\ \\ \end{tabular} \end{center} The ASCII representation of \\x = a\\ is \(| x = a |)\. A fixed record \\x = a, y = b\\ has field \x\ of value \a\ and field \y\ of value \b\. The corresponding type is \\x :: A, y :: B\\, assuming that \a :: A\ and \b :: B\. A record scheme like \\x = a, y = b, \ = m\\ contains fields \x\ and \y\ as before, but also possibly further fields as indicated by the ``\\\'' notation (which is actually part of the syntax). The improper field ``\\\'' of a record scheme is called the \<^emph>\more part\. Logically it is just a free variable, which is occasionally referred to as ``row variable'' in the literature. The more part of a record scheme may be instantiated by zero or more further components. For example, the previous scheme may get instantiated to \\x = a, y = b, z = c, \ = m'\\, where \m'\ refers to a different more part. Fixed records are special instances of record schemes, where ``\\\'' is properly terminated by the \() :: unit\ element. In fact, \\x = a, y = b\\ is just an abbreviation for \\x = a, y = b, \ = ()\\. \<^medskip> Two key observations make extensible records in a simply typed language like HOL work out: \<^enum> the more part is internalized, as a free term or type variable, \<^enum> field names are externalized, they cannot be accessed within the logic as first-class values. \<^medskip> In Isabelle/HOL record types have to be defined explicitly, fixing their field names and types, and their (optional) parent record. Afterwards, records may be formed using above syntax, while obeying the canonical order of fields as given by their declaration. The record package provides several standard operations like selectors and updates. The common setup for various generic proof tools enable succinct reasoning patterns. See also the - Isabelle/HOL tutorial @{cite "isabelle-hol-book"} for further instructions + Isabelle/HOL tutorial \<^cite>\"isabelle-hol-book"\ for further instructions on using records in practice. \ subsection \Record specifications\ text \ \begin{matharray}{rcl} @{command_def (HOL) "record"} & : & \theory \ theory\ \\ @{command_def (HOL) "print_record"} & : & \context \\ \\ \end{matharray} \<^rail>\ @@{command (HOL) record} @{syntax "overloaded"}? @{syntax typespec_sorts} '=' \ (@{syntax type} '+')? (constdecl +) ; constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}? ; @@{command (HOL) print_record} modes? @{syntax typespec_sorts} ; modes: '(' (@{syntax name} +) ')' \ \<^descr> @{command (HOL) "record"}~\(\\<^sub>1, \, \\<^sub>m) t = \ + c\<^sub>1 :: \\<^sub>1 \ c\<^sub>n :: \\<^sub>n\ defines extensible record type \(\\<^sub>1, \, \\<^sub>m) t\, derived from the optional parent record \\\ by adding new field components \c\<^sub>i :: \\<^sub>i\ etc. The type variables of \\\ and \\\<^sub>i\ need to be covered by the (distinct) parameters \\\<^sub>1, \, \\<^sub>m\. Type constructor \t\ has to be new, while \\\ needs to specify an instance of an existing record type. At least one new field \c\<^sub>i\ has to be specified. Basically, field names need to belong to a unique record. This is not a real restriction in practice, since fields are qualified by the record name internally. The parent record specification \\\ is optional; if omitted \t\ becomes a root record. The hierarchy of all records declared within a theory context forms a forest structure, i.e.\ a set of trees starting with a root record each. There is no way to merge multiple parent records! For convenience, \(\\<^sub>1, \, \\<^sub>m) t\ is made a type abbreviation for the fixed record type \\c\<^sub>1 :: \\<^sub>1, \, c\<^sub>n :: \\<^sub>n\\, likewise is \(\\<^sub>1, \, \\<^sub>m, \) t_scheme\ made an abbreviation for \\c\<^sub>1 :: \\<^sub>1, \, c\<^sub>n :: \\<^sub>n, \ :: \\\. \<^descr> @{command (HOL) "print_record"}~\(\\<^sub>1, \, \\<^sub>m) t\ prints the definition of record \(\\<^sub>1, \, \\<^sub>m) t\. Optionally \modes\ can be specified, which are appended to the current print mode; see \secref{sec:print-modes}. \ subsection \Record operations\ text \ Any record definition of the form presented above produces certain standard operations. Selectors and updates are provided for any field, including the improper one ``\more\''. There are also cumulative record constructor functions. To simplify the presentation below, we assume for now that \(\\<^sub>1, \, \\<^sub>m) t\ is a root record with fields \c\<^sub>1 :: \\<^sub>1, \, c\<^sub>n :: \\<^sub>n\. \<^medskip> \<^bold>\Selectors\ and \<^bold>\updates\ are available for any field (including ``\more\''): \begin{matharray}{lll} \c\<^sub>i\ & \::\ & \\\<^vec>c :: \<^vec>\, \ :: \\ \ \\<^sub>i\ \\ \c\<^sub>i_update\ & \::\ & \\\<^sub>i \ \\<^vec>c :: \<^vec>\, \ :: \\ \ \\<^vec>c :: \<^vec>\, \ :: \\\ \\ \end{matharray} There is special syntax for application of updates: \r\x := a\\ abbreviates term \x_update a r\. Further notation for repeated updates is also available: \r\x := a\\y := b\\z := c\\ may be written \r\x := a, y := b, z := c\\. Note that because of postfix notation the order of fields shown here is reverse than in the actual term. Since repeated updates are just function applications, fields may be freely permuted in \\x := a, y := b, z := c\\, as far as logical equality is concerned. Thus commutativity of independent updates can be proven within the logic for any two fields, but not as a general theorem. \<^medskip> The \<^bold>\make\ operation provides a cumulative record constructor function: \begin{matharray}{lll} \t.make\ & \::\ & \\\<^sub>1 \ \ \\<^sub>n \ \\<^vec>c :: \<^vec>\\\ \\ \end{matharray} \<^medskip> We now reconsider the case of non-root records, which are derived of some parent. In general, the latter may depend on another parent as well, resulting in a list of \<^emph>\ancestor records\. Appending the lists of fields of all ancestors results in a certain field prefix. The record package automatically takes care of this by lifting operations over this context of ancestor fields. Assuming that \(\\<^sub>1, \, \\<^sub>m) t\ has ancestor fields \b\<^sub>1 :: \\<^sub>1, \, b\<^sub>k :: \\<^sub>k\, the above record operations will get the following types: \<^medskip> \begin{tabular}{lll} \c\<^sub>i\ & \::\ & \\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ \\<^sub>i\ \\ \c\<^sub>i_update\ & \::\ & \\\<^sub>i \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\\ \\ \t.make\ & \::\ & \\\<^sub>1 \ \ \\<^sub>k \ \\<^sub>1 \ \ \\<^sub>n \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\\ \\ \end{tabular} \<^medskip> Some further operations address the extension aspect of a derived record scheme specifically: \t.fields\ produces a record fragment consisting of exactly the new fields introduced here (the result may serve as a more part elsewhere); \t.extend\ takes a fixed record and adds a given more part; \t.truncate\ restricts a record scheme to a fixed record. \<^medskip> \begin{tabular}{lll} \t.fields\ & \::\ & \\\<^sub>1 \ \ \\<^sub>n \ \\<^vec>c :: \<^vec>\\\ \\ \t.extend\ & \::\ & \\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\ \ \ \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\\ \\ \t.truncate\ & \::\ & \\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\\ \\ \end{tabular} \<^medskip> Note that \t.make\ and \t.fields\ coincide for root records. \ subsection \Derived rules and proof tools\ text \ The record package proves several results internally, declaring these facts to appropriate proof tools. This enables users to reason about record structures quite conveniently. Assume that \t\ is a record type as specified above. \<^enum> Standard conversions for selectors or updates applied to record constructor terms are made part of the default Simplifier context; thus proofs by reduction of basic operations merely require the @{method simp} method without further arguments. These rules are available as \t.simps\, too. \<^enum> Selectors applied to updated records are automatically reduced by an internal simplification procedure, which is also part of the standard Simplifier setup. \<^enum> Inject equations of a form analogous to \<^prop>\(x, y) = (x', y') \ x = x' \ y = y'\ are declared to the Simplifier and Classical Reasoner as @{attribute iff} rules. These rules are available as \t.iffs\. \<^enum> The introduction rule for record equality analogous to \x r = x r' \ y r = y r' \ \ r = r'\ is declared to the Simplifier, and as the basic rule context as ``@{attribute intro}\?\''. The rule is called \t.equality\. \<^enum> Representations of arbitrary record expressions as canonical constructor terms are provided both in @{method cases} and @{method induct} format (cf.\ the generic proof methods of the same name, \secref{sec:cases-induct}). Several variations are available, for fixed records, record schemes, more parts etc. The generic proof methods are sufficiently smart to pick the most sensible rule according to the type of the indicated record expression: users just need to apply something like ``\(cases r)\'' to a certain proof problem. \<^enum> The derived record operations \t.make\, \t.fields\, \t.extend\, \t.truncate\ are \<^emph>\not\ treated automatically, but usually need to be expanded by hand, using the collective fact \t.defs\. \ subsubsection \Examples\ text \See \<^file>\~~/src/HOL/Examples/Records.thy\, for example.\ section \Semantic subtype definitions \label{sec:hol-typedef}\ text \ \begin{matharray}{rcl} @{command_def (HOL) "typedef"} & : & \local_theory \ proof(prove)\ \\ \end{matharray} A type definition identifies a new type with a non-empty subset of an existing type. More precisely, the new type is defined by exhibiting an existing type \\\, a set \A :: \ set\, and proving \<^prop>\\x. x \ A\. Thus \A\ is a non-empty subset of \\\, and the new type denotes this subset. New functions are postulated that establish an isomorphism between the new type and the subset. In general, the type \\\ may involve type variables \\\<^sub>1, \, \\<^sub>n\ which means that the type definition produces a type constructor \(\\<^sub>1, \, \\<^sub>n) t\ depending on those type arguments. \<^rail>\ @@{command (HOL) typedef} @{syntax "overloaded"}? abs_type '=' rep_set ; @{syntax_def "overloaded"}: ('(' @'overloaded' ')') ; abs_type: @{syntax typespec_sorts} @{syntax mixfix}? ; rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})? \ To understand the concept of type definition better, we need to recount its somewhat complex history. The HOL logic goes back to the ``Simple Theory of - Types'' (STT) of A. Church @{cite "church40"}, which is further explained in - the book by P. Andrews @{cite "andrews86"}. The overview article by W. - Farmer @{cite "Farmer:2008"} points out the ``seven virtues'' of this + Types'' (STT) of A. Church \<^cite>\"church40"\, which is further explained in + the book by P. Andrews \<^cite>\"andrews86"\. The overview article by W. + Farmer \<^cite>\"Farmer:2008"\ points out the ``seven virtues'' of this relatively simple family of logics. STT has only ground types, without polymorphism and without type definitions. \<^medskip> - M. Gordon @{cite "Gordon:1985:HOL"} augmented Church's STT by adding + M. Gordon \<^cite>\"Gordon:1985:HOL"\ augmented Church's STT by adding schematic polymorphism (type variables and type constructors) and a facility to introduce new types as semantic subtypes from existing types. This genuine extension of the logic was explained semantically by A. Pitts in the - book of the original Cambridge HOL88 system @{cite "pitts93"}. Type + book of the original Cambridge HOL88 system \<^cite>\"pitts93"\. Type definitions work in this setting, because the general model-theory of STT is restricted to models that ensure that the universe of type interpretations is closed by forming subsets (via predicates taken from the logic). \<^medskip> Isabelle/HOL goes beyond Gordon-style HOL by admitting overloaded constant - definitions @{cite "Wenzel:1997:TPHOL" and "Haftmann-Wenzel:2006:classes"}, + definitions \<^cite>\"Wenzel:1997:TPHOL" and "Haftmann-Wenzel:2006:classes"\, which are actually a concept of Isabelle/Pure and do not depend on particular set-theoretic semantics of HOL. Over many years, there was no formal checking of semantic type definitions in Isabelle/HOL versus syntactic constant definitions in Isabelle/Pure. So the @{command typedef} command was described as ``axiomatic'' in the sense of \secref{sec:axiomatizations}, only with some local checks of the given type and its representing set. - Recent clarification of overloading in the HOL logic proper @{cite - "Kuncar-Popescu:2015"} demonstrates how the dissimilar concepts of constant + Recent clarification of overloading in the HOL logic proper \<^cite>\"Kuncar-Popescu:2015"\ demonstrates how the dissimilar concepts of constant definitions versus type definitions may be understood uniformly. This requires an interpretation of Isabelle/HOL that substantially reforms the - set-theoretic model of A. Pitts @{cite "pitts93"}, by taking a schematic + set-theoretic model of A. Pitts \<^cite>\"pitts93"\, by taking a schematic view on polymorphism and interpreting only ground types in the set-theoretic sense of HOL88. Moreover, type-constructors may be explicitly overloaded, e.g.\ by making the subset depend on type-class parameters (cf.\ \secref{sec:class}). This is semantically like a dependent type: the meaning relies on the operations provided by different type-class instances. \<^descr> @{command (HOL) "typedef"}~\(\\<^sub>1, \, \\<^sub>n) t = A\ defines a new type \(\\<^sub>1, \, \\<^sub>n) t\ from the set \A\ over an existing type. The set \A\ may contain type variables \\\<^sub>1, \, \\<^sub>n\ as specified on the LHS, but no term variables. Non-emptiness of \A\ needs to be proven on the spot, in order to turn the internal conditional characterization into usable theorems. The ``\(overloaded)\'' option allows the @{command "typedef"} specification to depend on constants that are not (yet) specified and thus left open as parameters, e.g.\ type-class parameters. Within a local theory specification, the newly introduced type constructor cannot depend on parameters or assumptions of the context: this is syntactically impossible in HOL. The non-emptiness proof may formally depend on local assumptions, but this has little practical relevance. For @{command (HOL) "typedef"}~\t = A\ the newly introduced type \t\ is accompanied by a pair of morphisms to relate it to the representing set over the old type. By default, the injection from type to set is called \Rep_t\ and its inverse \Abs_t\: An explicit @{keyword (HOL) "morphisms"} specification allows to provide alternative names. The logical characterization of @{command typedef} uses the predicate of locale \<^const>\type_definition\ that is defined in Isabelle/HOL. Various basic consequences of that are instantiated accordingly, re-using the locale facts with names derived from the new type constructor. Thus the generic theorem @{thm type_definition.Rep} is turned into the specific \Rep_t\, for example. Theorems @{thm type_definition.Rep}, @{thm type_definition.Rep_inverse}, and @{thm type_definition.Abs_inverse} provide the most basic characterization as a corresponding injection/surjection pair (in both directions). The derived rules @{thm type_definition.Rep_inject} and @{thm type_definition.Abs_inject} provide a more convenient version of injectivity, suitable for automated proof tools (e.g.\ in declarations involving @{attribute simp} or @{attribute iff}). Furthermore, the rules @{thm type_definition.Rep_cases}~/ @{thm type_definition.Rep_induct}, and @{thm type_definition.Abs_cases}~/ @{thm type_definition.Abs_induct} provide alternative views on surjectivity. These rules are already declared as set or type rules for the generic @{method cases} and @{method induct} methods, respectively. \ subsubsection \Examples\ text \ The following trivial example pulls a three-element type into existence within the formal logical environment of Isabelle/HOL.\ (*<*)experiment begin(*>*) typedef three = "{(True, True), (True, False), (False, True)}" by blast definition "One = Abs_three (True, True)" definition "Two = Abs_three (True, False)" definition "Three = Abs_three (False, True)" lemma three_distinct: "One \ Two" "One \ Three" "Two \ Three" by (simp_all add: One_def Two_def Three_def Abs_three_inject) lemma three_cases: fixes x :: three obtains "x = One" | "x = Two" | "x = Three" by (cases x) (auto simp: One_def Two_def Three_def Abs_three_inject) (*<*)end(*>*) text \Note that such trivial constructions are better done with derived specification mechanisms such as @{command datatype}:\ (*<*)experiment begin(*>*) datatype three = One | Two | Three (*<*)end(*>*) text \This avoids re-doing basic definitions and proofs from the primitive @{command typedef} above.\ section \Functorial structure of types\ text \ \begin{matharray}{rcl} @{command_def (HOL) "functor"} & : & \local_theory \ proof(prove)\ \end{matharray} \<^rail>\ @@{command (HOL) functor} (@{syntax name} ':')? @{syntax term} \ \<^descr> @{command (HOL) "functor"}~\prefix: m\ allows to prove and register properties about the functorial structure of type constructors. These properties then can be used by other packages to deal with those type constructors in certain type constructions. Characteristic theorems are noted in the current local theory. By default, they are prefixed with the base name of the type constructor, an explicit prefix can be given alternatively. The given term \m\ is considered as \<^emph>\mapper\ for the corresponding type constructor and must conform to the following type pattern: \begin{matharray}{lll} \m\ & \::\ & \\\<^sub>1 \ \ \\<^sub>k \ (\<^vec>\\<^sub>n) t \ (\<^vec>\\<^sub>n) t\ \\ \end{matharray} where \t\ is the type constructor, \\<^vec>\\<^sub>n\ and \\<^vec>\\<^sub>n\ are distinct type variables free in the local theory and \\\<^sub>1\, \ldots, \\\<^sub>k\ is a subsequence of \\\<^sub>1 \ \\<^sub>1\, \\\<^sub>1 \ \\<^sub>1\, \ldots, \\\<^sub>n \ \\<^sub>n\, \\\<^sub>n \ \\<^sub>n\. \ section \Quotient types with lifting and transfer\ text \ The quotient package defines a new quotient type given a raw type and a partial equivalence relation (\secref{sec:quotient-type}). The package also historically includes automation for transporting definitions and theorems (\secref{sec:old-quotient}), but most of this automation was superseded by the Lifting (\secref{sec:lifting}) and Transfer (\secref{sec:transfer}) packages. \ subsection \Quotient type definition \label{sec:quotient-type}\ text \ \begin{matharray}{rcl} @{command_def (HOL) "quotient_type"} & : & \local_theory \ proof(prove)\\\ \end{matharray} \<^rail>\ @@{command (HOL) quotient_type} @{syntax "overloaded"}? \ @{syntax typespec} @{syntax mixfix}? '=' quot_type \ quot_morphisms? quot_parametric? ; quot_type: @{syntax type} '/' ('partial' ':')? @{syntax term} ; quot_morphisms: @'morphisms' @{syntax name} @{syntax name} ; quot_parametric: @'parametric' @{syntax thm} \ \<^descr> @{command (HOL) "quotient_type"} defines a new quotient type \\\. The injection from a quotient type to a raw type is called \rep_\\, its inverse \abs_\\ unless explicit @{keyword (HOL) "morphisms"} specification provides alternative names. @{command (HOL) "quotient_type"} requires the user to prove that the relation is an equivalence relation (predicate \equivp\), unless the user specifies explicitly \partial\ in which case the obligation is \part_equivp\. A quotient defined with \partial\ is weaker in the sense that less things can be proved automatically. The command internally proves a Quotient theorem and sets up the Lifting package by the command @{command (HOL) setup_lifting}. Thus the Lifting and Transfer packages can be used also with quotient types defined by @{command (HOL) "quotient_type"} without any extra set-up. The parametricity theorem for the equivalence relation R can be provided as an extra argument of the command and is passed to the corresponding internal call of @{command (HOL) setup_lifting}. This theorem allows the Lifting package to generate a stronger transfer rule for equality. \ subsection \Lifting package \label{sec:lifting}\ text \ The Lifting package allows users to lift terms of the raw type to the abstract type, which is a necessary step in building a library for an abstract type. Lifting defines a new constant by combining coercion functions (\<^term>\Abs\ and \<^term>\Rep\) with the raw term. It also proves an appropriate transfer rule for the Transfer (\secref{sec:transfer}) package and, if possible, an equation for the code generator. The Lifting package provides two main commands: @{command (HOL) "setup_lifting"} for initializing the package to work with a new type, and @{command (HOL) "lift_definition"} for lifting constants. The Lifting package works with all four kinds of type abstraction: type copies, subtypes, total quotients and partial quotients. - Theoretical background can be found in @{cite - "Huffman-Kuncar:2013:lifting_transfer"}. + Theoretical background can be found in \<^cite>\"Huffman-Kuncar:2013:lifting_transfer"\. \begin{matharray}{rcl} @{command_def (HOL) "setup_lifting"} & : & \local_theory \ local_theory\\\ @{command_def (HOL) "lift_definition"} & : & \local_theory \ proof(prove)\\\ @{command_def (HOL) "lifting_forget"} & : & \local_theory \ local_theory\\\ @{command_def (HOL) "lifting_update"} & : & \local_theory \ local_theory\\\ @{command_def (HOL) "print_quot_maps"} & : & \context \\\\ @{command_def (HOL) "print_quotients"} & : & \context \\\\ @{attribute_def (HOL) "quot_map"} & : & \attribute\ \\ @{attribute_def (HOL) "relator_eq_onp"} & : & \attribute\ \\ @{attribute_def (HOL) "relator_mono"} & : & \attribute\ \\ @{attribute_def (HOL) "relator_distr"} & : & \attribute\ \\ @{attribute_def (HOL) "quot_del"} & : & \attribute\ \\ @{attribute_def (HOL) "lifting_restore"} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{command (HOL) setup_lifting} @{syntax thm} @{syntax thm}? \ (@'parametric' @{syntax thm})? ; @@{command (HOL) lift_definition} ('(' 'code_dt' ')')? \ @{syntax name} '::' @{syntax type} @{syntax mixfix}? 'is' @{syntax term} \ (@'parametric' (@{syntax thm}+))? ; @@{command (HOL) lifting_forget} @{syntax name} ; @@{command (HOL) lifting_update} @{syntax name} ; @@{attribute (HOL) lifting_restore} @{syntax thm} (@{syntax thm} @{syntax thm})? \ \<^descr> @{command (HOL) "setup_lifting"} Sets up the Lifting package to work with a user-defined type. The command supports two modes. \<^enum> The first one is a low-level mode when the user must provide as a first argument of @{command (HOL) "setup_lifting"} a quotient theorem \<^term>\Quotient R Abs Rep T\. The package configures a transfer rule for equality, a domain transfer rules and sets up the @{command_def (HOL) "lift_definition"} command to work with the abstract type. An optional theorem \<^term>\reflp R\, which certifies that the equivalence relation R is total, can be provided as a second argument. This allows the package to generate stronger transfer rules. And finally, the parametricity theorem for \<^term>\R\ can be provided as a third argument. This allows the package to generate a stronger transfer rule for equality. Users generally will not prove the \Quotient\ theorem manually for new types, as special commands exist to automate the process. \<^enum> When a new subtype is defined by @{command (HOL) typedef}, @{command (HOL) "lift_definition"} can be used in its second mode, where only the \<^term>\type_definition\ theorem \<^term>\type_definition Rep Abs A\ is used as an argument of the command. The command internally proves the corresponding \<^term>\Quotient\ theorem and registers it with @{command (HOL) setup_lifting} using its first mode. For quotients, the command @{command (HOL) quotient_type} can be used. The command defines a new quotient type and similarly to the previous case, the corresponding Quotient theorem is proved and registered by @{command (HOL) setup_lifting}. \<^medskip> The command @{command (HOL) "setup_lifting"} also sets up the code generator for the new type. Later on, when a new constant is defined by @{command (HOL) "lift_definition"}, the Lifting package proves and registers a code equation (if there is one) for the new constant. \<^descr> @{command (HOL) "lift_definition"} \f :: \\ @{keyword (HOL) "is"} \t\ Defines a new function \f\ with an abstract type \\\ in terms of a corresponding operation \t\ on a representation type. More formally, if \t :: \\, then the command builds a term \F\ as a corresponding combination of abstraction and representation functions such that \F :: \ \ \\ and defines \f \ F t\. The term \t\ does not have to be necessarily a constant but it can be any term. The command opens a proof and the user must discharge a respectfulness proof obligation. For a type copy, i.e.\ a typedef with \UNIV\, the obligation is discharged automatically. The proof goal is presented in a user-friendly, readable form. A respectfulness theorem in the standard format \f.rsp\ and a transfer rule \f.transfer\ for the Transfer package are generated by the package. The user can specify a parametricity theorems for \t\ after the keyword @{keyword "parametric"}, which allows the command to generate parametric transfer rules for \f\. For each constant defined through trivial quotients (type copies or subtypes) \f.rep_eq\ is generated. The equation is a code certificate that defines \f\ using the representation function. For each constant \f.abs_eq\ is generated. The equation is unconditional for total quotients. The equation defines \f\ using the abstraction function. \<^medskip> Integration with [@{attribute code} abstract]: For subtypes (e.g.\ corresponding to a datatype invariant, such as \<^typ>\'a dlist\), @{command (HOL) "lift_definition"} uses a code certificate theorem \f.rep_eq\ as a code equation. Because of the limitation of the code generator, \f.rep_eq\ cannot be used as a code equation if the subtype occurs inside the result type rather than at the top level (e.g.\ function returning \<^typ>\'a dlist option\ vs. \<^typ>\'a dlist\). In this case, an extension of @{command (HOL) "lift_definition"} can be invoked by specifying the flag \code_dt\. This extension enables code execution through series of internal type and lifting definitions if the return type \\\ meets the following inductive conditions: \<^descr> \\\ is a type variable \<^descr> \\ = \\<^sub>1 \ \\<^sub>n \\, where \\\ is an abstract type constructor and \\\<^sub>1 \ \\<^sub>n\ do not contain abstract types (i.e.\ \<^typ>\int dlist\ is allowed whereas \<^typ>\int dlist dlist\ not) \<^descr> \\ = \\<^sub>1 \ \\<^sub>n \\, \\\ is a type constructor that was defined as a (co)datatype whose constructor argument types do not contain either non-free datatypes or the function type. Integration with [@{attribute code} equation]: For total quotients, @{command (HOL) "lift_definition"} uses \f.abs_eq\ as a code equation. \<^descr> @{command (HOL) lifting_forget} and @{command (HOL) lifting_update} These two commands serve for storing and deleting the set-up of the Lifting package and corresponding transfer rules defined by this package. This is useful for hiding of type construction details of an abstract type when the construction is finished but it still allows additions to this construction when this is later necessary. Whenever the Lifting package is set up with a new abstract type \\\ by @{command_def (HOL) "lift_definition"}, the package defines a new bundle that is called \\.lifting\. This bundle already includes set-up for the Lifting package. The new transfer rules introduced by @{command (HOL) "lift_definition"} can be stored in the bundle by the command @{command (HOL) "lifting_update"} \\.lifting\. The command @{command (HOL) "lifting_forget"} \\.lifting\ deletes set-up of the Lifting package for \\\ and deletes all the transfer rules that were introduced by @{command (HOL) "lift_definition"} using \\\ as an abstract type. The stored set-up in a bundle can be reintroduced by the Isar commands for including a bundle (@{command "include"}, @{keyword "includes"} and @{command "including"}). \<^descr> @{command (HOL) "print_quot_maps"} prints stored quotient map theorems. \<^descr> @{command (HOL) "print_quotients"} prints stored quotient theorems. \<^descr> @{attribute (HOL) quot_map} registers a quotient map theorem, a theorem showing how to ``lift'' quotients over type constructors. E.g.\ \<^term>\Quotient R Abs Rep T \ Quotient (rel_set R) (image Abs) (image Rep) (rel_set T)\. For examples see \<^file>\~~/src/HOL/Lifting_Set.thy\ or \<^file>\~~/src/HOL/Lifting.thy\. This property is proved automatically if the involved type is BNF without dead variables. \<^descr> @{attribute (HOL) relator_eq_onp} registers a theorem that shows that a relator applied to an equality restricted by a predicate \<^term>\P\ (i.e.\ \<^term>\eq_onp P\) is equal to a predicator applied to the \<^term>\P\. The combinator \<^const>\eq_onp\ is used for internal encoding of proper subtypes. Such theorems allows the package to hide \eq_onp\ from a user in a user-readable form of a respectfulness theorem. For examples see \<^file>\~~/src/HOL/Lifting_Set.thy\ or \<^file>\~~/src/HOL/Lifting.thy\. This property is proved automatically if the involved type is BNF without dead variables. \<^descr> @{attribute (HOL) "relator_mono"} registers a property describing a monotonicity of a relator. E.g.\ \<^prop>\A \ B \ rel_set A \ rel_set B\. This property is needed for proving a stronger transfer rule in @{command_def (HOL) "lift_definition"} when a parametricity theorem for the raw term is specified and also for the reflexivity prover. For examples see \<^file>\~~/src/HOL/Lifting_Set.thy\ or \<^file>\~~/src/HOL/Lifting.thy\. This property is proved automatically if the involved type is BNF without dead variables. \<^descr> @{attribute (HOL) "relator_distr"} registers a property describing a distributivity of the relation composition and a relator. E.g.\ \rel_set R \\ rel_set S = rel_set (R \\ S)\. This property is needed for proving a stronger transfer rule in @{command_def (HOL) "lift_definition"} when a parametricity theorem for the raw term is specified. When this equality does not hold unconditionally (e.g.\ for the function type), the user can specified each direction separately and also register multiple theorems with different set of assumptions. This attribute can be used only after the monotonicity property was already registered by @{attribute (HOL) "relator_mono"}. For examples see \<^file>\~~/src/HOL/Lifting_Set.thy\ or \<^file>\~~/src/HOL/Lifting.thy\. This property is proved automatically if the involved type is BNF without dead variables. \<^descr> @{attribute (HOL) quot_del} deletes a corresponding Quotient theorem from the Lifting infrastructure and thus de-register the corresponding quotient. This effectively causes that @{command (HOL) lift_definition} will not do any lifting for the corresponding type. This attribute is rather used for low-level manipulation with set-up of the Lifting package because @{command (HOL) lifting_forget} is preferred for normal usage. \<^descr> @{attribute (HOL) lifting_restore} \Quotient_thm pcr_def pcr_cr_eq_thm\ registers the Quotient theorem \Quotient_thm\ in the Lifting infrastructure and thus sets up lifting for an abstract type \\\ (that is defined by \Quotient_thm\). Optional theorems \pcr_def\ and \pcr_cr_eq_thm\ can be specified to register the parametrized correspondence relation for \\\. E.g.\ for \<^typ>\'a dlist\, \pcr_def\ is \pcr_dlist A \ list_all2 A \\ cr_dlist\ and \pcr_cr_eq_thm\ is \pcr_dlist (=) = (=)\. This attribute is rather used for low-level manipulation with set-up of the Lifting package because using of the bundle \\.lifting\ together with the commands @{command (HOL) lifting_forget} and @{command (HOL) lifting_update} is preferred for normal usage. - \<^descr> Integration with the BNF package @{cite "isabelle-datatypes"}: As already + \<^descr> Integration with the BNF package \<^cite>\"isabelle-datatypes"\: As already mentioned, the theorems that are registered by the following attributes are proved and registered automatically if the involved type is BNF without dead variables: @{attribute (HOL) quot_map}, @{attribute (HOL) relator_eq_onp}, @{attribute (HOL) "relator_mono"}, @{attribute (HOL) "relator_distr"}. Also the definition of a relator and predicator is provided automatically. Moreover, if the BNF represents a datatype, simplification rules for a predicator are again proved automatically. \ subsection \Transfer package \label{sec:transfer}\ text \ \begin{matharray}{rcl} @{method_def (HOL) "transfer"} & : & \method\ \\ @{method_def (HOL) "transfer'"} & : & \method\ \\ @{method_def (HOL) "transfer_prover"} & : & \method\ \\ @{attribute_def (HOL) "Transfer.transferred"} & : & \attribute\ \\ @{attribute_def (HOL) "untransferred"} & : & \attribute\ \\ @{method_def (HOL) "transfer_start"} & : & \method\ \\ @{method_def (HOL) "transfer_prover_start"} & : & \method\ \\ @{method_def (HOL) "transfer_step"} & : & \method\ \\ @{method_def (HOL) "transfer_end"} & : & \method\ \\ @{method_def (HOL) "transfer_prover_end"} & : & \method\ \\ @{attribute_def (HOL) "transfer_rule"} & : & \attribute\ \\ @{attribute_def (HOL) "transfer_domain_rule"} & : & \attribute\ \\ @{attribute_def (HOL) "relator_eq"} & : & \attribute\ \\ @{attribute_def (HOL) "relator_domain"} & : & \attribute\ \\ \end{matharray} \<^descr> @{method (HOL) "transfer"} method replaces the current subgoal with a logically equivalent one that uses different types and constants. The replacement of types and constants is guided by the database of transfer rules. Goals are generalized over all free variables by default; this is necessary for variables whose types change, but can be overridden for specific variables with e.g. \transfer fixing: x y z\. \<^descr> @{method (HOL) "transfer'"} is a variant of @{method (HOL) transfer} that allows replacing a subgoal with one that is logically stronger (rather than equivalent). For example, a subgoal involving equality on a quotient type could be replaced with a subgoal involving equality (instead of the corresponding equivalence relation) on the underlying raw type. \<^descr> @{method (HOL) "transfer_prover"} method assists with proving a transfer rule for a new constant, provided the constant is defined in terms of other constants that already have transfer rules. It should be applied after unfolding the constant definitions. \<^descr> @{method (HOL) "transfer_start"}, @{method (HOL) "transfer_step"}, @{method (HOL) "transfer_end"}, @{method (HOL) "transfer_prover_start"} and @{method (HOL) "transfer_prover_end"} methods are meant to be used for debugging of @{method (HOL) "transfer"} and @{method (HOL) "transfer_prover"}, which we can decompose as follows: @{method (HOL) "transfer"} = (@{method (HOL) "transfer_start"}, @{method (HOL) "transfer_step"}+, @{method (HOL) "transfer_end"}) and @{method (HOL) "transfer_prover"} = (@{method (HOL) "transfer_prover_start"}, @{method (HOL) "transfer_step"}+, @{method (HOL) "transfer_prover_end"}). For usage examples see \<^file>\~~/src/HOL/ex/Transfer_Debug.thy\. \<^descr> @{attribute (HOL) "untransferred"} proves the same equivalent theorem as @{method (HOL) "transfer"} internally does. \<^descr> @{attribute (HOL) Transfer.transferred} works in the opposite direction than @{method (HOL) "transfer'"}. E.g.\ given the transfer relation \ZN x n \ (x = int n)\, corresponding transfer rules and the theorem \\x::int \ {0..}. x < x + 1\, the attribute would prove \\n::nat. n < n + 1\. The attribute is still in experimental phase of development. \<^descr> @{attribute (HOL) "transfer_rule"} attribute maintains a collection of transfer rules, which relate constants at two different types. Typical transfer rules may relate different type instances of the same polymorphic constant, or they may relate an operation on a raw type to a corresponding operation on an abstract type (quotient or subtype). For example: \((A ===> B) ===> list_all2 A ===> list_all2 B) map map\ \\ \(cr_int ===> cr_int ===> cr_int) (\(x,y) (u,v). (x+u, y+v)) plus\ Lemmas involving predicates on relations can also be registered using the same attribute. For example: \bi_unique A \ (list_all2 A ===> (=)) distinct distinct\ \\ \\bi_unique A; bi_unique B\ \ bi_unique (rel_prod A B)\ Preservation of predicates on relations (\bi_unique, bi_total, right_unique, right_total, left_unique, left_total\) with the respect to a relator is - proved automatically if the involved type is BNF @{cite - "isabelle-datatypes"} without dead variables. + proved automatically if the involved type is BNF \<^cite>\"isabelle-datatypes"\ without dead variables. \<^descr> @{attribute (HOL) "transfer_domain_rule"} attribute maintains a collection of rules, which specify a domain of a transfer relation by a predicate. E.g.\ given the transfer relation \ZN x n \ (x = int n)\, one can register the following transfer domain rule: \Domainp ZN = (\x. x \ 0)\. The rules allow the package to produce more readable transferred goals, e.g.\ when quantifiers are transferred. \<^descr> @{attribute (HOL) relator_eq} attribute collects identity laws for relators of various type constructors, e.g. \<^term>\rel_set (=) = (=)\. The @{method (HOL) transfer} method uses these lemmas to infer transfer rules for non-polymorphic constants on the fly. For examples see \<^file>\~~/src/HOL/Lifting_Set.thy\ or \<^file>\~~/src/HOL/Lifting.thy\. This property is proved automatically if the involved type is BNF without dead variables. \<^descr> @{attribute_def (HOL) "relator_domain"} attribute collects rules describing domains of relators by predicators. E.g.\ \<^term>\Domainp (rel_set T) = (\A. Ball A (Domainp T))\. This allows the package to lift transfer domain rules through type constructors. For examples see \<^file>\~~/src/HOL/Lifting_Set.thy\ or \<^file>\~~/src/HOL/Lifting.thy\. This property is proved automatically if the involved type is BNF without dead variables. - Theoretical background can be found in @{cite - "Huffman-Kuncar:2013:lifting_transfer"}. + Theoretical background can be found in \<^cite>\"Huffman-Kuncar:2013:lifting_transfer"\. \ subsection \Old-style definitions for quotient types \label{sec:old-quotient}\ text \ \begin{matharray}{rcl} @{command_def (HOL) "quotient_definition"} & : & \local_theory \ proof(prove)\\\ @{command_def (HOL) "print_quotmapsQ3"} & : & \context \\\\ @{command_def (HOL) "print_quotientsQ3"} & : & \context \\\\ @{command_def (HOL) "print_quotconsts"} & : & \context \\\\ @{method_def (HOL) "lifting"} & : & \method\ \\ @{method_def (HOL) "lifting_setup"} & : & \method\ \\ @{method_def (HOL) "descending"} & : & \method\ \\ @{method_def (HOL) "descending_setup"} & : & \method\ \\ @{method_def (HOL) "partiality_descending"} & : & \method\ \\ @{method_def (HOL) "partiality_descending_setup"} & : & \method\ \\ @{method_def (HOL) "regularize"} & : & \method\ \\ @{method_def (HOL) "injection"} & : & \method\ \\ @{method_def (HOL) "cleaning"} & : & \method\ \\ @{attribute_def (HOL) "quot_thm"} & : & \attribute\ \\ @{attribute_def (HOL) "quot_lifted"} & : & \attribute\ \\ @{attribute_def (HOL) "quot_respect"} & : & \attribute\ \\ @{attribute_def (HOL) "quot_preserve"} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{command (HOL) quotient_definition} constdecl? @{syntax thmdecl}? \ @{syntax term} 'is' @{syntax term} ; constdecl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? ; @@{method (HOL) lifting} @{syntax thms}? ; @@{method (HOL) lifting_setup} @{syntax thms}? \ \<^descr> @{command (HOL) "quotient_definition"} defines a constant on the quotient type. \<^descr> @{command (HOL) "print_quotmapsQ3"} prints quotient map functions. \<^descr> @{command (HOL) "print_quotientsQ3"} prints quotients. \<^descr> @{command (HOL) "print_quotconsts"} prints quotient constants. \<^descr> @{method (HOL) "lifting"} and @{method (HOL) "lifting_setup"} methods match the current goal with the given raw theorem to be lifted producing three new subgoals: regularization, injection and cleaning subgoals. @{method (HOL) "lifting"} tries to apply the heuristics for automatically solving these three subgoals and leaves only the subgoals unsolved by the heuristics to the user as opposed to @{method (HOL) "lifting_setup"} which leaves the three subgoals unsolved. \<^descr> @{method (HOL) "descending"} and @{method (HOL) "descending_setup"} try to guess a raw statement that would lift to the current subgoal. Such statement is assumed as a new subgoal and @{method (HOL) "descending"} continues in the same way as @{method (HOL) "lifting"} does. @{method (HOL) "descending"} tries to solve the arising regularization, injection and cleaning subgoals with the analogous method @{method (HOL) "descending_setup"} which leaves the four unsolved subgoals. \<^descr> @{method (HOL) "partiality_descending"} finds the regularized theorem that would lift to the current subgoal, lifts it and leaves as a subgoal. This method can be used with partial equivalence quotients where the non regularized statements would not be true. @{method (HOL) "partiality_descending_setup"} leaves the injection and cleaning subgoals unchanged. \<^descr> @{method (HOL) "regularize"} applies the regularization heuristics to the current subgoal. \<^descr> @{method (HOL) "injection"} applies the injection heuristics to the current goal using the stored quotient respectfulness theorems. \<^descr> @{method (HOL) "cleaning"} applies the injection cleaning heuristics to the current subgoal using the stored quotient preservation theorems. \<^descr> @{attribute (HOL) quot_lifted} attribute tries to automatically transport the theorem to the quotient type. The attribute uses all the defined quotients types and quotient constants often producing undesired results or theorems that cannot be lifted. \<^descr> @{attribute (HOL) quot_respect} and @{attribute (HOL) quot_preserve} attributes declare a theorem as a respectfulness and preservation theorem respectively. These are stored in the local theory store and used by the @{method (HOL) "injection"} and @{method (HOL) "cleaning"} methods respectively. \<^descr> @{attribute (HOL) quot_thm} declares that a certain theorem is a quotient extension theorem. Quotient extension theorems allow for quotienting inside container types. Given a polymorphic type that serves as a container, a map function defined for this container using @{command (HOL) "functor"} and a relation map defined for for the container type, the quotient extension theorem should be \<^term>\Quotient3 R Abs Rep \ Quotient3 (rel_map R) (map Abs) (map Rep)\. Quotient extension theorems are stored in a database and are used all the steps of lifting theorems. \ chapter \Proof tools\ section \Proving propositions\ text \ In addition to the standard proof methods, a number of diagnosis tools search for proofs and provide an Isar proof snippet on success. These tools are available via the following commands. \begin{matharray}{rcl} @{command_def (HOL) "solve_direct"}\\<^sup>*\ & : & \proof \\ \\ @{command_def (HOL) "try"}\\<^sup>*\ & : & \proof \\ \\ @{command_def (HOL) "try0"}\\<^sup>*\ & : & \proof \\ \\ @{command_def (HOL) "sledgehammer"}\\<^sup>*\ & : & \proof \\ \\ @{command_def (HOL) "sledgehammer_params"} & : & \theory \ theory\ \end{matharray} \<^rail>\ @@{command (HOL) try} ; @@{command (HOL) try0} ( ( ( 'simp' | 'intro' | 'elim' | 'dest' ) ':' @{syntax thms} ) + ) ? @{syntax nat}? ; @@{command (HOL) sledgehammer} ( '[' args ']' )? facts? @{syntax nat}? ; @@{command (HOL) sledgehammer_params} ( ( '[' args ']' ) ? ) ; args: ( @{syntax name} '=' value + ',' ) ; facts: '(' ( ( ( ( 'add' | 'del' ) ':' ) ? @{syntax thms} ) + ) ? ')' \ % FIXME check args "value" \<^descr> @{command (HOL) "solve_direct"} checks whether the current subgoals can be solved directly by an existing theorem. Duplicate lemmas can be detected in this way. \<^descr> @{command (HOL) "try0"} attempts to prove a subgoal using a combination of standard proof methods (@{method auto}, @{method simp}, @{method blast}, etc.). Additional facts supplied via \simp:\, \intro:\, \elim:\, and \dest:\ are passed to the appropriate proof methods. \<^descr> @{command (HOL) "try"} attempts to prove or disprove a subgoal using a combination of provers and disprovers (@{command (HOL) "solve_direct"}, @{command (HOL) "quickcheck"}, @{command (HOL) "try0"}, @{command (HOL) "sledgehammer"}, @{command (HOL) "nitpick"}). \<^descr> @{command (HOL) "sledgehammer"} attempts to prove a subgoal using external automatic provers (resolution provers and SMT solvers). See the Sledgehammer - manual @{cite "isabelle-sledgehammer"} for details. + manual \<^cite>\"isabelle-sledgehammer"\ for details. \<^descr> @{command (HOL) "sledgehammer_params"} changes @{command (HOL) "sledgehammer"} configuration options persistently. \ section \Checking and refuting propositions\ text \ Identifying incorrect propositions usually involves evaluation of particular assignments and systematic counterexample search. This is supported by the following commands. \begin{matharray}{rcl} @{command_def (HOL) "value"}\\<^sup>*\ & : & \context \\ \\ @{command_def (HOL) "values"}\\<^sup>*\ & : & \context \\ \\ @{command_def (HOL) "quickcheck"}\\<^sup>*\ & : & \proof \\ \\ @{command_def (HOL) "nitpick"}\\<^sup>*\ & : & \proof \\ \\ @{command_def (HOL) "quickcheck_params"} & : & \theory \ theory\ \\ @{command_def (HOL) "nitpick_params"} & : & \theory \ theory\ \\ @{command_def (HOL) "quickcheck_generator"} & : & \theory \ theory\ \\ @{command_def (HOL) "find_unused_assms"} & : & \context \\ \end{matharray} \<^rail>\ @@{command (HOL) value} ( '[' @{syntax name} ']' )? modes? @{syntax term} ; @@{command (HOL) values} modes? @{syntax nat}? @{syntax term} ; (@@{command (HOL) quickcheck} | @@{command (HOL) nitpick}) ( '[' args ']' )? @{syntax nat}? ; (@@{command (HOL) quickcheck_params} | @@{command (HOL) nitpick_params}) ( '[' args ']' )? ; @@{command (HOL) quickcheck_generator} @{syntax name} \ 'operations:' ( @{syntax term} +) ; @@{command (HOL) find_unused_assms} @{syntax name}? ; modes: '(' (@{syntax name} +) ')' ; args: ( @{syntax name} '=' value + ',' ) \ % FIXME check "value" \<^descr> @{command (HOL) "value"}~\t\ evaluates and prints a term; optionally \modes\ can be specified, which are appended to the current print mode; see \secref{sec:print-modes}. Evaluation is tried first using ML, falling back to normalization by evaluation if this fails. Alternatively a specific evaluator can be selected using square brackets; typical evaluators use the current set of code equations to normalize and include \simp\ for fully symbolic evaluation using the simplifier, \nbe\ for \<^emph>\normalization by evaluation\ and \<^emph>\code\ for code generation in SML. \<^descr> @{command (HOL) "values"}~\t\ enumerates a set comprehension by evaluation and prints its values up to the given number of solutions; optionally \modes\ can be specified, which are appended to the current print mode; see \secref{sec:print-modes}. \<^descr> @{command (HOL) "quickcheck"} tests the current goal for counterexamples using a series of assignments for its free variables; by default the first subgoal is tested, an other can be selected explicitly using an optional goal index. Assignments can be chosen exhausting the search space up to a given size, or using a fixed number of random assignments in the search space, or exploring the search space symbolically using narrowing. By default, quickcheck uses exhaustive testing. A number of configuration options are supported for @{command (HOL) "quickcheck"}, notably: \<^descr>[\tester\] specifies which testing approach to apply. There are three testers, \exhaustive\, \random\, and \narrowing\. An unknown configuration option is treated as an argument to tester, making \tester =\ optional. When multiple testers are given, these are applied in parallel. If no tester is specified, quickcheck uses the testers that are set active, i.e.\ configurations @{attribute quickcheck_exhaustive_active}, @{attribute quickcheck_random_active}, @{attribute quickcheck_narrowing_active} are set to true. \<^descr>[\size\] specifies the maximum size of the search space for assignment values. \<^descr>[\genuine_only\] sets quickcheck only to return genuine counterexample, but not potentially spurious counterexamples due to underspecified functions. \<^descr>[\abort_potential\] sets quickcheck to abort once it found a potentially spurious counterexample and to not continue to search for a further genuine counterexample. For this option to be effective, the \genuine_only\ option must be set to false. \<^descr>[\eval\] takes a term or a list of terms and evaluates these terms under the variable assignment found by quickcheck. This option is currently only supported by the default (exhaustive) tester. \<^descr>[\iterations\] sets how many sets of assignments are generated for each particular size. \<^descr>[\no_assms\] specifies whether assumptions in structured proofs should be ignored. \<^descr>[\locale\] specifies how to process conjectures in a locale context, i.e.\ they can be interpreted or expanded. The option is a whitespace-separated list of the two words \interpret\ and \expand\. The list determines the order they are employed. The default setting is to first use interpretations and then test the expanded conjecture. The option is only provided as attribute declaration, but not as parameter to the command. \<^descr>[\timeout\] sets the time limit in seconds. \<^descr>[\default_type\] sets the type(s) generally used to instantiate type variables. \<^descr>[\report\] if set quickcheck reports how many tests fulfilled the preconditions. \<^descr>[\use_subtype\] if set quickcheck automatically lifts conjectures to registered subtypes if possible, and tests the lifted conjecture. \<^descr>[\quiet\] if set quickcheck does not output anything while testing. \<^descr>[\verbose\] if set quickcheck informs about the current size and cardinality while testing. \<^descr>[\expect\] can be used to check if the user's expectation was met (\no_expectation\, \no_counterexample\, or \counterexample\). These option can be given within square brackets. Using the following type classes, the testers generate values and convert them back into Isabelle terms for displaying counterexamples. \<^descr>[\exhaustive\] The parameters of the type classes \<^class>\exhaustive\ and \<^class>\full_exhaustive\ implement the testing. They take a testing function as a parameter, which takes a value of type \<^typ>\'a\ and optionally produces a counterexample, and a size parameter for the test values. In \<^class>\full_exhaustive\, the testing function parameter additionally expects a lazy term reconstruction in the type \<^typ>\Code_Evaluation.term\ of the tested value. The canonical implementation for \exhaustive\ testers calls the given testing function on all values up to the given size and stops as soon as a counterexample is found. \<^descr>[\random\] The operation \<^const>\Quickcheck_Random.random\ of the type class \<^class>\random\ generates a pseudo-random value of the given size and a lazy term reconstruction of the value in the type \<^typ>\Code_Evaluation.term\. A pseudo-randomness generator is defined in theory \<^theory>\HOL.Random\. - \<^descr>[\narrowing\] implements Haskell's Lazy Smallcheck @{cite - "runciman-naylor-lindblad"} using the type classes \<^class>\narrowing\ and + \<^descr>[\narrowing\] implements Haskell's Lazy Smallcheck \<^cite>\"runciman-naylor-lindblad"\ using the type classes \<^class>\narrowing\ and \<^class>\partial_term_of\. Variables in the current goal are initially represented as symbolic variables. If the execution of the goal tries to evaluate one of them, the test engine replaces it with refinements provided by \<^const>\narrowing\. Narrowing views every value as a sum-of-products which is expressed using the operations \<^const>\Quickcheck_Narrowing.cons\ (embedding a value), \<^const>\Quickcheck_Narrowing.apply\ (product) and \<^const>\Quickcheck_Narrowing.sum\ (sum). The refinement should enable further evaluation of the goal. For example, \<^const>\narrowing\ for the list type \<^typ>\'a :: narrowing list\ can be recursively defined as \<^term>\Quickcheck_Narrowing.sum (Quickcheck_Narrowing.cons []) (Quickcheck_Narrowing.apply (Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons (#)) narrowing) narrowing)\. If a symbolic variable of type \<^typ>\_ list\ is evaluated, it is replaced by (i)~the empty list \<^term>\[]\ and (ii)~by a non-empty list whose head and tail can then be recursively refined if needed. To reconstruct counterexamples, the operation \<^const>\partial_term_of\ transforms \narrowing\'s deep representation of terms to the type \<^typ>\Code_Evaluation.term\. The deep representation models symbolic variables as \<^const>\Quickcheck_Narrowing.Narrowing_variable\, which are normally converted to \<^const>\Code_Evaluation.Free\, and refined values as \<^term>\Quickcheck_Narrowing.Narrowing_constructor i args\, where \<^term>\i :: integer\ denotes the index in the sum of refinements. In the above example for lists, \<^term>\0\ corresponds to \<^term>\[]\ and \<^term>\1\ to \<^term>\(#)\. The command @{command (HOL) "code_datatype"} sets up \<^const>\partial_term_of\ such that the \<^term>\i\-th refinement is interpreted as the \<^term>\i\-th constructor, but it does not ensures consistency with \<^const>\narrowing\. \<^descr> @{command (HOL) "quickcheck_params"} changes @{command (HOL) "quickcheck"} configuration options persistently. \<^descr> @{command (HOL) "quickcheck_generator"} creates random and exhaustive value generators for a given type and operations. It generates values by using the operations as if they were constructors of that type. \<^descr> @{command (HOL) "nitpick"} tests the current goal for counterexamples using a reduction to first-order relational logic. See the Nitpick manual - @{cite "isabelle-nitpick"} for details. + \<^cite>\"isabelle-nitpick"\ for details. \<^descr> @{command (HOL) "nitpick_params"} changes @{command (HOL) "nitpick"} configuration options persistently. \<^descr> @{command (HOL) "find_unused_assms"} finds potentially superfluous assumptions in theorems using quickcheck. It takes the theory name to be checked for superfluous assumptions as optional argument. If not provided, it checks the current theory. Options to the internal quickcheck invocations can be changed with common configuration declarations. \ section \Coercive subtyping\ text \ \begin{matharray}{rcl} @{attribute_def (HOL) coercion} & : & \attribute\ \\ @{attribute_def (HOL) coercion_delete} & : & \attribute\ \\ @{attribute_def (HOL) coercion_enabled} & : & \attribute\ \\ @{attribute_def (HOL) coercion_map} & : & \attribute\ \\ @{attribute_def (HOL) coercion_args} & : & \attribute\ \\ \end{matharray} Coercive subtyping allows the user to omit explicit type conversions, also called \<^emph>\coercions\. Type inference will add them as necessary when parsing - a term. See @{cite "traytel-berghofer-nipkow-2011"} for details. + a term. See \<^cite>\"traytel-berghofer-nipkow-2011"\ for details. \<^rail>\ @@{attribute (HOL) coercion} (@{syntax term}) ; @@{attribute (HOL) coercion_delete} (@{syntax term}) ; @@{attribute (HOL) coercion_map} (@{syntax term}) ; @@{attribute (HOL) coercion_args} (@{syntax const}) (('+' | '0' | '-')+) \ \<^descr> @{attribute (HOL) "coercion"}~\f\ registers a new coercion function \f :: \\<^sub>1 \ \\<^sub>2\ where \\\<^sub>1\ and \\\<^sub>2\ are type constructors without arguments. Coercions are composed by the inference algorithm if needed. Note that the type inference algorithm is complete only if the registered coercions form a lattice. \<^descr> @{attribute (HOL) "coercion_delete"}~\f\ deletes a preceding declaration (using @{attribute (HOL) "coercion"}) of the function \f :: \\<^sub>1 \ \\<^sub>2\ as a coercion. \<^descr> @{attribute (HOL) "coercion_map"}~\map\ registers a new map function to lift coercions through type constructors. The function \map\ must conform to the following type pattern \begin{matharray}{lll} \map\ & \::\ & \f\<^sub>1 \ \ \ f\<^sub>n \ (\\<^sub>1, \, \\<^sub>n) t \ (\\<^sub>1, \, \\<^sub>n) t\ \\ \end{matharray} where \t\ is a type constructor and \f\<^sub>i\ is of type \\\<^sub>i \ \\<^sub>i\ or \\\<^sub>i \ \\<^sub>i\. Registering a map function overwrites any existing map function for this particular type constructor. \<^descr> @{attribute (HOL) "coercion_args"} can be used to disallow coercions to be inserted in certain positions in a term. For example, given the constant \c :: \\<^sub>1 \ \\<^sub>2 \ \\<^sub>3 \ \\<^sub>4\ and the list of policies \- + 0\ as arguments, coercions will not be inserted in the first argument of \c\ (policy \-\); they may be inserted in the second argument (policy \+\) even if the constant \c\ itself is in a position where coercions are disallowed; the third argument inherits the allowance of coercsion insertion from the position of the constant \c\ (policy \0\). The standard usage of policies is the definition of syntatic constructs (usually extralogical, i.e., processed and stripped during type inference), that should not be destroyed by the insertion of coercions (see, for example, the setup for the case syntax in \<^theory>\HOL.Ctr_Sugar\). \<^descr> @{attribute (HOL) "coercion_enabled"} enables the coercion inference algorithm. \ section \Arithmetic proof support\ text \ \begin{matharray}{rcl} @{method_def (HOL) arith} & : & \method\ \\ @{attribute_def (HOL) arith} & : & \attribute\ \\ @{attribute_def (HOL) linarith_split} & : & \attribute\ \\ \end{matharray} \<^descr> @{method (HOL) arith} decides linear arithmetic problems (on types \nat\, \int\, \real\). Any current facts are inserted into the goal before running the procedure. \<^descr> @{attribute (HOL) arith} declares facts that are supplied to the arithmetic provers implicitly. \<^descr> @{attribute (HOL) linarith_split} attribute declares case split rules to be expanded before @{method (HOL) arith} is invoked. Note that a simpler (but faster) arithmetic prover is already invoked by the Simplifier. \ section \Intuitionistic proof search\ text \ \begin{matharray}{rcl} @{method_def (HOL) iprover} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method (HOL) iprover} (@{syntax rulemod} *) \ \<^descr> @{method (HOL) iprover} performs intuitionistic proof search, depending on specifically declared rules from the context, or given as explicit arguments. Chained facts are inserted into the goal before commencing proof search. Rules need to be classified as @{attribute (Pure) intro}, @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the ``\!\'' indicator refers to ``safe'' rules, which may be applied aggressively (without considering back-tracking later). Rules declared with ``\?\'' are ignored in proof search (the single-step @{method (Pure) rule} method still observes these). An explicit weight annotation may be given as well; otherwise the number of rule premises will be taken into account here. \ section \Model Elimination and Resolution\ text \ \begin{matharray}{rcl} @{method_def (HOL) "meson"} & : & \method\ \\ @{method_def (HOL) "metis"} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method (HOL) meson} @{syntax thms}? ; @@{method (HOL) metis} ('(' ('partial_types' | 'full_types' | 'no_types' | @{syntax name}) ')')? @{syntax thms}? \ \<^descr> @{method (HOL) meson} implements Loveland's model elimination procedure - @{cite "loveland-78"}. See \<^file>\~~/src/HOL/ex/Meson_Test.thy\ for examples. + \<^cite>\"loveland-78"\. See \<^file>\~~/src/HOL/ex/Meson_Test.thy\ for examples. \<^descr> @{method (HOL) metis} combines ordered resolution and ordered paramodulation to find first-order (or mildly higher-order) proofs. The first optional argument specifies a type encoding; see the Sledgehammer - manual @{cite "isabelle-sledgehammer"} for details. The directory + manual \<^cite>\"isabelle-sledgehammer"\ for details. The directory \<^dir>\~~/src/HOL/Metis_Examples\ contains several small theories developed to a large extent using @{method (HOL) metis}. \ section \Algebraic reasoning via Gröbner bases\ text \ \begin{matharray}{rcl} @{method_def (HOL) "algebra"} & : & \method\ \\ @{attribute_def (HOL) algebra} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{method (HOL) algebra} ('add' ':' @{syntax thms})? ('del' ':' @{syntax thms})? ; @@{attribute (HOL) algebra} (() | 'add' | 'del') \ \<^descr> @{method (HOL) algebra} performs algebraic reasoning via Gröbner bases, - see also @{cite "Chaieb-Wenzel:2007"} and @{cite \\S3.2\ "Chaieb-thesis"}. + see also \<^cite>\"Chaieb-Wenzel:2007"\ and \<^cite>\\\S3.2\ in "Chaieb-thesis"\. The method handles deals with two main classes of problems: \<^enum> Universal problems over multivariate polynomials in a (semi)-ring/field/idom; the capabilities of the method are augmented according to properties of these structures. For this problem class the method is only complete for algebraically closed fields, since the underlying method is based on Hilbert's Nullstellensatz, where the equivalence only holds for algebraically closed fields. The problems can contain equations \p = 0\ or inequations \q \ 0\ anywhere within a universal problem statement. \<^enum> All-exists problems of the following restricted (but useful) form: @{text [display] "\x\<^sub>1 \ x\<^sub>n. e\<^sub>1(x\<^sub>1, \, x\<^sub>n) = 0 \ \ \ e\<^sub>m(x\<^sub>1, \, x\<^sub>n) = 0 \ (\y\<^sub>1 \ y\<^sub>k. p\<^sub>1\<^sub>1(x\<^sub>1, \ ,x\<^sub>n) * y\<^sub>1 + \ + p\<^sub>1\<^sub>k(x\<^sub>1, \, x\<^sub>n) * y\<^sub>k = 0 \ \ \ p\<^sub>t\<^sub>1(x\<^sub>1, \, x\<^sub>n) * y\<^sub>1 + \ + p\<^sub>t\<^sub>k(x\<^sub>1, \, x\<^sub>n) * y\<^sub>k = 0)"} Here \e\<^sub>1, \, e\<^sub>n\ and the \p\<^sub>i\<^sub>j\ are multivariate polynomials only in the variables mentioned as arguments. The proof method is preceded by a simplification step, which may be modified by using the form \(algebra add: ths\<^sub>1 del: ths\<^sub>2)\. This acts like declarations for the Simplifier (\secref{sec:simplifier}) on a private simpset for this tool. \<^descr> @{attribute algebra} (as attribute) manages the default collection of pre-simplification rules of the above proof method. \ subsubsection \Example\ text \ The subsequent example is from geometry: collinearity is invariant by rotation. \ (*<*)experiment begin(*>*) type_synonym point = "int \ int" fun collinear :: "point \ point \ point \ bool" where "collinear (Ax, Ay) (Bx, By) (Cx, Cy) \ (Ax - Bx) * (By - Cy) = (Ay - By) * (Bx - Cx)" lemma collinear_inv_rotation: assumes "collinear (Ax, Ay) (Bx, By) (Cx, Cy)" and "c\<^sup>2 + s\<^sup>2 = 1" shows "collinear (Ax * c - Ay * s, Ay * c + Ax * s) (Bx * c - By * s, By * c + Bx * s) (Cx * c - Cy * s, Cy * c + Cx * s)" using assms by (algebra add: collinear.simps) (*<*)end(*>*) text \ See also \<^file>\~~/src/HOL/Examples/Groebner_Examples.thy\. \ section \Coherent Logic\ text \ \begin{matharray}{rcl} @{method_def (HOL) "coherent"} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method (HOL) coherent} @{syntax thms}? \ - \<^descr> @{method (HOL) coherent} solves problems of \<^emph>\Coherent Logic\ @{cite - "Bezem-Coquand:2005"}, which covers applications in confluence theory, + \<^descr> @{method (HOL) coherent} solves problems of \<^emph>\Coherent Logic\ \<^cite>\"Bezem-Coquand:2005"\, which covers applications in confluence theory, lattice theory and projective geometry. See \<^file>\~~/src/HOL/Examples/Coherent.thy\ for some examples. \ section \Unstructured case analysis and induction \label{sec:hol-induct-tac}\ text \ The following tools of Isabelle/HOL support cases analysis and induction in unstructured tactic scripts; see also \secref{sec:cases-induct} for proper Isar versions of similar ideas. \begin{matharray}{rcl} @{method_def (HOL) case_tac}\\<^sup>*\ & : & \method\ \\ @{method_def (HOL) induct_tac}\\<^sup>*\ & : & \method\ \\ @{method_def (HOL) ind_cases}\\<^sup>*\ & : & \method\ \\ @{command_def (HOL) "inductive_cases"}\\<^sup>*\ & : & \local_theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{method (HOL) case_tac} @{syntax goal_spec}? @{syntax term} rule? ; @@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule? ; @@{method (HOL) ind_cases} (@{syntax prop}+) @{syntax for_fixes} ; @@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and') ; rule: 'rule' ':' @{syntax thm} \ \<^descr> @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit to reason about inductive types. Rules are selected according to the declarations by the @{attribute cases} and @{attribute induct} attributes, cf.\ \secref{sec:cases-induct}. The @{command (HOL) datatype} package already takes care of this. These unstructured tactics feature both goal addressing and dynamic instantiation. Note that named rule cases are \<^emph>\not\ provided as would be by the proper @{method cases} and @{method induct} proof methods (see \secref{sec:cases-induct}). Unlike the @{method induct} method, @{method induct_tac} does not handle structured rule statements, only the compact object-logic conclusion of the subgoal being addressed. \<^descr> @{method (HOL) ind_cases} and @{command (HOL) "inductive_cases"} provide an interface to the internal \<^ML_text>\mk_cases\ operation. Rules are simplified in an unrestricted forward manner. While @{method (HOL) ind_cases} is a proof method to apply the result immediately as elimination rules, @{command (HOL) "inductive_cases"} provides case split theorems at the theory level for later use. The @{keyword "for"} argument of the @{method (HOL) ind_cases} method allows to specify a list of variables that should be generalized before applying the resulting rule. \ section \Adhoc tuples\ text \ \begin{matharray}{rcl} @{attribute_def (HOL) split_format}\\<^sup>*\ & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute (HOL) split_format} ('(' 'complete' ')')? \ \<^descr> @{attribute (HOL) split_format}\ \(complete)\ causes arguments in function applications to be represented canonically according to their tuple type structure. Note that this operation tends to invent funny names for new local parameters introduced. \ chapter \Executable code \label{ch:export-code}\ text \ For validation purposes, it is often useful to \<^emph>\execute\ specifications. In principle, execution could be simulated by Isabelle's inference kernel, i.e. by a combination of resolution and simplification. Unfortunately, this approach is rather inefficient. A more efficient way of executing specifications is to translate them into a functional programming language such as ML. Isabelle provides a generic framework to support code generation from executable specifications. Isabelle/HOL instantiates these mechanisms in a way that is amenable to end-user applications. Code can be generated for functional programs (including overloading using type classes) targeting SML - @{cite SML}, OCaml @{cite OCaml}, Haskell @{cite "haskell-revised-report"} - and Scala @{cite "scala-overview-tech-report"}. Conceptually, code + \<^cite>\SML\, OCaml \<^cite>\OCaml\, Haskell \<^cite>\"haskell-revised-report"\ + and Scala \<^cite>\"scala-overview-tech-report"\. Conceptually, code generation is split up in three steps: \<^emph>\selection\ of code theorems, \<^emph>\translation\ into an abstract executable view and \<^emph>\serialization\ to a specific \<^emph>\target language\. Inductive specifications can be executed using - the predicate compiler which operates within HOL. See @{cite - "isabelle-codegen"} for an introduction. + the predicate compiler which operates within HOL. See \<^cite>\"isabelle-codegen"\ for an introduction. \begin{matharray}{rcl} @{command_def (HOL) "export_code"}\\<^sup>*\ & : & \local_theory \ local_theory\ \\ @{attribute_def (HOL) code} & : & \attribute\ \\ @{command_def (HOL) "code_datatype"} & : & \theory \ theory\ \\ @{command_def (HOL) "print_codesetup"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def (HOL) code_unfold} & : & \attribute\ \\ @{attribute_def (HOL) code_post} & : & \attribute\ \\ @{attribute_def (HOL) code_abbrev} & : & \attribute\ \\ @{command_def (HOL) "print_codeproc"}\\<^sup>*\ & : & \context \\ \\ @{command_def (HOL) "code_thms"}\\<^sup>*\ & : & \context \\ \\ @{command_def (HOL) "code_deps"}\\<^sup>*\ & : & \context \\ \\ @{command_def (HOL) "code_reserved"} & : & \theory \ theory\ \\ @{command_def (HOL) "code_printing"} & : & \theory \ theory\ \\ @{command_def (HOL) "code_identifier"} & : & \theory \ theory\ \\ @{command_def (HOL) "code_monad"} & : & \theory \ theory\ \\ @{command_def (HOL) "code_reflect"} & : & \theory \ theory\ \\ @{command_def (HOL) "code_pred"} & : & \theory \ proof(prove)\ \\ @{attribute_def (HOL) code_timing} & : & \attribute\ \\ @{attribute_def (HOL) code_simp_trace} & : & \attribute\ \\ @{attribute_def (HOL) code_runtime_trace} & : & \attribute\ \end{matharray} \<^rail>\ @@{command (HOL) export_code} @'open'? \ (const_expr+) (export_target*) ; export_target: @'in' target (@'module_name' @{syntax name})? \ (@'file_prefix' @{syntax path})? ('(' args ')')? ; target: 'SML' | 'OCaml' | 'Haskell' | 'Scala' | 'Eval' ; const_expr: (const | 'name._' | '_') ; const: @{syntax term} ; type_constructor: @{syntax name} ; class: @{syntax name} ; path: @{syntax embedded} ; @@{attribute (HOL) code} ('equation' | 'nbe' | 'abstype' | 'abstract' | 'del' | 'drop:' (const+) | 'abort:' (const+))? ; @@{command (HOL) code_datatype} (const+) ; @@{attribute (HOL) code_unfold} 'del'? ; @@{attribute (HOL) code_post} 'del'? ; @@{attribute (HOL) code_abbrev} 'del'? ; @@{command (HOL) code_thms} (const_expr+) ; @@{command (HOL) code_deps} (const_expr+) ; @@{command (HOL) code_reserved} target (@{syntax string}+) ; symbol_const: @'constant' const ; symbol_type_constructor: @'type_constructor' type_constructor ; symbol_class: @'type_class' class ; symbol_class_relation: @'class_relation' class ('<' | '\') class ; symbol_class_instance: @'class_instance' type_constructor @'::' class ; symbol_module: @'code_module' name ; syntax: @{syntax string} | (@'infix' | @'infixl' | @'infixr') @{syntax nat} @{syntax string} ; printing_const: symbol_const ('\' | '=>') \ ('(' target ')' syntax ? + @'and') ; printing_type_constructor: symbol_type_constructor ('\' | '=>') \ ('(' target ')' syntax ? + @'and') ; printing_class: symbol_class ('\' | '=>') \ ('(' target ')' @{syntax string} ? + @'and') ; printing_class_relation: symbol_class_relation ('\' | '=>') \ ('(' target ')' @{syntax string} ? + @'and') ; printing_class_instance: symbol_class_instance ('\'| '=>') \ ('(' target ')' '-' ? + @'and') ; printing_module: symbol_module ('\' | '=>') \ ('(' target ')' (@{syntax string} for_symbol?)? + @'and') ; for_symbol: @'for' ((symbol_const | symbol_typeconstructor | symbol_class | symbol_class_relation | symbol_class_instance)+) ; @@{command (HOL) code_printing} ((printing_const | printing_type_constructor | printing_class | printing_class_relation | printing_class_instance | printing_module) + '|') ; @@{command (HOL) code_identifier} ((symbol_const | symbol_type_constructor | symbol_class | symbol_class_relation | symbol_class_instance | symbol_module ) ('\' | '=>') \ ('(' target ')' @{syntax string} ? + @'and') + '|') ; @@{command (HOL) code_monad} const const target ; @@{command (HOL) code_reflect} @{syntax string} \ (@'datatypes' (@{syntax string} '=' ('_' | (@{syntax string} + '|') + @'and')))? \ (@'functions' (@{syntax string} +))? (@'file_prefix' @{syntax path})? ; @@{command (HOL) code_pred} \ ('(' @'modes' ':' modedecl ')')? \ const ; modedecl: (modes | ((const ':' modes) \ (@'and' ((const ':' modes @'and')+))?)) ; modes: mode @'as' const \ \<^descr> @{command (HOL) "export_code"} generates code for a given list of constants in the specified target language(s). If no serialization instruction is given, only abstract code is generated internally. Constants may be specified by giving them literally, referring to all executable constants within a certain theory by giving \name._\, or referring to \<^emph>\all\ executable constants currently available by giving \_\. By default, exported identifiers are minimized per module. This can be suppressed by prepending @{keyword "open"} before the list of constants. By default, for each involved theory one corresponding name space module is generated. Alternatively, a module name may be specified after the @{keyword "module_name"} keyword; then \<^emph>\all\ code is placed in this module. Generated code is output as logical files within the theory context, as well as session exports that can be retrieved using @{tool_ref export}, or @{tool build} with option \<^verbatim>\-e\ and suitable \isakeyword{export\_files} specifications in the session \<^verbatim>\ROOT\ entry. All files have a common directory prefix: the long theory name plus ``\<^verbatim>\code\''. The actual file name is determined by the target language together with an optional \<^theory_text>\file_prefix\ (the default is ``\<^verbatim>\export\'' with a consecutive number within the current theory). For \SML\, \OCaml\ and \Scala\, the file prefix becomes a plain file with extension (e.g.\ ``\<^verbatim>\.ML\'' for SML). For \Haskell\ the file prefix becomes a directory that is populated with a separate file for each module (with extension ``\<^verbatim>\.hs\''). Serializers take an optional list of arguments in parentheses. \<^item> For \<^emph>\Haskell\ a module name prefix may be given using the ``\root:\'' argument; ``\string_classes\'' adds a ``\<^verbatim>\deriving (Read, Show)\'' clause to each appropriate datatype declaration. \<^item> For \<^emph>\Scala\, ``\case_insensitive\'' avoids name clashes on case-insensitive file systems. \<^descr> @{attribute (HOL) code} declares code equations for code generation. Variant \code equation\ declares a conventional equation as code equation. Variants \code abstype\ and \code abstract\ declare abstract datatype certificates or code equations on abstract datatype representations respectively. Vanilla \code\ falls back to \code equation\ or \code abstract\ depending on the syntactic shape of the underlying equation. Variant \code del\ deselects a code equation for code generation. Variant \code nbe\ accepts also non-left-linear equations for \<^emph>\normalization by evaluation\ only. Variants \code drop:\ and \code abort:\ take a list of constants as arguments and drop all code equations declared for them. In the case of \abort\, these constants if needed are implemented by program abort (exception). Packages declaring code equations usually provide a reasonable default setup. \<^descr> @{command (HOL) "code_datatype"} specifies a constructor set for a logical type. \<^descr> @{command (HOL) "print_codesetup"} gives an overview on selected code equations and code generator datatypes. \<^descr> @{attribute (HOL) code_unfold} declares (or with option ``\del\'' removes) theorems which during preprocessing are applied as rewrite rules to any code equation or evaluation input. \<^descr> @{attribute (HOL) code_post} declares (or with option ``\del\'' removes) theorems which are applied as rewrite rules to any result of an evaluation. \<^descr> @{attribute (HOL) code_abbrev} declares (or with option ``\del\'' removes) equations which are applied as rewrite rules to any result of an evaluation and symmetrically during preprocessing to any code equation or evaluation input. \<^descr> @{command (HOL) "print_codeproc"} prints the setup of the code generator preprocessor. \<^descr> @{command (HOL) "code_thms"} prints a list of theorems representing the corresponding program containing all given constants after preprocessing. \<^descr> @{command (HOL) "code_deps"} visualizes dependencies of theorems representing the corresponding program containing all given constants after preprocessing. \<^descr> @{command (HOL) "code_reserved"} declares a list of names as reserved for a given target, preventing it to be shadowed by any generated code. \<^descr> @{command (HOL) "code_printing"} associates a series of symbols (constants, type constructors, classes, class relations, instances, module names) with target-specific serializations; omitting a serialization deletes an existing serialization. \<^descr> @{command (HOL) "code_monad"} provides an auxiliary mechanism to generate monadic code for Haskell. \<^descr> @{command (HOL) "code_identifier"} associates a a series of symbols (constants, type constructors, classes, class relations, instances, module names) with target-specific hints how these symbols shall be named. These hints gain precedence over names for symbols with no hints at all. Conflicting hints are subject to name disambiguation. \<^emph>\Warning:\ It is at the discretion of the user to ensure that name prefixes of identifiers in compound statements like type classes or datatypes are still the same. \<^descr> @{command (HOL) "code_reflect"} without a ``\<^theory_text>\file_prefix\'' argument compiles code into the system runtime environment and modifies the code generator setup that future invocations of system runtime code generation referring to one of the ``\datatypes\'' or ``\functions\'' entities use these precompiled entities. With a ``\<^theory_text>\file_prefix\'' argument, the corresponding code is generated/exported to the specified file (as for \<^theory_text>\export_code\) without modifying the code generator setup. \<^descr> @{command (HOL) "code_pred"} creates code equations for a predicate given a set of introduction rules. Optional mode annotations determine which arguments are supposed to be input or output. If alternative introduction rules are declared, one must prove a corresponding elimination rule. \<^descr> @{attribute (HOL) "code_timing"} scrapes timing samples from different stages of the code generator. \<^descr> @{attribute (HOL) "code_simp_trace"} traces the simplifier when it is used with code equations. \<^descr> @{attribute (HOL) "code_runtime_trace"} traces ML code generated dynamically for execution. \ end diff --git a/src/Doc/Isar_Ref/Inner_Syntax.thy b/src/Doc/Isar_Ref/Inner_Syntax.thy --- a/src/Doc/Isar_Ref/Inner_Syntax.thy +++ b/src/Doc/Isar_Ref/Inner_Syntax.thy @@ -1,1531 +1,1530 @@ (*:maxLineLen=78:*) theory Inner_Syntax imports Main Base begin chapter \Inner syntax --- the term language \label{ch:inner-syntax}\ text \ The inner syntax of Isabelle provides concrete notation for the main entities of the logical framework, notably \\\-terms with types and type classes. Applications may either extend existing syntactic categories by additional notation, or define new sub-languages that are linked to the standard term language via some explicit markers. For example \<^verbatim>\FOO\~\foo\ could embed the syntax corresponding for some user-defined nonterminal \foo\ --- within the bounds of the given lexical syntax of Isabelle/Pure. The most basic way to specify concrete syntax for logical entities works via mixfix annotations (\secref{sec:mixfix}), which may be usually given as part of the original declaration or via explicit notation commands later on (\secref{sec:notation}). This already covers many needs of concrete syntax without having to understand the full complexity of inner syntax layers. Further details of the syntax engine involves the classical distinction of lexical language versus context-free grammar (see \secref{sec:pure-syntax}), and various mechanisms for \<^emph>\syntax transformations\ (see \secref{sec:syntax-transformations}). \ section \Printing logical entities\ subsection \Diagnostic commands \label{sec:print-diag}\ text \ \begin{matharray}{rcl} @{command_def "typ"}\\<^sup>*\ & : & \context \\ \\ @{command_def "term"}\\<^sup>*\ & : & \context \\ \\ @{command_def "prop"}\\<^sup>*\ & : & \context \\ \\ @{command_def "thm"}\\<^sup>*\ & : & \context \\ \\ @{command_def "prf"}\\<^sup>*\ & : & \context \\ \\ @{command_def "full_prf"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_state"}\\<^sup>*\ & : & \any \\ \\ \end{matharray} These diagnostic commands assist interactive development by printing internal logical entities in a human-readable fashion. \<^rail>\ @@{command typ} @{syntax modes}? @{syntax type} ('::' @{syntax sort})? ; @@{command term} @{syntax modes}? @{syntax term} ; @@{command prop} @{syntax modes}? @{syntax prop} ; @@{command thm} @{syntax modes}? @{syntax thms} ; ( @@{command prf} | @@{command full_prf} ) @{syntax modes}? @{syntax thms}? ; @@{command print_state} @{syntax modes}? ; @{syntax_def modes}: '(' (@{syntax name} + ) ')' \ \<^descr> @{command "typ"}~\\\ reads and prints a type expression according to the current context. \<^descr> @{command "typ"}~\\ :: s\ uses type-inference to determine the most general way to make \\\ conform to sort \s\. For concrete \\\ this checks if the type belongs to that sort. Dummy type parameters ``\_\'' (underscore) are assigned to fresh type variables with most general sorts, according the the principles of type-inference. \<^descr> @{command "term"}~\t\ and @{command "prop"}~\\\ read, type-check and print terms or propositions according to the current theory or proof context; the inferred type of \t\ is output as well. Note that these commands are also useful in inspecting the current environment of term abbreviations. \<^descr> @{command "thm"}~\a\<^sub>1 \ a\<^sub>n\ retrieves theorems from the current theory or proof context. Note that any attributes included in the theorem specifications are applied to a temporary context derived from the current theory or proof; the result is discarded, i.e.\ attributes involved in \a\<^sub>1, \, a\<^sub>n\ do not have any permanent effect. \<^descr> @{command "prf"} displays the (compact) proof term of the current proof state (if present), or of the given theorems. Note that this requires an underlying logic image with proof terms enabled, e.g. \HOL-Proofs\. \<^descr> @{command "full_prf"} is like @{command "prf"}, but displays the full proof term, i.e.\ also displays information omitted in the compact proof term, which is denoted by ``\_\'' placeholders there. \<^descr> @{command "print_state"} prints the current proof state (if present), including current facts and goals. All of the diagnostic commands above admit a list of \modes\ to be specified, which is appended to the current print mode; see also \secref{sec:print-modes}. Thus the output behavior may be modified according particular print mode features. For example, @{command "print_state"}~\(latex)\ prints the current proof state with mathematical symbols and special characters represented in {\LaTeX} source, according to - the Isabelle style @{cite "isabelle-system"}. + the Isabelle style \<^cite>\"isabelle-system"\. Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more systematic way to include formal items into the printed text document. \ subsection \Details of printed content\ text \ \begin{tabular}{rcll} @{attribute_def show_markup} & : & \attribute\ \\ @{attribute_def show_types} & : & \attribute\ & default \false\ \\ @{attribute_def show_sorts} & : & \attribute\ & default \false\ \\ @{attribute_def show_consts} & : & \attribute\ & default \false\ \\ @{attribute_def show_abbrevs} & : & \attribute\ & default \true\ \\ @{attribute_def show_brackets} & : & \attribute\ & default \false\ \\ @{attribute_def names_long} & : & \attribute\ & default \false\ \\ @{attribute_def names_short} & : & \attribute\ & default \false\ \\ @{attribute_def names_unique} & : & \attribute\ & default \true\ \\ @{attribute_def eta_contract} & : & \attribute\ & default \true\ \\ @{attribute_def goals_limit} & : & \attribute\ & default \10\ \\ @{attribute_def show_main_goal} & : & \attribute\ & default \false\ \\ @{attribute_def show_hyps} & : & \attribute\ & default \false\ \\ @{attribute_def show_tags} & : & \attribute\ & default \false\ \\ @{attribute_def show_question_marks} & : & \attribute\ & default \true\ \\ \end{tabular} \<^medskip> These configuration options control the detail of information that is displayed for types, terms, theorems, goals etc. See also \secref{sec:config}. \<^descr> @{attribute show_markup} controls direct inlining of markup into the printed representation of formal entities --- notably type and sort constraints. This enables Prover IDE users to retrieve that information via tooltips or popups while hovering with the mouse over the output window, for example. Consequently, this option is enabled by default for Isabelle/jEdit. \<^descr> @{attribute show_types} and @{attribute show_sorts} control printing of type constraints for term variables, and sort constraints for type variables. By default, neither of these are shown in output. If @{attribute show_sorts} is enabled, types are always shown as well. In Isabelle/jEdit, manual setting of these options is normally not required thanks to @{attribute show_markup} above. Note that displaying types and sorts may explain why a polymorphic inference rule fails to resolve with some goal, or why a rewrite rule does not apply as expected. \<^descr> @{attribute show_consts} controls printing of types of constants when displaying a goal state. Note that the output can be enormous, because polymorphic constants often occur at several different type instances. \<^descr> @{attribute show_abbrevs} controls folding of constant abbreviations. \<^descr> @{attribute show_brackets} controls bracketing in pretty printed output. If enabled, all sub-expressions of the pretty printing tree will be parenthesized, even if this produces malformed term syntax! This crude way of showing the internal structure of pretty printed entities may occasionally help to diagnose problems with operator priorities, for example. \<^descr> @{attribute names_long}, @{attribute names_short}, and @{attribute names_unique} control the way of printing fully qualified internal names in external form. See also \secref{sec:antiq} for the document antiquotation options of the same names. \<^descr> @{attribute eta_contract} controls \\\-contracted printing of terms. The \\\-contraction law asserts \<^prop>\(\x. f x) \ f\, provided \x\ is not free in \f\. It asserts \<^emph>\extensionality\ of functions: \<^prop>\f \ g\ if \<^prop>\f x \ g x\ for all \x\. Higher-order unification frequently puts terms into a fully \\\-expanded form. For example, if \F\ has type \(\ \ \) \ \\ then its expanded form is \<^term>\\h. F (\x. h x)\. Enabling @{attribute eta_contract} makes Isabelle perform \\\-contractions before printing, so that \<^term>\\h. F (\x. h x)\ appears simply as \F\. Note that the distinction between a term and its \\\-expanded form occasionally matters. While higher-order resolution and rewriting operate modulo \\\\\-conversion, some other tools might look at terms more discretely. \<^descr> @{attribute goals_limit} controls the maximum number of subgoals to be printed. \<^descr> @{attribute show_main_goal} controls whether the main result to be proven should be displayed. This information might be relevant for schematic goals, to inspect the current claim that has been synthesized so far. \<^descr> @{attribute show_hyps} controls printing of implicit hypotheses of local facts. Normally, only those hypotheses are displayed that are \<^emph>\not\ covered by the assumptions of the current context: this situation indicates a fault in some tool being used. By enabling @{attribute show_hyps}, output of \<^emph>\all\ hypotheses can be enforced, which is occasionally useful for diagnostic purposes. \<^descr> @{attribute show_tags} controls printing of extra annotations within theorems, such as internal position information, or the case names being attached by the attribute @{attribute case_names}. Note that the @{attribute tagged} and @{attribute untagged} attributes provide low-level access to the collection of tags associated with a theorem. \<^descr> @{attribute show_question_marks} controls printing of question marks for schematic variables, such as \?x\. Only the leading question mark is affected, the remaining text is unchanged (including proper markup for schematic variables that might be relevant for user interfaces). \ subsection \Alternative print modes \label{sec:print-modes}\ text \ \begin{mldecls} @{define_ML print_mode_value: "unit -> string list"} \\ @{define_ML Print_Mode.with_modes: "string list -> ('a -> 'b) -> 'a -> 'b"} \\ \end{mldecls} The \<^emph>\print mode\ facility allows to modify various operations for printing. Commands like @{command typ}, @{command term}, @{command thm} (see \secref{sec:print-diag}) take additional print modes as optional argument. The underlying ML operations are as follows. \<^descr> \<^ML>\print_mode_value ()\ yields the list of currently active print mode names. This should be understood as symbolic representation of certain individual features for printing (with precedence from left to right). \<^descr> \<^ML>\Print_Mode.with_modes\~\modes f x\ evaluates \f x\ in an execution context where the print mode is prepended by the given \modes\. This provides a thread-safe way to augment print modes. It is also monotonic in the set of mode names: it retains the default print mode that certain user-interfaces might have installed for their proper functioning! \<^medskip> The pretty printer for inner syntax maintains alternative mixfix productions for any print mode name invented by the user, say in commands like @{command notation} or @{command abbreviation}. Mode names can be arbitrary, but the following ones have a specific meaning by convention: \<^item> \<^verbatim>\""\ (the empty string): default mode; implicitly active as last element in the list of modes. \<^item> \<^verbatim>\input\: dummy print mode that is never active; may be used to specify notation that is only available for input. \<^item> \<^verbatim>\internal\ dummy print mode that is never active; used internally in Isabelle/Pure. \<^item> \<^verbatim>\ASCII\: prefer ASCII art over mathematical symbols. \<^item> \<^verbatim>\latex\: additional mode that is active in {\LaTeX} document preparation of Isabelle theory sources; allows to provide alternative output notation. \ section \Mixfix annotations \label{sec:mixfix}\ text \ Mixfix annotations specify concrete \<^emph>\inner syntax\ of Isabelle types and terms. Locally fixed parameters in toplevel theorem statements, locale and class specifications also admit mixfix annotations in a fairly uniform manner. A mixfix annotation describes the concrete syntax, the translation to abstract syntax, and the pretty printing. Special case annotations provide a simple means of specifying infix operators and binders. - Isabelle mixfix syntax is inspired by {\OBJ} @{cite OBJ}. It allows to + Isabelle mixfix syntax is inspired by {\OBJ} \<^cite>\OBJ\. It allows to specify any context-free priority grammar, which is more general than the fixity declarations of ML and Prolog. \<^rail>\ @{syntax_def mixfix}: '(' (@{syntax template} prios? @{syntax nat}? | (@'infix' | @'infixl' | @'infixr') @{syntax template} @{syntax nat} | @'binder' @{syntax template} prio? @{syntax nat} | @'structure') ')' ; @{syntax template}: (string | cartouche) ; prios: '[' (@{syntax nat} + ',') ']' ; prio: '[' @{syntax nat} ']' \ The mixfix \template\ may include literal text, spacing, blocks, and arguments (denoted by ``\_\''); the special symbol ``\<^verbatim>\\\'' (printed as ``\\\'') represents an index argument that specifies an implicit @{keyword "structure"} reference (see also \secref{sec:locale}). Only locally fixed variables may be declared as @{keyword "structure"}. Infix and binder declarations provide common abbreviations for particular mixfix declarations. So in practice, mixfix templates mostly degenerate to literal text for concrete syntax, such as ``\<^verbatim>\++\'' for an infix symbol. \ subsection \The general mixfix form\ text \ In full generality, mixfix declarations work as follows. Suppose a constant \c :: \\<^sub>1 \ \ \\<^sub>n \ \\ is annotated by \(mixfix [p\<^sub>1, \, p\<^sub>n] p)\, where \mixfix\ is a string \d\<^sub>0 _ d\<^sub>1 _ \ _ d\<^sub>n\ consisting of delimiters that surround argument positions as indicated by underscores. Altogether this determines a production for a context-free priority grammar, where for each argument \i\ the syntactic category is determined by \\\<^sub>i\ (with priority \p\<^sub>i\), and the result category is determined from \\\ (with priority \p\). Priority specifications are optional, with default 0 for arguments and 1000 for the result.\<^footnote>\Omitting priorities is prone to syntactic ambiguities unless the delimiter tokens determine fully bracketed notation, as in \if _ then _ else _ fi\.\ Since \\\ may be again a function type, the constant type scheme may have more argument positions than the mixfix pattern. Printing a nested application \c t\<^sub>1 \ t\<^sub>m\ for \m > n\ works by attaching concrete notation only to the innermost part, essentially by printing \(c t\<^sub>1 \ t\<^sub>n) \ t\<^sub>m\ instead. If a term has fewer arguments than specified in the mixfix template, the concrete syntax is ignored. \<^medskip> A mixfix template may also contain additional directives for pretty printing, notably spaces, blocks, and breaks. The general template format is a sequence over any of the following entities. \<^descr> \d\ is a delimiter, namely a non-empty sequence delimiter items of the following form: \<^enum> a control symbol followed by a cartouche \<^enum> a single symbol, excluding the following special characters: \<^medskip> \begin{tabular}{ll} \<^verbatim>\'\ & single quote \\ \<^verbatim>\_\ & underscore \\ \\\ & index symbol \\ \<^verbatim>\(\ & open parenthesis \\ \<^verbatim>\)\ & close parenthesis \\ \<^verbatim>\/\ & slash \\ \\ \\ & cartouche delimiters \\ \end{tabular} \<^medskip> \<^descr> \<^verbatim>\'\ escapes the special meaning of these meta-characters, producing a literal version of the following character, unless that is a blank. A single quote followed by a blank separates delimiters, without affecting printing, but input tokens may have additional white space here. \<^descr> \<^verbatim>\_\ is an argument position, which stands for a certain syntactic category in the underlying grammar. \<^descr> \\\ is an indexed argument position; this is the place where implicit structure arguments can be attached. \<^descr> \s\ is a non-empty sequence of spaces for printing. This and the following specifications do not affect parsing at all. \<^descr> \<^verbatim>\(\\n\ opens a pretty printing block. The optional natural number specifies the block indentation, i.e. how much spaces to add when a line break occurs within the block. The default indentation is 0. \<^descr> \<^verbatim>\(\\\properties\\ opens a pretty printing block, with properties specified within the given text cartouche. The syntax and semantics of the category @{syntax_ref mixfix_properties} is described below. \<^descr> \<^verbatim>\)\ closes a pretty printing block. \<^descr> \<^verbatim>\//\ forces a line break. \<^descr> \<^verbatim>\/\\s\ allows a line break. Here \s\ stands for the string of spaces (zero or more) right after the slash. These spaces are printed if the break is \<^emph>\not\ taken. \<^medskip> Block properties allow more control over the details of pretty-printed output. The concrete syntax is defined as follows. \<^rail>\ @{syntax_def "mixfix_properties"}: (entry *) ; entry: atom ('=' atom)? ; atom: @{syntax short_ident} | @{syntax int} | @{syntax float} | @{syntax cartouche} \ Each @{syntax entry} is a name-value pair: if the value is omitted, it defaults to \<^verbatim>\true\ (intended for Boolean properties). The following standard block properties are supported: \<^item> \indent\ (natural number): the block indentation --- the same as for the simple syntax without block properties. \<^item> \consistent\ (Boolean): this block has consistent breaks (if one break is taken, all breaks are taken). \<^item> \unbreakable\ (Boolean): all possible breaks of the block are disabled (turned into spaces). \<^item> \markup\ (string): the optional name of the markup node. If this is provided, all remaining properties are turned into its XML attributes. This allows to specify free-form PIDE markup, e.g.\ for specialized output. \<^medskip> Note that the general idea of pretty printing with blocks and breaks is - described in @{cite "paulson-ml2"}; it goes back to @{cite "Oppen:1980"}. + described in \<^cite>\"paulson-ml2"\; it goes back to \<^cite>\"Oppen:1980"\. \ subsection \Infixes\ text \ Infix operators are specified by convenient short forms that abbreviate general mixfix annotations as follows: \begin{center} \begin{tabular}{lll} \<^verbatim>\(\@{keyword_def "infix"}~\<^verbatim>\"\\sy\\<^verbatim>\"\ \p\\<^verbatim>\)\ & \\\ & \<^verbatim>\("(_\~\sy\\<^verbatim>\/ _)" [\\p + 1\\<^verbatim>\,\~\p + 1\\<^verbatim>\]\~\p\\<^verbatim>\)\ \\ \<^verbatim>\(\@{keyword_def "infixl"}~\<^verbatim>\"\\sy\\<^verbatim>\"\ \p\\<^verbatim>\)\ & \\\ & \<^verbatim>\("(_\~\sy\\<^verbatim>\/ _)" [\\p\\<^verbatim>\,\~\p + 1\\<^verbatim>\]\~\p\\<^verbatim>\)\ \\ \<^verbatim>\(\@{keyword_def "infixr"}~\<^verbatim>\"\\sy\\<^verbatim>\"\~\p\\<^verbatim>\)\ & \\\ & \<^verbatim>\("(_\~\sy\\<^verbatim>\/ _)" [\\p + 1\\<^verbatim>\,\~\p\\<^verbatim>\]\~\p\\<^verbatim>\)\ \\ \end{tabular} \end{center} The mixfix template \<^verbatim>\"(_\~\sy\\<^verbatim>\/ _)"\ specifies two argument positions; the delimiter is preceded by a space and followed by a space or line break; the entire phrase is a pretty printing block. The alternative notation \<^verbatim>\(\\sy\\<^verbatim>\)\ is introduced in addition. Thus any infix operator may be written in prefix form (as in Haskell), independently of the number of arguments. \ subsection \Binders\ text \ A \<^emph>\binder\ is a variable-binding construct such as a quantifier. The idea to formalize \\x. b\ as \All (\x. b)\ for \All :: ('a \ bool) \ bool\ - already goes back to @{cite church40}. Isabelle declarations of certain + already goes back to \<^cite>\church40\. Isabelle declarations of certain higher-order operators may be annotated with @{keyword_def "binder"} annotations as follows: \begin{center} \c ::\~\<^verbatim>\"\\(\\<^sub>1 \ \\<^sub>2) \ \\<^sub>3\\<^verbatim>\" (\@{keyword "binder"}~\<^verbatim>\"\\sy\\<^verbatim>\" [\\p\\<^verbatim>\]\~\q\\<^verbatim>\)\ \end{center} This introduces concrete binder syntax \sy x. b\, where \x\ is a bound variable of type \\\<^sub>1\, the body \b\ has type \\\<^sub>2\ and the whole term has type \\\<^sub>3\. The optional integer \p\ specifies the syntactic priority of the body; the default is \q\, which is also the priority of the whole construct. Internally, the binder syntax is expanded to something like this: \begin{center} \c_binder ::\~\<^verbatim>\"\\idts \ \\<^sub>2 \ \\<^sub>3\\<^verbatim>\" ("(3\\sy\\<^verbatim>\_./ _)" [0,\~\p\\<^verbatim>\]\~\q\\<^verbatim>\)\ \end{center} Here @{syntax (inner) idts} is the nonterminal symbol for a list of identifiers with optional type constraints (see also \secref{sec:pure-grammar}). The mixfix template \<^verbatim>\"(3\\sy\\<^verbatim>\_./ _)"\ defines argument positions for the bound identifiers and the body, separated by a dot with optional line break; the entire phrase is a pretty printing block of indentation level 3. Note that there is no extra space after \sy\, so it needs to be included user specification if the binder syntax ends with a token that may be continued by an identifier token at the start of @{syntax (inner) idts}. Furthermore, a syntax translation to transforms \c_binder x\<^sub>1 \ x\<^sub>n b\ into iterated application \c (\x\<^sub>1. \ c (\x\<^sub>n. b)\)\. This works in both directions, for parsing and printing. \ section \Explicit notation \label{sec:notation}\ text \ \begin{matharray}{rcll} @{command_def "type_notation"} & : & \local_theory \ local_theory\ \\ @{command_def "no_type_notation"} & : & \local_theory \ local_theory\ \\ @{command_def "notation"} & : & \local_theory \ local_theory\ \\ @{command_def "no_notation"} & : & \local_theory \ local_theory\ \\ @{command_def "write"} & : & \proof(state) \ proof(state)\ \\ \end{matharray} Commands that introduce new logical entities (terms or types) usually allow to provide mixfix annotations on the spot, which is convenient for default notation. Nonetheless, the syntax may be modified later on by declarations for explicit notation. This allows to add or delete mixfix annotations for of existing logical entities within the current context. \<^rail>\ (@@{command type_notation} | @@{command no_type_notation}) @{syntax mode}? \ (@{syntax name} @{syntax mixfix} + @'and') ; (@@{command notation} | @@{command no_notation}) @{syntax mode}? \ (@{syntax name} @{syntax mixfix} + @'and') ; @@{command write} @{syntax mode}? (@{syntax name} @{syntax mixfix} + @'and') \ \<^descr> @{command "type_notation"}~\c (mx)\ associates mixfix syntax with an existing type constructor. The arity of the constructor is retrieved from the context. \<^descr> @{command "no_type_notation"} is similar to @{command "type_notation"}, but removes the specified syntax annotation from the present context. \<^descr> @{command "notation"}~\c (mx)\ associates mixfix syntax with an existing constant or fixed variable. The type declaration of the given entity is retrieved from the context. \<^descr> @{command "no_notation"} is similar to @{command "notation"}, but removes the specified syntax annotation from the present context. \<^descr> @{command "write"} is similar to @{command "notation"}, but works within an Isar proof body. \ section \The Pure syntax \label{sec:pure-syntax}\ subsection \Lexical matters \label{sec:inner-lex}\ text \ The inner lexical syntax vaguely resembles the outer one (\secref{sec:outer-lex}), but some details are different. There are two main categories of inner syntax tokens: \<^enum> \<^emph>\delimiters\ --- the literal tokens occurring in productions of the given priority grammar (cf.\ \secref{sec:priority-grammar}); \<^enum> \<^emph>\named tokens\ --- various categories of identifiers etc. Delimiters override named tokens and may thus render certain identifiers inaccessible. Sometimes the logical context admits alternative ways to refer to the same entity, potentially via qualified names. \<^medskip> The categories for named tokens are defined once and for all as follows, reusing some categories of the outer token syntax (\secref{sec:outer-lex}). \begin{center} \begin{supertabular}{rcl} @{syntax_def (inner) id} & = & @{syntax_ref short_ident} \\ @{syntax_def (inner) longid} & = & @{syntax_ref long_ident} \\ @{syntax_def (inner) var} & = & @{syntax_ref var} \\ @{syntax_def (inner) tid} & = & @{syntax_ref type_ident} \\ @{syntax_def (inner) tvar} & = & @{syntax_ref type_var} \\ @{syntax_def (inner) num_token} & = & @{syntax_ref nat} \\ @{syntax_def (inner) float_token} & = & @{syntax_ref nat}\<^verbatim>\.\@{syntax_ref nat} \\ @{syntax_def (inner) str_token} & = & \<^verbatim>\''\ \\\ \<^verbatim>\''\ \\ @{syntax_def (inner) string_token} & = & \<^verbatim>\"\ \\\ \<^verbatim>\"\ \\ @{syntax_def (inner) cartouche} & = & @{verbatim "\"} \\\ @{verbatim "\"} \\ \end{supertabular} \end{center} The token categories @{syntax (inner) num_token}, @{syntax (inner) float_token}, @{syntax (inner) str_token}, @{syntax (inner) string_token}, and @{syntax (inner) cartouche} are not used in Pure. Object-logics may implement numerals and string literals by adding appropriate syntax declarations, together with some translation functions (e.g.\ see \<^file>\~~/src/HOL/Tools/string_syntax.ML\). The derived categories @{syntax_def (inner) num_const}, and @{syntax_def (inner) float_const}, provide robust access to the respective tokens: the syntax tree holds a syntactic constant instead of a free variable. Formal document comments (\secref{sec:comments}) may be also used within the inner syntax. \ subsection \Priority grammars \label{sec:priority-grammar}\ text \ A context-free grammar consists of a set of \<^emph>\terminal symbols\, a set of \<^emph>\nonterminal symbols\ and a set of \<^emph>\productions\. Productions have the form \A = \\, where \A\ is a nonterminal and \\\ is a string of terminals and nonterminals. One designated nonterminal is called the \<^emph>\root symbol\. The language defined by the grammar consists of all strings of terminals that can be derived from the root symbol by applying productions as rewrite rules. The standard Isabelle parser for inner syntax uses a \<^emph>\priority grammar\. Each nonterminal is decorated by an integer priority: \A\<^sup>(\<^sup>p\<^sup>)\. In a derivation, \A\<^sup>(\<^sup>p\<^sup>)\ may be rewritten using a production \A\<^sup>(\<^sup>q\<^sup>) = \\ only if \p \ q\. Any priority grammar can be translated into a normal context-free grammar by introducing new nonterminals and productions. \<^medskip> Formally, a set of context free productions \G\ induces a derivation relation \\\<^sub>G\ as follows. Let \\\ and \\\ denote strings of terminal or nonterminal symbols. Then \\ A\<^sup>(\<^sup>p\<^sup>) \ \\<^sub>G \ \ \\ holds if and only if \G\ contains some production \A\<^sup>(\<^sup>q\<^sup>) = \\ for \p \ q\. \<^medskip> The following grammar for arithmetic expressions demonstrates how binding power and associativity of operators can be enforced by priorities. \begin{center} \begin{tabular}{rclr} \A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\ & \=\ & \<^verbatim>\(\ \A\<^sup>(\<^sup>0\<^sup>)\ \<^verbatim>\)\ \\ \A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\ & \=\ & \<^verbatim>\0\ \\ \A\<^sup>(\<^sup>0\<^sup>)\ & \=\ & \A\<^sup>(\<^sup>0\<^sup>)\ \<^verbatim>\+\ \A\<^sup>(\<^sup>1\<^sup>)\ \\ \A\<^sup>(\<^sup>2\<^sup>)\ & \=\ & \A\<^sup>(\<^sup>3\<^sup>)\ \<^verbatim>\*\ \A\<^sup>(\<^sup>2\<^sup>)\ \\ \A\<^sup>(\<^sup>3\<^sup>)\ & \=\ & \<^verbatim>\-\ \A\<^sup>(\<^sup>3\<^sup>)\ \\ \end{tabular} \end{center} The choice of priorities determines that \<^verbatim>\-\ binds tighter than \<^verbatim>\*\, which binds tighter than \<^verbatim>\+\. Furthermore \<^verbatim>\+\ associates to the left and \<^verbatim>\*\ to the right. \<^medskip> For clarity, grammars obey these conventions: \<^item> All priorities must lie between 0 and 1000. \<^item> Priority 0 on the right-hand side and priority 1000 on the left-hand side may be omitted. \<^item> The production \A\<^sup>(\<^sup>p\<^sup>) = \\ is written as \A = \ (p)\, i.e.\ the priority of the left-hand side actually appears in a column on the far right. \<^item> Alternatives are separated by \|\. \<^item> Repetition is indicated by dots \(\)\ in an informal but obvious way. Using these conventions, the example grammar specification above takes the form: \begin{center} \begin{tabular}{rclc} \A\ & \=\ & \<^verbatim>\(\ \A\ \<^verbatim>\)\ \\ & \|\ & \<^verbatim>\0\ & \qquad\qquad \\ & \|\ & \A\ \<^verbatim>\+\ \A\<^sup>(\<^sup>1\<^sup>)\ & \(0)\ \\ & \|\ & \A\<^sup>(\<^sup>3\<^sup>)\ \<^verbatim>\*\ \A\<^sup>(\<^sup>2\<^sup>)\ & \(2)\ \\ & \|\ & \<^verbatim>\-\ \A\<^sup>(\<^sup>3\<^sup>)\ & \(3)\ \\ \end{tabular} \end{center} \ subsection \The Pure grammar \label{sec:pure-grammar}\ text \ The priority grammar of the \Pure\ theory is defined approximately like this: \begin{center} \begin{supertabular}{rclr} @{syntax_def (inner) any} & = & \prop | logic\ \\\\ @{syntax_def (inner) prop} & = & \<^verbatim>\(\ \prop\ \<^verbatim>\)\ \\ & \|\ & \prop\<^sup>(\<^sup>4\<^sup>)\ \<^verbatim>\::\ \type\ & \(3)\ \\ & \|\ & \any\<^sup>(\<^sup>3\<^sup>)\ \<^verbatim>\==\ \any\<^sup>(\<^sup>3\<^sup>)\ & \(2)\ \\ & \|\ & \any\<^sup>(\<^sup>3\<^sup>)\ \\\ \any\<^sup>(\<^sup>3\<^sup>)\ & \(2)\ \\ & \|\ & \prop\<^sup>(\<^sup>3\<^sup>)\ \<^verbatim>\&&&\ \prop\<^sup>(\<^sup>2\<^sup>)\ & \(2)\ \\ & \|\ & \prop\<^sup>(\<^sup>2\<^sup>)\ \<^verbatim>\==>\ \prop\<^sup>(\<^sup>1\<^sup>)\ & \(1)\ \\ & \|\ & \prop\<^sup>(\<^sup>2\<^sup>)\ \\\ \prop\<^sup>(\<^sup>1\<^sup>)\ & \(1)\ \\ & \|\ & \<^verbatim>\[|\ \prop\ \<^verbatim>\;\ \\\ \<^verbatim>\;\ \prop\ \<^verbatim>\|]\ \<^verbatim>\==>\ \prop\<^sup>(\<^sup>1\<^sup>)\ & \(1)\ \\ & \|\ & \\\ \prop\ \<^verbatim>\;\ \\\ \<^verbatim>\;\ \prop\ \\\ \\\ \prop\<^sup>(\<^sup>1\<^sup>)\ & \(1)\ \\ & \|\ & \<^verbatim>\!!\ \idts\ \<^verbatim>\.\ \prop\ & \(0)\ \\ & \|\ & \\\ \idts\ \<^verbatim>\.\ \prop\ & \(0)\ \\ & \|\ & \<^verbatim>\OFCLASS\ \<^verbatim>\(\ \type\ \<^verbatim>\,\ \logic\ \<^verbatim>\)\ \\ & \|\ & \<^verbatim>\SORT_CONSTRAINT\ \<^verbatim>\(\ \type\ \<^verbatim>\)\ \\ & \|\ & \<^verbatim>\TERM\ \logic\ \\ & \|\ & \<^verbatim>\PROP\ \aprop\ \\\\ @{syntax_def (inner) aprop} & = & \<^verbatim>\(\ \aprop\ \<^verbatim>\)\ \\ & \|\ & \id | longid | var |\~~\<^verbatim>\_\~~\|\~~\<^verbatim>\...\ \\ & \|\ & \<^verbatim>\CONST\ \id |\~~\<^verbatim>\CONST\ \longid\ \\ & \|\ & \<^verbatim>\XCONST\ \id |\~~\<^verbatim>\XCONST\ \longid\ \\ & \|\ & \logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \ any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\ & \(999)\ \\\\ @{syntax_def (inner) logic} & = & \<^verbatim>\(\ \logic\ \<^verbatim>\)\ \\ & \|\ & \logic\<^sup>(\<^sup>4\<^sup>)\ \<^verbatim>\::\ \type\ & \(3)\ \\ & \|\ & \id | longid | var |\~~\<^verbatim>\_\~~\|\~~\<^verbatim>\...\ \\ & \|\ & \<^verbatim>\CONST\ \id |\~~\<^verbatim>\CONST\ \longid\ \\ & \|\ & \<^verbatim>\XCONST\ \id |\~~\<^verbatim>\XCONST\ \longid\ \\ & \|\ & \logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \ any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)\ & \(999)\ \\ & \|\ & \<^verbatim>\%\ \pttrns\ \<^verbatim>\.\ \any\<^sup>(\<^sup>3\<^sup>)\ & \(3)\ \\ & \|\ & \\\ \pttrns\ \<^verbatim>\.\ \any\<^sup>(\<^sup>3\<^sup>)\ & \(3)\ \\ & \|\ & \<^verbatim>\(==)\~~\|\~~\<^verbatim>\(\\\\\<^verbatim>\)\~~\|\~~\<^verbatim>\(&&&)\ \\ & \|\ & \<^verbatim>\(==>)\~~\|\~~\<^verbatim>\(\\\\\<^verbatim>\)\ \\ & \|\ & \<^verbatim>\TYPE\ \<^verbatim>\(\ \type\ \<^verbatim>\)\ \\\\ @{syntax_def (inner) idt} & = & \<^verbatim>\(\ \idt\ \<^verbatim>\)\~~\| id |\~~\<^verbatim>\_\ \\ & \|\ & \id\ \<^verbatim>\::\ \type\ & \(0)\ \\ & \|\ & \<^verbatim>\_\ \<^verbatim>\::\ \type\ & \(0)\ \\\\ @{syntax_def (inner) index} & = & \<^verbatim>\\<^bsub>\ \logic\<^sup>(\<^sup>0\<^sup>)\ \<^verbatim>\\<^esub>\~~\| | \\ \\\\ @{syntax_def (inner) idts} & = & \idt | idt\<^sup>(\<^sup>1\<^sup>) idts\ & \(0)\ \\\\ @{syntax_def (inner) pttrn} & = & \idt\ \\\\ @{syntax_def (inner) pttrns} & = & \pttrn | pttrn\<^sup>(\<^sup>1\<^sup>) pttrns\ & \(0)\ \\\\ @{syntax_def (inner) type} & = & \<^verbatim>\(\ \type\ \<^verbatim>\)\ \\ & \|\ & \tid | tvar |\~~\<^verbatim>\_\ \\ & \|\ & \tid\ \<^verbatim>\::\ \sort | tvar\~~\<^verbatim>\::\ \sort |\~~\<^verbatim>\_\ \<^verbatim>\::\ \sort\ \\ & \|\ & \type_name | type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) type_name\ \\ & \|\ & \<^verbatim>\(\ \type\ \<^verbatim>\,\ \\\ \<^verbatim>\,\ \type\ \<^verbatim>\)\ \type_name\ \\ & \|\ & \type\<^sup>(\<^sup>1\<^sup>)\ \<^verbatim>\=>\ \type\ & \(0)\ \\ & \|\ & \type\<^sup>(\<^sup>1\<^sup>)\ \\\ \type\ & \(0)\ \\ & \|\ & \<^verbatim>\[\ \type\ \<^verbatim>\,\ \\\ \<^verbatim>\,\ \type\ \<^verbatim>\]\ \<^verbatim>\=>\ \type\ & \(0)\ \\ & \|\ & \<^verbatim>\[\ \type\ \<^verbatim>\,\ \\\ \<^verbatim>\,\ \type\ \<^verbatim>\]\ \\\ \type\ & \(0)\ \\ @{syntax_def (inner) type_name} & = & \id | longid\ \\\\ @{syntax_def (inner) sort} & = & @{syntax class_name}~~\|\~~\<^verbatim>\_\~~\|\~~\<^verbatim>\{}\ \\ & \|\ & \<^verbatim>\{\ @{syntax class_name} \<^verbatim>\,\ \\\ \<^verbatim>\,\ @{syntax class_name} \<^verbatim>\}\ \\ @{syntax_def (inner) class_name} & = & \id | longid\ \\ \end{supertabular} \end{center} \<^medskip> Here literal terminals are printed \<^verbatim>\verbatim\; see also \secref{sec:inner-lex} for further token categories of the inner syntax. The meaning of the nonterminals defined by the above grammar is as follows: \<^descr> @{syntax_ref (inner) any} denotes any term. \<^descr> @{syntax_ref (inner) prop} denotes meta-level propositions, which are terms of type \<^typ>\prop\. The syntax of such formulae of the meta-logic is carefully distinguished from usual conventions for object-logics. In particular, plain \\\-term notation is \<^emph>\not\ recognized as @{syntax (inner) prop}. \<^descr> @{syntax_ref (inner) aprop} denotes atomic propositions, which are embedded into regular @{syntax (inner) prop} by means of an explicit \<^verbatim>\PROP\ token. Terms of type \<^typ>\prop\ with non-constant head, e.g.\ a plain variable, are printed in this form. Constants that yield type \<^typ>\prop\ are expected to provide their own concrete syntax; otherwise the printed version will appear like @{syntax (inner) logic} and cannot be parsed again as @{syntax (inner) prop}. \<^descr> @{syntax_ref (inner) logic} denotes arbitrary terms of a logical type, excluding type \<^typ>\prop\. This is the main syntactic category of object-logic entities, covering plain \\\-term notation (variables, abstraction, application), plus anything defined by the user. When specifying notation for logical entities, all logical types (excluding \<^typ>\prop\) are \<^emph>\collapsed\ to this single category of @{syntax (inner) logic}. \<^descr> @{syntax_ref (inner) index} denotes an optional index term for indexed syntax. If omitted, it refers to the first @{keyword_ref "structure"} variable in the context. The special dummy ``\\\'' serves as pattern variable in mixfix annotations that introduce indexed notation. \<^descr> @{syntax_ref (inner) idt} denotes identifiers, possibly constrained by types. \<^descr> @{syntax_ref (inner) idts} denotes a sequence of @{syntax_ref (inner) idt}. This is the most basic category for variables in iterated binders, such as \\\ or \\\. \<^descr> @{syntax_ref (inner) pttrn} and @{syntax_ref (inner) pttrns} denote patterns for abstraction, cases bindings etc. In Pure, these categories start as a merely copy of @{syntax (inner) idt} and @{syntax (inner) idts}, respectively. Object-logics may add additional productions for binding forms. \<^descr> @{syntax_ref (inner) type} denotes types of the meta-logic. \<^descr> @{syntax_ref (inner) sort} denotes meta-level sorts. Here are some further explanations of certain syntax features. \<^item> In @{syntax (inner) idts}, note that \x :: nat y\ is parsed as \x :: (nat y)\, treating \y\ like a type constructor applied to \nat\. To avoid this interpretation, write \(x :: nat) y\ with explicit parentheses. \<^item> Similarly, \x :: nat y :: nat\ is parsed as \x :: (nat y :: nat)\. The correct form is \(x :: nat) (y :: nat)\, or \(x :: nat) y :: nat\ if \y\ is last in the sequence of identifiers. \<^item> Type constraints for terms bind very weakly. For example, \x < y :: nat\ is normally parsed as \(x < y) :: nat\, unless \<\ has a very low priority, in which case the input is likely to be ambiguous. The correct form is \x < (y :: nat)\. \<^item> Dummy variables (written as underscore) may occur in different roles. \<^descr> A sort ``\_\'' refers to a vacuous constraint for type variables, which is effectively ignored in type-inference. \<^descr> A type ``\_\'' or ``\_ :: sort\'' acts like an anonymous inference parameter, which is filled-in according to the most general type produced by the type-checking phase. \<^descr> A bound ``\_\'' refers to a vacuous abstraction, where the body does not refer to the binding introduced here. As in the term \<^term>\\x _. x\, which is \\\-equivalent to \\x y. x\. \<^descr> A free ``\_\'' refers to an implicit outer binding. Higher definitional packages usually allow forms like \f x _ = x\. \<^descr> A schematic ``\_\'' (within a term pattern, see \secref{sec:term-decls}) refers to an anonymous variable that is implicitly abstracted over its context of locally bound variables. For example, this allows pattern matching of \{x. f x = g x}\ against \{x. _ = _}\, or even \{_. _ = _}\ by using both bound and schematic dummies. \<^descr> The three literal dots ``\<^verbatim>\...\'' may be also written as ellipsis symbol \<^verbatim>\\\. In both cases this refers to a special schematic variable, which is bound in the context. This special term abbreviation works nicely with calculational reasoning (\secref{sec:calculation}). \<^descr> \<^verbatim>\CONST\ ensures that the given identifier is treated as constant term, and passed through the parse tree in fully internalized form. This is particularly relevant for translation rules (\secref{sec:syn-trans}), notably on the RHS. \<^descr> \<^verbatim>\XCONST\ is similar to \<^verbatim>\CONST\, but retains the constant name as given. This is only relevant to translation rules (\secref{sec:syn-trans}), notably on the LHS. \ subsection \Inspecting the syntax\ text \ \begin{matharray}{rcl} @{command_def "print_syntax"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} \<^descr> @{command "print_syntax"} prints the inner syntax of the current context. The output can be quite large; the most important sections are explained below. \<^descr> \lexicon\ lists the delimiters of the inner token language; see \secref{sec:inner-lex}. \<^descr> \productions\ lists the productions of the underlying priority grammar; see \secref{sec:priority-grammar}. Many productions have an extra \\ \<^bold>\ name\. These names later become the heads of parse trees; they also guide the pretty printer. Productions without such parse tree names are called \<^emph>\copy productions\. Their right-hand side must have exactly one nonterminal symbol (or named token). The parser does not create a new parse tree node for copy productions, but simply returns the parse tree of the right-hand symbol. If the right-hand side of a copy production consists of a single nonterminal without any delimiters, then it is called a \<^emph>\chain production\. Chain productions act as abbreviations: conceptually, they are removed from the grammar by adding new productions. Priority information attached to chain productions is ignored. \<^descr> \print modes\ lists the alternative print modes provided by this grammar; see \secref{sec:print-modes}. \<^descr> \parse_rules\ and \print_rules\ relate to syntax translations (macros); see \secref{sec:syn-trans}. \<^descr> \parse_ast_translation\ and \print_ast_translation\ list sets of constants that invoke translation functions for abstract syntax trees, which are only required in very special situations; see \secref{sec:tr-funs}. \<^descr> \parse_translation\ and \print_translation\ list the sets of constants that invoke regular translation functions; see \secref{sec:tr-funs}. \ subsection \Ambiguity of parsed expressions\ text \ \begin{tabular}{rcll} @{attribute_def syntax_ambiguity_warning} & : & \attribute\ & default \true\ \\ @{attribute_def syntax_ambiguity_limit} & : & \attribute\ & default \10\ \\ \end{tabular} Depending on the grammar and the given input, parsing may be ambiguous. Isabelle lets the Earley parser enumerate all possible parse trees, and then tries to make the best out of the situation. Terms that cannot be type-checked are filtered out, which often leads to a unique result in the end. Unlike regular type reconstruction, which is applied to the whole collection of input terms simultaneously, the filtering stage only treats each given term in isolation. Filtering is also not attempted for individual types or raw ASTs (as required for @{command translations}). Certain warning or error messages are printed, depending on the situation and the given configuration options. Parsing ultimately fails, if multiple results remain after the filtering phase. \<^descr> @{attribute syntax_ambiguity_warning} controls output of explicit warning messages about syntax ambiguity. \<^descr> @{attribute syntax_ambiguity_limit} determines the number of resulting parse trees that are shown as part of the printed message in case of an ambiguity. \ section \Syntax transformations \label{sec:syntax-transformations}\ text \ The inner syntax engine of Isabelle provides separate mechanisms to transform parse trees either via rewrite systems on first-order ASTs (\secref{sec:syn-trans}), or ML functions on ASTs or syntactic \\\-terms (\secref{sec:tr-funs}). This works both for parsing and printing, as outlined in \figref{fig:parse-print}. \begin{figure}[htbp] \begin{center} \begin{tabular}{cl} string & \\ \\\ & lexer + parser \\ parse tree & \\ \\\ & parse AST translation \\ AST & \\ \\\ & AST rewriting (macros) \\ AST & \\ \\\ & parse translation \\ --- pre-term --- & \\ \\\ & print translation \\ AST & \\ \\\ & AST rewriting (macros) \\ AST & \\ \\\ & print AST translation \\ string & \end{tabular} \end{center} \caption{Parsing and printing with translations}\label{fig:parse-print} \end{figure} These intermediate syntax tree formats eventually lead to a pre-term with all names and binding scopes resolved, but most type information still missing. Explicit type constraints might be given by the user, or implicit position information by the system --- both need to be passed-through carefully by syntax transformations. Pre-terms are further processed by the so-called \<^emph>\check\ and \<^emph>\uncheck\ - phases that are intertwined with type-inference (see also @{cite - "isabelle-implementation"}). The latter allows to operate on higher-order + phases that are intertwined with type-inference (see also \<^cite>\"isabelle-implementation"\). The latter allows to operate on higher-order abstract syntax with proper binding and type information already available. As a rule of thumb, anything that manipulates bindings of variables or constants needs to be implemented as syntax transformation (see below). Anything else is better done via check/uncheck: a prominent example application is the @{command abbreviation} concept of Isabelle/Pure. \ subsection \Abstract syntax trees \label{sec:ast}\ text \ The ML datatype \<^ML_type>\Ast.ast\ explicitly represents the intermediate AST format that is used for syntax rewriting (\secref{sec:syn-trans}). It is defined in ML as follows: @{verbatim [display] \datatype ast = Constant of string | Variable of string | Appl of ast list\} An AST is either an atom (constant or variable) or a list of (at least two) subtrees. Occasional diagnostic output of ASTs uses notation that resembles S-expression of LISP. Constant atoms are shown as quoted strings, variable atoms as non-quoted strings and applications as a parenthesized list of subtrees. For example, the AST @{ML [display] \Ast.Appl [Ast.Constant "_abs", Ast.Variable "x", Ast.Variable "t"]\} is pretty-printed as \<^verbatim>\("_abs" x t)\. Note that \<^verbatim>\()\ and \<^verbatim>\(x)\ are excluded as ASTs, because they have too few subtrees. \<^medskip> AST application is merely a pro-forma mechanism to indicate certain syntactic structures. Thus \<^verbatim>\(c a b)\ could mean either term application or type application, depending on the syntactic context. Nested application like \<^verbatim>\(("_abs" x t) u)\ is also possible, but ASTs are definitely first-order: the syntax constant \<^verbatim>\"_abs"\ does not bind the \<^verbatim>\x\ in any way. Proper bindings are introduced in later stages of the term syntax, where \<^verbatim>\("_abs" x t)\ becomes an \<^ML>\Abs\ node and occurrences of \<^verbatim>\x\ in \<^verbatim>\t\ are replaced by bound variables (represented as de-Bruijn indices). \ subsubsection \AST constants versus variables\ text \ Depending on the situation --- input syntax, output syntax, translation patterns --- the distinction of atomic ASTs as \<^ML>\Ast.Constant\ versus \<^ML>\Ast.Variable\ serves slightly different purposes. Input syntax of a term such as \f a b = c\ does not yet indicate the scopes of atomic entities \f, a, b, c\: they could be global constants or local variables, even bound ones depending on the context of the term. \<^ML>\Ast.Variable\ leaves this choice still open: later syntax layers (or translation functions) may capture such a variable to determine its role specifically, to make it a constant, bound variable, free variable etc. In contrast, syntax translations that introduce already known constants would rather do it via \<^ML>\Ast.Constant\ to prevent accidental re-interpretation later on. Output syntax turns term constants into \<^ML>\Ast.Constant\ and variables (free or schematic) into \<^ML>\Ast.Variable\. This information is precise when printing fully formal \\\-terms. \<^medskip> AST translation patterns (\secref{sec:syn-trans}) that represent terms cannot distinguish constants and variables syntactically. Explicit indication of \CONST c\ inside the term language is required, unless \c\ is known as special \<^emph>\syntax constant\ (see also @{command syntax}). It is also possible to use @{command syntax} declarations (without mixfix annotation) to enforce that certain unqualified names are always treated as constant within the syntax machinery. The situation is simpler for ASTs that represent types or sorts, since the concrete syntax already distinguishes type variables from type constants (constructors). So \('a, 'b) foo\ corresponds to an AST application of some constant for \foo\ and variable arguments for \'a\ and \'b\. Note that the postfix application is merely a feature of the concrete syntax, while in the AST the constructor occurs in head position. \ subsubsection \Authentic syntax names\ text \ Naming constant entities within ASTs is another delicate issue. Unqualified names are resolved in the name space tables in the last stage of parsing, after all translations have been applied. Since syntax transformations do not know about this later name resolution, there can be surprises in boundary cases. \<^emph>\Authentic syntax names\ for \<^ML>\Ast.Constant\ avoid this problem: the fully-qualified constant name with a special prefix for its formal category (\class\, \type\, \const\, \fixed\) represents the information faithfully within the untyped AST format. Accidental overlap with free or bound variables is excluded as well. Authentic syntax names work implicitly in the following situations: \<^item> Input of term constants (or fixed variables) that are introduced by concrete syntax via @{command notation}: the correspondence of a particular grammar production to some known term entity is preserved. \<^item> Input of type constants (constructors) and type classes --- thanks to explicit syntactic distinction independently on the context. \<^item> Output of term constants, type constants, type classes --- this information is already available from the internal term to be printed. In other words, syntax transformations that operate on input terms written as prefix applications are difficult to make robust. Luckily, this case rarely occurs in practice, because syntax forms to be translated usually correspond to some concrete notation. \ subsection \Raw syntax and translations \label{sec:syn-trans}\ text \ \begin{tabular}{rcll} @{command_def "nonterminal"} & : & \theory \ theory\ \\ @{command_def "syntax"} & : & \local_theory \ local_theory\ \\ @{command_def "no_syntax"} & : & \local_theory \ local_theory\ \\ @{command_def "translations"} & : & \theory \ theory\ \\ @{command_def "no_translations"} & : & \theory \ theory\ \\ @{attribute_def syntax_ast_trace} & : & \attribute\ & default \false\ \\ @{attribute_def syntax_ast_stats} & : & \attribute\ & default \false\ \\ \end{tabular} \<^medskip> Unlike mixfix notation for existing formal entities (\secref{sec:notation}), raw syntax declarations provide full access to the priority grammar of the inner syntax, without any sanity checks. This includes additional syntactic categories (via @{command nonterminal}) and free-form grammar productions (via @{command syntax}). Additional syntax translations (or macros, via @{command translations}) are required to turn resulting parse trees into proper representations of formal entities again. \<^rail>\ @@{command nonterminal} (@{syntax name} + @'and') ; (@@{command syntax} | @@{command no_syntax}) @{syntax mode}? (constdecl +) ; (@@{command translations} | @@{command no_translations}) (transpat ('==' | '=>' | '<=' | '\' | '\' | '\') transpat +) ; constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}? ; mode: ('(' ( @{syntax name} | @'output' | @{syntax name} @'output' ) ')') ; transpat: ('(' @{syntax name} ')')? @{syntax string} \ \<^descr> @{command "nonterminal"}~\c\ declares a type constructor \c\ (without arguments) to act as purely syntactic type: a nonterminal symbol of the inner syntax. \<^descr> @{command "syntax"}~\(mode) c :: \ (mx)\ augments the priority grammar and the pretty printer table for the given print mode (default \<^verbatim>\""\). An optional keyword @{keyword_ref "output"} means that only the pretty printer table is affected. Following \secref{sec:mixfix}, the mixfix annotation \mx = template ps q\ together with type \\ = \\<^sub>1 \ \ \\<^sub>n \ \\ and specify a grammar production. The \template\ contains delimiter tokens that surround \n\ argument positions (\<^verbatim>\_\). The latter correspond to nonterminal symbols \A\<^sub>i\ derived from the argument types \\\<^sub>i\ as follows: \<^item> \prop\ if \\\<^sub>i = prop\ \<^item> \logic\ if \\\<^sub>i = (\)\\ for logical type constructor \\ \ prop\ \<^item> \any\ if \\\<^sub>i = \\ for type variables \<^item> \\\ if \\\<^sub>i = \\ for nonterminal \\\ (syntactic type constructor) Each \A\<^sub>i\ is decorated by priority \p\<^sub>i\ from the given list \ps\; missing priorities default to 0. The resulting nonterminal of the production is determined similarly from type \\\, with priority \q\ and default 1000. \<^medskip> Parsing via this production produces parse trees \t\<^sub>1, \, t\<^sub>n\ for the argument slots. The resulting parse tree is composed as \c t\<^sub>1 \ t\<^sub>n\, by using the syntax constant \c\ of the syntax declaration. Such syntactic constants are invented on the spot, without formal check wrt.\ existing declarations. It is conventional to use plain identifiers prefixed by a single underscore (e.g.\ \_foobar\). Names should be chosen with care, to avoid clashes with other syntax declarations. \<^medskip> The special case of copy production is specified by \c =\~\<^verbatim>\""\ (empty string). It means that the resulting parse tree \t\ is copied directly, without any further decoration. \<^descr> @{command "no_syntax"}~\(mode) decls\ removes grammar declarations (and translations) resulting from \decls\, which are interpreted in the same manner as for @{command "syntax"} above. \<^descr> @{command "translations"}~\rules\ specifies syntactic translation rules (i.e.\ macros) as first-order rewrite rules on ASTs (\secref{sec:ast}). The theory context maintains two independent lists translation rules: parse rules (\<^verbatim>\=>\ or \\\) and print rules (\<^verbatim>\<=\ or \\\). For convenience, both can be specified simultaneously as parse~/ print rules (\<^verbatim>\==\ or \\\). Translation patterns may be prefixed by the syntactic category to be used for parsing; the default is \logic\ which means that regular term syntax is used. Both sides of the syntax translation rule undergo parsing and parse AST translations \secref{sec:tr-funs}, in order to perform some fundamental normalization like \\x y. b \ \x. \y. b\, but other AST translation rules are \<^emph>\not\ applied recursively here. When processing AST patterns, the inner syntax lexer runs in a different mode that allows identifiers to start with underscore. This accommodates the usual naming convention for auxiliary syntax constants --- those that do not have a logical counter part --- by allowing to specify arbitrary AST applications within the term syntax, independently of the corresponding concrete syntax. Atomic ASTs are distinguished as \<^ML>\Ast.Constant\ versus \<^ML>\Ast.Variable\ as follows: a qualified name or syntax constant declared via @{command syntax}, or parse tree head of concrete notation becomes \<^ML>\Ast.Constant\, anything else \<^ML>\Ast.Variable\. Note that \CONST\ and \XCONST\ within the term language (\secref{sec:pure-grammar}) allow to enforce treatment as constants. AST rewrite rules \(lhs, rhs)\ need to obey the following side-conditions: \<^item> Rules must be left linear: \lhs\ must not contain repeated variables.\<^footnote>\The deeper reason for this is that AST equality is not well-defined: different occurrences of the ``same'' AST could be decorated differently by accidental type-constraints or source position information, for example.\ \<^item> Every variable in \rhs\ must also occur in \lhs\. \<^descr> @{command "no_translations"}~\rules\ removes syntactic translation rules, which are interpreted in the same manner as for @{command "translations"} above. \<^descr> @{attribute syntax_ast_trace} and @{attribute syntax_ast_stats} control diagnostic output in the AST normalization process, when translation rules are applied to concrete input or output. Raw syntax and translations provides a slightly more low-level access to the grammar and the form of resulting parse trees. It is often possible to avoid this untyped macro mechanism, and use type-safe @{command abbreviation} or @{command notation} instead. Some important situations where @{command syntax} and @{command translations} are really need are as follows: \<^item> Iterated replacement via recursive @{command translations}. For example, consider list enumeration \<^term>\[a, b, c, d]\ as defined in theory \<^theory>\HOL.List\. \<^item> Change of binding status of variables: anything beyond the built-in @{keyword "binder"} mixfix annotation requires explicit syntax translations. For example, consider the set comprehension syntax \<^term>\{x. P}\ as defined in theory \<^theory>\HOL.Set\. \ subsubsection \Applying translation rules\ text \ As a term is being parsed or printed, an AST is generated as an intermediate form according to \figref{fig:parse-print}. The AST is normalized by applying translation rules in the manner of a first-order term rewriting system. We first examine how a single rule is applied. Let \t\ be the abstract syntax tree to be normalized and \(lhs, rhs)\ some translation rule. A subtree \u\ of \t\ is called \<^emph>\redex\ if it is an instance of \lhs\; in this case the pattern \lhs\ is said to match the object \u\. A redex matched by \lhs\ may be replaced by the corresponding instance of \rhs\, thus \<^emph>\rewriting\ the AST \t\. Matching requires some notion of \<^emph>\place-holders\ in rule patterns: \<^ML>\Ast.Variable\ serves this purpose. More precisely, the matching of the object \u\ against the pattern \lhs\ is performed as follows: \<^item> Objects of the form \<^ML>\Ast.Variable\~\x\ or \<^ML>\Ast.Constant\~\x\ are matched by pattern \<^ML>\Ast.Constant\~\x\. Thus all atomic ASTs in the object are treated as (potential) constants, and a successful match makes them actual constants even before name space resolution (see also \secref{sec:ast}). \<^item> Object \u\ is matched by pattern \<^ML>\Ast.Variable\~\x\, binding \x\ to \u\. \<^item> Object \<^ML>\Ast.Appl\~\us\ is matched by \<^ML>\Ast.Appl\~\ts\ if \us\ and \ts\ have the same length and each corresponding subtree matches. \<^item> In every other case, matching fails. A successful match yields a substitution that is applied to \rhs\, generating the instance that replaces \u\. Normalizing an AST involves repeatedly applying translation rules until none are applicable. This works yoyo-like: top-down, bottom-up, top-down, etc. At each subtree position, rules are chosen in order of appearance in the theory definitions. The configuration options @{attribute syntax_ast_trace} and @{attribute syntax_ast_stats} might help to understand this process and diagnose problems. \begin{warn} If syntax translation rules work incorrectly, the output of @{command_ref print_syntax} with its \<^emph>\rules\ sections reveals the actual internal forms of AST pattern, without potentially confusing concrete syntax. Recall that AST constants appear as quoted strings and variables without quotes. \end{warn} \begin{warn} If @{attribute_ref eta_contract} is set to \true\, terms will be \\\-contracted \<^emph>\before\ the AST rewriter sees them. Thus some abstraction nodes needed for print rules to match may vanish. For example, \Ball A (\x. P x)\ would contract to \Ball A P\ and the standard print rule would fail to apply. This problem can be avoided by hand-written ML translation functions (see also \secref{sec:tr-funs}), which is in fact the same mechanism used in built-in @{keyword "binder"} declarations. \end{warn} \ subsection \Syntax translation functions \label{sec:tr-funs}\ text \ \begin{matharray}{rcl} @{command_def "parse_ast_translation"} & : & \theory \ theory\ \\ @{command_def "parse_translation"} & : & \theory \ theory\ \\ @{command_def "print_translation"} & : & \theory \ theory\ \\ @{command_def "typed_print_translation"} & : & \theory \ theory\ \\ @{command_def "print_ast_translation"} & : & \theory \ theory\ \\ @{ML_antiquotation_def "class_syntax"} & : & \ML antiquotation\ \\ @{ML_antiquotation_def "type_syntax"} & : & \ML antiquotation\ \\ @{ML_antiquotation_def "const_syntax"} & : & \ML antiquotation\ \\ @{ML_antiquotation_def "syntax_const"} & : & \ML antiquotation\ \\ \end{matharray} Syntax translation functions written in ML admit almost arbitrary manipulations of inner syntax, at the expense of some complexity and obscurity in the implementation. \<^rail>\ ( @@{command parse_ast_translation} | @@{command parse_translation} | @@{command print_translation} | @@{command typed_print_translation} | @@{command print_ast_translation}) @{syntax text} ; (@@{ML_antiquotation class_syntax} | @@{ML_antiquotation type_syntax} | @@{ML_antiquotation const_syntax} | @@{ML_antiquotation syntax_const}) embedded \ \<^descr> @{command parse_translation} etc. declare syntax translation functions to the theory. Any of these commands have a single @{syntax text} argument that refers to an ML expression of appropriate type as follows: \<^medskip> {\footnotesize \begin{tabular}{l} @{command parse_ast_translation} : \\ \quad \<^ML_type>\(string * (Proof.context -> Ast.ast list -> Ast.ast)) list\ \\ @{command parse_translation} : \\ \quad \<^ML_type>\(string * (Proof.context -> term list -> term)) list\ \\ @{command print_translation} : \\ \quad \<^ML_type>\(string * (Proof.context -> term list -> term)) list\ \\ @{command typed_print_translation} : \\ \quad \<^ML_type>\(string * (Proof.context -> typ -> term list -> term)) list\ \\ @{command print_ast_translation} : \\ \quad \<^ML_type>\(string * (Proof.context -> Ast.ast list -> Ast.ast)) list\ \\ \end{tabular}} \<^medskip> The argument list consists of \(c, tr)\ pairs, where \c\ is the syntax name of the formal entity involved, and \tr\ a function that translates a syntax form \c args\ into \tr ctxt args\ (depending on the context). The Isabelle/ML naming convention for parse translations is \c_tr\ and for print translations \c_tr'\. The @{command_ref print_syntax} command displays the sets of names associated with the translation functions of a theory under \parse_ast_translation\ etc. \<^descr> \@{class_syntax c}\, \@{type_syntax c}\, \@{const_syntax c}\ inline the authentic syntax name of the given formal entities into the ML source. This is the fully-qualified logical name prefixed by a special marker to indicate its kind: thus different logical name spaces are properly distinguished within parse trees. \<^descr> \@{const_syntax c}\ inlines the name \c\ of the given syntax constant, having checked that it has been declared via some @{command syntax} commands within the theory context. Note that the usual naming convention makes syntax constants start with underscore, to reduce the chance of accidental clashes with other names occurring in parse trees (unqualified constants etc.). \ subsubsection \The translation strategy\ text \ The different kinds of translation functions are invoked during the transformations between parse trees, ASTs and syntactic terms (cf.\ \figref{fig:parse-print}). Whenever a combination of the form \c x\<^sub>1 \ x\<^sub>n\ is encountered, and a translation function \f\ of appropriate kind is declared for \c\, the result is produced by evaluation of \f [x\<^sub>1, \, x\<^sub>n]\ in ML. For AST translations, the arguments \x\<^sub>1, \, x\<^sub>n\ are ASTs. A combination has the form \<^ML>\Ast.Constant\~\c\ or \<^ML>\Ast.Appl\~\[\\<^ML>\Ast.Constant\~\c, x\<^sub>1, \, x\<^sub>n]\. For term translations, the arguments are terms and a combination has the form \<^ML>\Const\~\(c, \)\ or \<^ML>\Const\~\(c, \) $ x\<^sub>1 $ \ $ x\<^sub>n\. Terms allow more sophisticated transformations than ASTs do, typically involving abstractions and bound variables. \<^emph>\Typed\ print translations may even peek at the type \\\ of the constant they are invoked on, although some information might have been suppressed for term output already. Regardless of whether they act on ASTs or terms, translation functions called during the parsing process differ from those for printing in their overall behaviour: \<^descr>[Parse translations] are applied bottom-up. The arguments are already in translated form. The translations must not fail; exceptions trigger an error message. There may be at most one function associated with any syntactic name. \<^descr>[Print translations] are applied top-down. They are supplied with arguments that are partly still in internal form. The result again undergoes translation; therefore a print translation should not introduce as head the very constant that invoked it. The function may raise exception \<^ML>\Match\ to indicate failure; in this event it has no effect. Multiple functions associated with some syntactic name are tried in the order of declaration in the theory. Only constant atoms --- constructor \<^ML>\Ast.Constant\ for ASTs and \<^ML>\Const\ for terms --- can invoke translation functions. This means that parse translations can only be associated with parse tree heads of concrete syntax, or syntactic constants introduced via other translations. For plain identifiers within the term language, the status of constant versus variable is not yet know during parsing. This is in contrast to print translations, where constants are explicitly known from the given term in its fully internal form. \ subsection \Built-in syntax transformations\ text \ Here are some further details of the main syntax transformation phases of \figref{fig:parse-print}. \ subsubsection \Transforming parse trees to ASTs\ text \ The parse tree is the raw output of the parser. It is transformed into an AST according to some basic scheme that may be augmented by AST translation functions as explained in \secref{sec:tr-funs}. The parse tree is constructed by nesting the right-hand sides of the productions used to recognize the input. Such parse trees are simply lists of tokens and constituent parse trees, the latter representing the nonterminals of the productions. Ignoring AST translation functions, parse trees are transformed to ASTs by stripping out delimiters and copy productions, while retaining some source position information from input tokens. The Pure syntax provides predefined AST translations to make the basic \\\-term structure more apparent within the (first-order) AST representation, and thus facilitate the use of @{command translations} (see also \secref{sec:syn-trans}). This covers ordinary term application, type application, nested abstraction, iterated meta implications and function types. The effect is illustrated on some representative input strings is as follows: \begin{center} \begin{tabular}{ll} input source & AST \\ \hline \f x y z\ & \<^verbatim>\(f x y z)\ \\ \'a ty\ & \<^verbatim>\(ty 'a)\ \\ \('a, 'b)ty\ & \<^verbatim>\(ty 'a 'b)\ \\ \\x y z. t\ & \<^verbatim>\("_abs" x ("_abs" y ("_abs" z t)))\ \\ \\x :: 'a. t\ & \<^verbatim>\("_abs" ("_constrain" x 'a) t)\ \\ \\P; Q; R\ \ S\ & \<^verbatim>\("Pure.imp" P ("Pure.imp" Q ("Pure.imp" R S)))\ \\ \['a, 'b, 'c] \ 'd\ & \<^verbatim>\("fun" 'a ("fun" 'b ("fun" 'c 'd)))\ \\ \end{tabular} \end{center} Note that type and sort constraints may occur in further places --- translations need to be ready to cope with them. The built-in syntax transformation from parse trees to ASTs insert additional constraints that represent source positions. \ subsubsection \Transforming ASTs to terms\ text \ After application of macros (\secref{sec:syn-trans}), the AST is transformed into a term. This term still lacks proper type information, but it might contain some constraints consisting of applications with head \<^verbatim>\_constrain\, where the second argument is a type encoded as a pre-term within the syntax. Type inference later introduces correct types, or indicates type errors in the input. Ignoring parse translations, ASTs are transformed to terms by mapping AST constants to term constants, AST variables to term variables or constants (according to the name space), and AST applications to iterated term applications. The outcome is still a first-order term. Proper abstractions and bound variables are introduced by parse translations associated with certain syntax constants. Thus \<^verbatim>\("_abs" x x)\ eventually becomes a de-Bruijn term \<^verbatim>\Abs ("x", _, Bound 0)\. \ subsubsection \Printing of terms\ text \ The output phase is essentially the inverse of the input phase. Terms are translated via abstract syntax trees into pretty-printed text. Ignoring print translations, the transformation maps term constants, variables and applications to the corresponding constructs on ASTs. Abstractions are mapped to applications of the special constant \<^verbatim>\_abs\ as seen before. Type constraints are represented via special \<^verbatim>\_constrain\ forms, according to various policies of type annotation determined elsewhere. Sort constraints of type variables are handled in a similar fashion. After application of macros (\secref{sec:syn-trans}), the AST is finally pretty-printed. The built-in print AST translations reverse the corresponding parse AST translations. \<^medskip> For the actual printing process, the priority grammar (\secref{sec:priority-grammar}) plays a vital role: productions are used as templates for pretty printing, with argument slots stemming from nonterminals, and syntactic sugar stemming from literal tokens. Each AST application with constant head \c\ and arguments \t\<^sub>1\, \dots, \t\<^sub>n\ (for \n = 0\ the AST is just the constant \c\ itself) is printed according to the first grammar production of result name \c\. The required syntax priority of the argument slot is given by its nonterminal \A\<^sup>(\<^sup>p\<^sup>)\. The argument \t\<^sub>i\ that corresponds to the position of \A\<^sup>(\<^sup>p\<^sup>)\ is printed recursively, and then put in parentheses \<^emph>\if\ its priority \p\ requires this. The resulting output is concatenated with the syntactic sugar according to the grammar production. If an AST application \(c x\<^sub>1 \ x\<^sub>m)\ has more arguments than the corresponding production, it is first split into \((c x\<^sub>1 \ x\<^sub>n) x\<^sub>n\<^sub>+\<^sub>1 \ x\<^sub>m)\ and then printed recursively as above. Applications with too few arguments or with non-constant head or without a corresponding production are printed in prefix-form like \f t\<^sub>1 \ t\<^sub>n\ for terms. Multiple productions associated with some name \c\ are tried in order of appearance within the grammar. An occurrence of some AST variable \x\ is printed as \x\ outright. \<^medskip> White space is \<^emph>\not\ inserted automatically. If blanks (or breaks) are required to separate tokens, they need to be specified in the mixfix declaration (\secref{sec:mixfix}). \ end diff --git a/src/Doc/Isar_Ref/Outer_Syntax.thy b/src/Doc/Isar_Ref/Outer_Syntax.thy --- a/src/Doc/Isar_Ref/Outer_Syntax.thy +++ b/src/Doc/Isar_Ref/Outer_Syntax.thy @@ -1,606 +1,605 @@ (*:maxLineLen=78:*) theory Outer_Syntax imports Main Base begin chapter \Outer syntax --- the theory language \label{ch:outer-syntax}\ text \ The rather generic framework of Isabelle/Isar syntax emerges from three main syntactic categories: \<^emph>\commands\ of the top-level Isar engine (covering theory and proof elements), \<^emph>\methods\ for general goal refinements (analogous to traditional ``tactics''), and \<^emph>\attributes\ for operations on facts (within a certain context). Subsequently we give a reference of basic syntactic entities underlying Isabelle/Isar syntax in a bottom-up manner. Concrete theory and proof language elements will be introduced later on. \<^medskip> In order to get started with writing well-formed Isabelle/Isar documents, the most important aspect to be noted is the difference of \<^emph>\inner\ versus \<^emph>\outer\ syntax. Inner syntax is that of Isabelle types and terms of the logic, while outer syntax is that of Isabelle/Isar theory sources (specifications and proofs). As a general rule, inner syntax entities may occur only as \<^emph>\atomic entities\ within outer syntax. For example, the string \<^verbatim>\"x + y"\ and identifier \<^verbatim>\z\ are legal term specifications within a theory, while \<^verbatim>\x + y\ without quotes is not. Printed theory documents usually omit quotes to gain readability (this is a - matter of {\LaTeX} macro setup, say via \<^verbatim>\\isabellestyle\, see also @{cite - "isabelle-system"}). Experienced users of Isabelle/Isar may easily + matter of {\LaTeX} macro setup, say via \<^verbatim>\\isabellestyle\, see also \<^cite>\"isabelle-system"\). Experienced users of Isabelle/Isar may easily reconstruct the lost technical information, while mere readers need not care about quotes at all. \ section \Commands\ text \ \begin{matharray}{rcl} @{command_def "print_commands"}\\<^sup>*\ & : & \any \\ \\ @{command_def "help"}\\<^sup>*\ & : & \any \\ \\ \end{matharray} \<^rail>\ @@{command help} (@{syntax name} * ) \ \<^descr> @{command "print_commands"} prints all outer syntax keywords and commands. \<^descr> @{command "help"}~\pats\ retrieves outer syntax commands according to the specified name patterns. \ subsubsection \Examples\ text \ Some common diagnostic commands are retrieved like this (according to usual naming conventions): \ help "print" help "find" section \Lexical matters \label{sec:outer-lex}\ text \ The outer lexical syntax consists of three main categories of syntax tokens: \<^enum> \<^emph>\major keywords\ --- the command names that are available in the present logic session; \<^enum> \<^emph>\minor keywords\ --- additional literal tokens required by the syntax of commands; \<^enum> \<^emph>\named tokens\ --- various categories of identifiers etc. Major keywords and minor keywords are guaranteed to be disjoint. This helps user-interfaces to determine the overall structure of a theory text, without knowing the full details of command syntax. Internally, there is some additional information about the kind of major keywords, which approximates the command type (theory command, proof command etc.). Keywords override named tokens. For example, the presence of a command called \<^verbatim>\term\ inhibits the identifier \<^verbatim>\term\, but the string \<^verbatim>\"term"\ can be used instead. By convention, the outer syntax always allows quoted strings in addition to identifiers, wherever a named entity is expected. When tokenizing a given input sequence, the lexer repeatedly takes the longest prefix of the input that forms a valid token. Spaces, tabs, newlines and formfeeds between tokens serve as explicit separators. \<^medskip> The categories for named tokens are defined once and for all as follows. \begin{center} \begin{supertabular}{rcl} @{syntax_def short_ident} & = & \letter (subscript\<^sup>? quasiletter)\<^sup>*\ \\ @{syntax_def long_ident} & = & \short_ident(\\<^verbatim>\.\\short_ident)\<^sup>+\ \\ @{syntax_def sym_ident} & = & \sym\<^sup>+ |\~~\<^verbatim>\\\\<^verbatim>\<\\short_ident\\<^verbatim>\>\ \\ @{syntax_def nat} & = & \digit\<^sup>+\ \\ @{syntax_def float} & = & @{syntax_ref nat}\<^verbatim>\.\@{syntax_ref nat}~~\|\~~\<^verbatim>\-\@{syntax_ref nat}\<^verbatim>\.\@{syntax_ref nat} \\ @{syntax_def term_var} & = & \<^verbatim>\?\\short_ident |\~~\<^verbatim>\?\\short_ident\\<^verbatim>\.\\nat\ \\ @{syntax_def type_ident} & = & \<^verbatim>\'\\short_ident\ \\ @{syntax_def type_var} & = & \<^verbatim>\?\\type_ident |\~~\<^verbatim>\?\\type_ident\\<^verbatim>\.\\nat\ \\ @{syntax_def string} & = & \<^verbatim>\"\ \\\ \<^verbatim>\"\ \\ @{syntax_def altstring} & = & \<^verbatim>\`\ \\\ \<^verbatim>\`\ \\ @{syntax_def cartouche} & = & @{verbatim "\"} \\\ @{verbatim "\"} \\ @{syntax_def verbatim} & = & \<^verbatim>\{*\ \\\ \<^verbatim>\*}\ \\[1ex] \letter\ & = & \latin |\~~\<^verbatim>\\\\<^verbatim>\<\\latin\\<^verbatim>\>\~~\|\~~\<^verbatim>\\\\<^verbatim>\<\\latin latin\\<^verbatim>\>\~~\| greek |\ \\ \subscript\ & = & \<^verbatim>\\<^sub>\ \\ \quasiletter\ & = & \letter | digit |\~~\<^verbatim>\_\~~\|\~~\<^verbatim>\'\ \\ \latin\ & = & \<^verbatim>\a\~~\| \ |\~~\<^verbatim>\z\~~\|\~~\<^verbatim>\A\~~\| \ |\~~\<^verbatim>\Z\ \\ \digit\ & = & \<^verbatim>\0\~~\| \ |\~~\<^verbatim>\9\ \\ \sym\ & = & \<^verbatim>\!\~~\|\~~\<^verbatim>\#\~~\|\~~\<^verbatim>\$\~~\|\~~\<^verbatim>\%\~~\|\~~\<^verbatim>\&\~~\|\~~\<^verbatim>\*\~~\|\~~\<^verbatim>\+\~~\|\~~\<^verbatim>\-\~~\|\~~\<^verbatim>\/\~~\|\ \\ & & \<^verbatim>\<\~~\|\~~\<^verbatim>\=\~~\|\~~\<^verbatim>\>\~~\|\~~\<^verbatim>\?\~~\|\~~\<^verbatim>\@\~~\|\~~\<^verbatim>\^\~~\|\~~\<^verbatim>\_\~~\|\~~\<^verbatim>\|\~~\|\~~\<^verbatim>\~\ \\ \greek\ & = & \<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\ \\ & & \<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\ \\ & & \<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\ \\ & & \<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\ \\ & & \<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\ \\ & & \<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\ \\ & & \<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\ \\ & & \<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\~~\|\~~\<^verbatim>\\\ \\ \end{supertabular} \end{center} A @{syntax_ref term_var} or @{syntax_ref type_var} describes an unknown, which is internally a pair of base name and index (ML type \<^ML_type>\indexname\). These components are either separated by a dot as in \?x.1\ or \?x7.3\ or run together as in \?x1\. The latter form is possible if the base name does not end with digits. If the index is 0, it may be dropped altogether: \?x\ and \?x0\ and \?x.0\ all refer to the same unknown, with basename \x\ and index 0. The syntax of @{syntax_ref string} admits any characters, including newlines; ``\<^verbatim>\"\'' (double-quote) and ``\<^verbatim>\\\'' (backslash) need to be escaped by a backslash; arbitrary character codes may be specified as ``\<^verbatim>\\\\ddd\'', with three decimal digits. Alternative strings according to @{syntax_ref altstring} are analogous, using single back-quotes instead. The body of @{syntax_ref verbatim} may consist of any text not containing ``\<^verbatim>\*}\''; this allows to include quotes without further escapes, but there is no way to escape ``\<^verbatim>\*}\''. Cartouches do not have this limitation. A @{syntax_ref cartouche} consists of arbitrary text, with properly balanced blocks of ``@{verbatim "\"}~\\\~@{verbatim "\"}''. Note that the rendering of cartouche delimiters is usually like this: ``\\ \ \\''. Source comments take the form \<^verbatim>\(*\~\\\~\<^verbatim>\*)\ and may be nested: the text is removed after lexical analysis of the input and thus not suitable for documentation. The Isar syntax also provides proper \<^emph>\document comments\ that are considered as part of the text (see \secref{sec:comments}). Common mathematical symbols such as \\\ are represented in Isabelle as \<^verbatim>\\\. There are infinitely many Isabelle symbols like this, although proper presentation is left to front-end tools such as {\LaTeX} or Isabelle/jEdit. A list of predefined Isabelle symbols that work well with these tools is given in \appref{app:symbols}. Note that \<^verbatim>\\\ does not belong to the \letter\ category, since it is already used differently in the Pure term language. \ section \Common syntax entities\ text \ We now introduce several basic syntactic entities, such as names, terms, and theorem specifications, which are factored out of the actual Isar language elements to be described later. \ subsection \Names\ text \ Entity @{syntax name} usually refers to any name of types, constants, theorems etc.\ Quoted strings provide an escape for non-identifier names or those ruled out by outer syntax keywords (e.g.\ quoted \<^verbatim>\"let"\). \<^rail>\ @{syntax_def name}: @{syntax short_ident} | @{syntax long_ident} | @{syntax sym_ident} | @{syntax nat} | @{syntax string} ; @{syntax_def par_name}: '(' @{syntax name} ')' \ A @{syntax_def system_name} is like @{syntax name}, but it excludes white-space characters and needs to conform to file-name notation. Name components that are special on Windows (e.g.\ \<^verbatim>\CON\, \<^verbatim>\PRN\, \<^verbatim>\AUX\) are excluded on all platforms. \ subsection \Numbers\ text \ The outer lexical syntax (\secref{sec:outer-lex}) admits natural numbers and floating point numbers. These are combined as @{syntax int} and @{syntax real} as follows. \<^rail>\ @{syntax_def int}: @{syntax nat} | '-' @{syntax nat} ; @{syntax_def real}: @{syntax float} | @{syntax int} \ Note that there is an overlap with the category @{syntax name}, which also includes @{syntax nat}. \ subsection \Embedded content\ text \ Entity @{syntax embedded} refers to content of other languages: cartouches allow arbitrary nesting of sub-languages that respect the recursive balancing of cartouche delimiters. Quoted strings are possible as well, but require escaped quotes when nested. As a shortcut, tokens that appear as plain identifiers in the outer language may be used as inner language content without delimiters. \<^rail>\ @{syntax_def embedded}: @{syntax cartouche} | @{syntax string} | @{syntax short_ident} | @{syntax long_ident} | @{syntax sym_ident} | @{syntax term_var} | @{syntax type_ident} | @{syntax type_var} | @{syntax nat} \ \ subsection \Document text\ text \ A chunk of document @{syntax text} is usually given as @{syntax cartouche} \\\\\. For convenience, any of the smaller text unit that conforms to @{syntax name} is admitted as well. \<^rail>\ @{syntax_def text}: @{syntax embedded} \ Typical uses are document markup commands, like \<^theory_text>\chapter\, \<^theory_text>\section\ etc. (\secref{sec:markup}). \ subsection \Document comments \label{sec:comments}\ text \ Formal comments are an integral part of the document, but are logically void and removed from the resulting theory or term content. The output of document preparation (\chref{ch:document-prep}) supports various styles, according to the following kinds of comments. \<^item> Marginal comment of the form \<^verbatim>\\\~\\text\\ or \\\~\\text\\, usually with a single space between the comment symbol and the argument cartouche. The given argument is typeset as regular text, with formal antiquotations (\secref{sec:antiq}). \<^item> Canceled text of the form \<^verbatim>\\<^cancel>\\\text\\ (no white space between the control symbol and the argument cartouche). The argument is typeset as formal Isabelle source and overlaid with a ``strike-through'' pattern, e.g. \<^theory_text>\\<^cancel>\bad\\. \<^item> Raw {\LaTeX} source of the form \<^verbatim>\\<^latex>\\\argument\\ (no white space between the control symbol and the argument cartouche). This allows to augment the generated {\TeX} source arbitrarily, without any sanity checks! These formal comments work uniformly in outer syntax, inner syntax (term language), Isabelle/ML, and some other embedded languages of Isabelle. \ subsection \Type classes, sorts and arities\ text \ Classes are specified by plain names. Sorts have a very simple inner syntax, which is either a single class name \c\ or a list \{c\<^sub>1, \, c\<^sub>n}\ referring to the intersection of these classes. The syntax of type arities is given directly at the outer level. \<^rail>\ @{syntax_def classdecl}: @{syntax name} (('<' | '\') (@{syntax name} + ','))? ; @{syntax_def sort}: @{syntax embedded} ; @{syntax_def arity}: ('(' (@{syntax sort} + ',') ')')? @{syntax sort} \ \ subsection \Types and terms \label{sec:types-terms}\ text \ The actual inner Isabelle syntax, that of types and terms of the logic, is far too sophisticated in order to be modelled explicitly at the outer theory level. Basically, any such entity has to be quoted to turn it into a single token (the parsing and type-checking is performed internally later). For convenience, a slightly more liberal convention is adopted: quotes may be omitted for any type or term that is already atomic at the outer level. For example, one may just write \<^verbatim>\x\ instead of quoted \<^verbatim>\"x"\. Note that symbolic identifiers (e.g.\ \<^verbatim>\++\ or \\\ are available as well, provided these have not been superseded by commands or other keywords already (such as \<^verbatim>\=\ or \<^verbatim>\+\). \<^rail>\ @{syntax_def type}: @{syntax embedded} ; @{syntax_def term}: @{syntax embedded} ; @{syntax_def prop}: @{syntax embedded} \ Positional instantiations are specified as a sequence of terms, or the placeholder ``\_\'' (underscore), which means to skip a position. \<^rail>\ @{syntax_def inst}: '_' | @{syntax term} ; @{syntax_def insts}: (@{syntax inst} *) \ Named instantiations are specified as pairs of assignments \v = t\, which refer to schematic variables in some theorem that is instantiated. Both type and terms instantiations are admitted, and distinguished by the usual syntax of variable names. \<^rail>\ @{syntax_def named_inst}: variable '=' (type | term) ; @{syntax_def named_insts}: (named_inst @'and' +) ; variable: @{syntax name} | @{syntax term_var} | @{syntax type_ident} | @{syntax type_var} \ Type declarations and definitions usually refer to @{syntax typespec} on the left-hand side. This models basic type constructor application at the outer syntax level. Note that only plain postfix notation is available here, but no infixes. \<^rail>\ @{syntax_def typeargs}: (() | @{syntax type_ident} | '(' ( @{syntax type_ident} + ',' ) ')') ; @{syntax_def typeargs_sorts}: (() | (@{syntax type_ident} ('::' @{syntax sort})?) | '(' ( (@{syntax type_ident} ('::' @{syntax sort})?) + ',' ) ')') ; @{syntax_def typespec}: @{syntax typeargs} @{syntax name} ; @{syntax_def typespec_sorts}: @{syntax typeargs_sorts} @{syntax name} \ \ subsection \Term patterns and declarations \label{sec:term-decls}\ text \ Wherever explicit propositions (or term fragments) occur in a proof text, casual binding of schematic term variables may be given specified via patterns of the form ``\<^theory_text>\(is p\<^sub>1 \ p\<^sub>n)\''. This works both for @{syntax term} and @{syntax prop}. \<^rail>\ @{syntax_def term_pat}: '(' (@'is' @{syntax term} +) ')' ; @{syntax_def prop_pat}: '(' (@'is' @{syntax prop} +) ')' \ \<^medskip> Declarations of local variables \x :: \\ and logical propositions \a : \\ represent different views on the same principle of introducing a local scope. In practice, one may usually omit the typing of @{syntax vars} (due to type-inference), and the naming of propositions (due to implicit references of current facts). In any case, Isar proof elements usually admit to introduce multiple such items simultaneously. \<^rail>\ @{syntax_def vars}: (((@{syntax name} +) ('::' @{syntax type})? | @{syntax name} ('::' @{syntax type})? @{syntax mixfix}) + @'and') ; @{syntax_def props}: @{syntax thmdecl}? (@{syntax prop} @{syntax prop_pat}? +) ; @{syntax_def props'}: (@{syntax prop} @{syntax prop_pat}? +) \ The treatment of multiple declarations corresponds to the complementary focus of @{syntax vars} versus @{syntax props}. In ``\x\<^sub>1 \ x\<^sub>n :: \\'' the typing refers to all variables, while in \a: \\<^sub>1 \ \\<^sub>n\ the naming refers to all propositions collectively. Isar language elements that refer to @{syntax vars} or @{syntax props} typically admit separate typings or namings via another level of iteration, with explicit @{keyword_ref "and"} separators; e.g.\ see @{command "fix"} and @{command "assume"} in \secref{sec:proof-context}. \ subsection \Attributes and theorems \label{sec:syn-att}\ text \ Attributes have their own ``semi-inner'' syntax, in the sense that input conforming to @{syntax args} below is parsed by the attribute a second time. The attribute argument specifications may be any sequence of atomic entities (identifiers, strings etc.), or properly bracketed argument lists. Below @{syntax atom} refers to any atomic entity, including any @{syntax keyword} conforming to @{syntax sym_ident}. \<^rail>\ @{syntax_def atom}: @{syntax name} | @{syntax type_ident} | @{syntax type_var} | @{syntax term_var} | @{syntax nat} | @{syntax float} | @{syntax keyword} | @{syntax cartouche} ; arg: @{syntax atom} | '(' @{syntax args} ')' | '[' @{syntax args} ']' ; @{syntax_def args}: arg * ; @{syntax_def attributes}: '[' (@{syntax name} @{syntax args} * ',') ']' \ Theorem specifications come in several flavors: @{syntax axmdecl} and @{syntax thmdecl} usually refer to axioms, assumptions or results of goal statements, while @{syntax thmdef} collects lists of existing theorems. Existing theorems are given by @{syntax thm} and @{syntax thms}, the former requires an actual singleton result. There are three forms of theorem references: \<^enum> named facts \a\, \<^enum> selections from named facts \a(i)\ or \a(j - k)\, \<^enum> literal fact propositions using token syntax @{syntax_ref altstring} \<^verbatim>\`\\\\\<^verbatim>\`\ or @{syntax_ref cartouche} \\\\\ (see also method @{method_ref fact}). Any kind of theorem specification may include lists of attributes both on the left and right hand sides; attributes are applied to any immediately preceding fact. If names are omitted, the theorems are not stored within the theorem database of the theory or proof context, but any given attributes are applied nonetheless. An extra pair of brackets around attributes (like ``\[[simproc a]]\'') abbreviates a theorem reference involving an internal dummy fact, which will be ignored later on. So only the effect of the attribute on the background context will persist. This form of in-place declarations is particularly useful with commands like @{command "declare"} and @{command "using"}. \<^rail>\ @{syntax_def axmdecl}: @{syntax name} @{syntax attributes}? ':' ; @{syntax_def thmbind}: @{syntax name} @{syntax attributes} | @{syntax name} | @{syntax attributes} ; @{syntax_def thmdecl}: thmbind ':' ; @{syntax_def thmdef}: thmbind '=' ; @{syntax_def thm}: (@{syntax name} selection? | @{syntax altstring} | @{syntax cartouche}) @{syntax attributes}? | '[' @{syntax attributes} ']' ; @{syntax_def thms}: @{syntax thm} + ; selection: '(' ((@{syntax nat} | @{syntax nat} '-' @{syntax nat}?) + ',') ')' \ \ subsection \Structured specifications\ text \ Structured specifications use propositions with explicit notation for the ``eigen-context'' to describe rule structure: \\x. A x \ \ \ B x\ is expressed as \<^theory_text>\B x if A x and \ for x\. It is also possible to use dummy terms ``\_\'' (underscore) to refer to locally fixed variables anonymously. Multiple specifications are delimited by ``\|\'' to emphasize separate cases: each with its own scope of inferred types for free variables. \<^rail>\ @{syntax_def for_fixes}: (@'for' @{syntax vars})? ; @{syntax_def multi_specs}: (@{syntax structured_spec} + '|') ; @{syntax_def structured_spec}: @{syntax thmdecl}? @{syntax prop} @{syntax spec_prems} @{syntax for_fixes} ; @{syntax_def spec_prems}: (@'if' ((@{syntax prop}+) + @'and'))? ; @{syntax_def specification}: @{syntax vars} @'where' @{syntax multi_specs} \ \ section \Diagnostic commands\ text \ \begin{matharray}{rcl} @{command_def "print_theory"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_definitions"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_methods"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_attributes"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_theorems"}\\<^sup>*\ & : & \context \\ \\ @{command_def "find_theorems"}\\<^sup>*\ & : & \context \\ \\ @{command_def "find_consts"}\\<^sup>*\ & : & \context \\ \\ @{command_def "thm_deps"}\\<^sup>*\ & : & \context \\ \\ @{command_def "unused_thms"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_facts"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_term_bindings"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} \<^rail>\ (@@{command print_theory} | @@{command print_definitions} | @@{command print_methods} | @@{command print_attributes} | @@{command print_theorems} | @@{command print_facts}) ('!'?) ; @@{command find_theorems} ('(' @{syntax nat}? 'with_dups'? ')')? \ (thm_criterion*) ; thm_criterion: ('-'?) ('name' ':' @{syntax name} | 'intro' | 'elim' | 'dest' | 'solves' | 'simp' ':' @{syntax term} | @{syntax term}) ; @@{command find_consts} (const_criterion*) ; const_criterion: ('-'?) ('name' ':' @{syntax name} | 'strict' ':' @{syntax type} | @{syntax type}) ; @@{command thm_deps} @{syntax thmrefs} ; @@{command unused_thms} ((@{syntax name} +) '-' (@{syntax name} * ))? \ These commands print certain parts of the theory and proof context. Note that there are some further ones available, such as for the set of rules declared for simplifications. \<^descr> @{command "print_theory"} prints the main logical content of the background theory; the ``\!\'' option indicates extra verbosity. \<^descr> @{command "print_definitions"} prints dependencies of definitional specifications within the background theory, which may be constants (\secref{sec:term-definitions}, \secref{sec:overloading}) or types (\secref{sec:types-pure}, \secref{sec:hol-typedef}); the ``\!\'' option indicates extra verbosity. \<^descr> @{command "print_methods"} prints all proof methods available in the current theory context; the ``\!\'' option indicates extra verbosity. \<^descr> @{command "print_attributes"} prints all attributes available in the current theory context; the ``\!\'' option indicates extra verbosity. \<^descr> @{command "print_theorems"} prints theorems of the background theory resulting from the last command; the ``\!\'' option indicates extra verbosity. \<^descr> @{command "print_facts"} prints all local facts of the current context, both named and unnamed ones; the ``\!\'' option indicates extra verbosity. \<^descr> @{command "print_term_bindings"} prints all term bindings that are present in the context. \<^descr> @{command "find_theorems"}~\criteria\ retrieves facts from the theory or proof context matching all of given search criteria. The criterion \name: p\ selects all theorems whose fully qualified name matches pattern \p\, which may contain ``\*\'' wildcards. The criteria \intro\, \elim\, and \dest\ select theorems that match the current goal as introduction, elimination or destruction rules, respectively. The criterion \solves\ returns all rules that would directly solve the current goal. The criterion \simp: t\ selects all rewrite rules whose left-hand side matches the given term. The criterion term \t\ selects all theorems that contain the pattern \t\ -- as usual, patterns may contain occurrences of the dummy ``\_\'', schematic variables, and type constraints. Criteria can be preceded by ``\-\'' to select theorems that do \<^emph>\not\ match. Note that giving the empty list of criteria yields \<^emph>\all\ currently known facts. An optional limit for the number of printed facts may be given; the default is 40. By default, duplicates are removed from the search result. Use \with_dups\ to display duplicates. \<^descr> @{command "find_consts"}~\criteria\ prints all constants whose type meets all of the given criteria. The criterion \strict: ty\ is met by any type that matches the type pattern \ty\. Patterns may contain both the dummy type ``\_\'' and sort constraints. The criterion \ty\ is similar, but it also matches against subtypes. The criterion \name: p\ and the prefix ``\-\'' function as described for @{command "find_theorems"}. \<^descr> @{command "thm_deps"}~\thms\ prints immediate theorem dependencies, i.e.\ the union of all theorems that are used directly to prove the argument facts, without going deeper into the dependency graph. \<^descr> @{command "unused_thms"}~\A\<^sub>1 \ A\<^sub>m - B\<^sub>1 \ B\<^sub>n\ displays all theorems that are proved in theories \B\<^sub>1 \ B\<^sub>n\ or their parents but not in \A\<^sub>1 \ A\<^sub>m\ or their parents and that are never used. If \n\ is \0\, the end of the range of theories defaults to the current theory. If no range is specified, only the unused theorems in the current theory are displayed. \ end diff --git a/src/Doc/Isar_Ref/Preface.thy b/src/Doc/Isar_Ref/Preface.thy --- a/src/Doc/Isar_Ref/Preface.thy +++ b/src/Doc/Isar_Ref/Preface.thy @@ -1,76 +1,76 @@ (*:maxLineLen=78:*) theory Preface imports Main Base begin text \ The \<^emph>\Isabelle\ system essentially provides a generic infrastructure for building deductive systems (programmed in Standard ML), with a special focus on interactive theorem proving in higher-order logics. Many years ago, even end-users would refer to certain ML functions (goal commands, tactics, tacticals etc.) to pursue their everyday theorem proving tasks. In contrast \<^emph>\Isar\ provides an interpreted language environment of its own, which has been specifically tailored for the needs of theory and proof development. Compared to raw ML, the Isabelle/Isar top-level provides a more robust and comfortable development platform, with proper support for theory development graphs, managed transactions with unlimited undo etc. In its pioneering times, the Isabelle/Isar version of the - \<^emph>\Proof~General\ user interface @{cite proofgeneral and - "Aspinall:TACAS:2000"} has contributed to the + \<^emph>\Proof~General\ user interface \<^cite>\proofgeneral and + "Aspinall:TACAS:2000"\ has contributed to the success of for interactive theory and proof development in this advanced theorem proving environment, even though it was somewhat biased towards old-style proof scripts. The more recent - Isabelle/jEdit Prover IDE @{cite "Wenzel:2012"} emphasizes the + Isabelle/jEdit Prover IDE \<^cite>\"Wenzel:2012"\ emphasizes the document-oriented approach of Isabelle/Isar again more explicitly. \<^medskip> Apart from the technical advances over bare-bones ML programming, the main purpose of the Isar language is to provide a conceptually different view on machine-checked proofs - @{cite "Wenzel:1999:TPHOL" and "Wenzel-PhD"}. \<^emph>\Isar\ stands for + \<^cite>\"Wenzel:1999:TPHOL" and "Wenzel-PhD"\. \<^emph>\Isar\ stands for \<^emph>\Intelligible semi-automated reasoning\. Drawing from both the traditions of informal mathematical proof texts and high-level programming languages, Isar offers a versatile environment for structured formal proof documents. Thus properly written Isar proofs become accessible to a broader audience than unstructured tactic scripts (which typically only provide operational information for the machine). Writing human-readable proof texts certainly requires some additional efforts by the writer to achieve a good presentation, both of formal and informal parts of the text. On the other hand, human-readable formal texts gain some value in their own right, independently of the mechanic proof-checking process. Despite its grand design of structured proof texts, Isar is able to assimilate the old tactical style as an ``improper'' sub-language. This provides an easy upgrade path for existing tactic scripts, as well as some means for interactive experimentation and debugging of structured proofs. Isabelle/Isar supports a broad range of proof styles, both readable and unreadable ones. \<^medskip> The generic Isabelle/Isar framework (see \chref{ch:isar-framework}) works reasonably well for any Isabelle object-logic that conforms to the natural deduction view of the Isabelle/Pure framework. Specific language elements introduced by Isabelle/HOL are described in \partref{part:hol}. Although the main language elements are already provided by the Isabelle/Pure framework, examples given in the generic parts will usually refer to Isabelle/HOL. \<^medskip> Isar commands may be either \<^emph>\proper\ document constructors, or \<^emph>\improper commands\. Some proof methods and attributes introduced later are classified as improper as well. Improper Isar language elements, which are marked by ``\\<^sup>*\'' in the subsequent chapters; they are often helpful when developing proof documents, but their use is discouraged for the final human-readable outcome. Typical examples are diagnostic commands that print terms or theorems according to the current context; other commands emulate old-style tactical theorem proving. \ end diff --git a/src/Doc/Isar_Ref/Proof.thy b/src/Doc/Isar_Ref/Proof.thy --- a/src/Doc/Isar_Ref/Proof.thy +++ b/src/Doc/Isar_Ref/Proof.thy @@ -1,1421 +1,1421 @@ (*:maxLineLen=78:*) theory Proof imports Main Base begin chapter \Proofs \label{ch:proofs}\ text \ Proof commands perform transitions of Isar/VM machine configurations, which are block-structured, consisting of a stack of nodes with three main components: logical proof context, current facts, and open goals. Isar/VM transitions are typed according to the following three different modes of operation: \<^descr> \proof(prove)\ means that a new goal has just been stated that is now to be \<^emph>\proven\; the next command may refine it by some proof method, and enter a sub-proof to establish the actual result. \<^descr> \proof(state)\ is like a nested theory mode: the context may be augmented by \<^emph>\stating\ additional assumptions, intermediate results etc. \<^descr> \proof(chain)\ is intermediate between \proof(state)\ and \proof(prove)\: existing facts (i.e.\ the contents of the special @{fact_ref this} register) have been just picked up in order to be used when refining the goal claimed next. The proof mode indicator may be understood as an instruction to the writer, telling what kind of operation may be performed next. The corresponding typings of proof commands restricts the shape of well-formed proof texts to particular command sequences. So dynamic arrangements of commands eventually turn out as static texts of a certain structure. \Appref{ap:refcard} gives a simplified grammar of the (extensible) language emerging that way from the different types of proof commands. The main ideas of the overall Isar framework are explained in \chref{ch:isar-framework}. \ section \Proof structure\ subsection \Formal notepad\ text \ \begin{matharray}{rcl} @{command_def "notepad"} & : & \local_theory \ proof(state)\ \\ \end{matharray} \<^rail>\ @@{command notepad} @'begin' ; @@{command end} \ \<^descr> @{command "notepad"}~@{keyword "begin"} opens a proof state without any goal statement. This allows to experiment with Isar, without producing any persistent result. The notepad is closed by @{command "end"}. \ subsection \Blocks\ text \ \begin{matharray}{rcl} @{command_def "next"} & : & \proof(state) \ proof(state)\ \\ @{command_def "{"} & : & \proof(state) \ proof(state)\ \\ @{command_def "}"} & : & \proof(state) \ proof(state)\ \\ \end{matharray} While Isar is inherently block-structured, opening and closing blocks is mostly handled rather casually, with little explicit user-intervention. Any local goal statement automatically opens \<^emph>\two\ internal blocks, which are closed again when concluding the sub-proof (by @{command "qed"} etc.). Sections of different context within a sub-proof may be switched via @{command "next"}, which is just a single block-close followed by block-open again. The effect of @{command "next"} is to reset the local proof context; there is no goal focus involved here! For slightly more advanced applications, there are explicit block parentheses as well. These typically achieve a stronger forward style of reasoning. \<^descr> @{command "next"} switches to a fresh block within a sub-proof, resetting the local context to the initial one. \<^descr> @{command "{"} and @{command "}"} explicitly open and close blocks. Any current facts pass through ``@{command "{"}'' unchanged, while ``@{command "}"}'' causes any result to be \<^emph>\exported\ into the enclosing context. Thus fixed variables are generalized, assumptions discharged, and local definitions unfolded (cf.\ \secref{sec:proof-context}). There is no difference of @{command "assume"} and @{command "presume"} in this mode of forward reasoning --- in contrast to plain backward reasoning with the result exported at @{command "show"} time. \ subsection \Omitting proofs\ text \ \begin{matharray}{rcl} @{command_def "oops"} & : & \proof \ local_theory | theory\ \\ \end{matharray} The @{command "oops"} command discontinues the current proof attempt, while considering the partial proof text as properly processed. This is conceptually quite different from ``faking'' actual proofs via @{command_ref "sorry"} (see \secref{sec:proof-steps}): @{command "oops"} does not observe the proof structure at all, but goes back right to the theory level. Furthermore, @{command "oops"} does not produce any result theorem --- there is no intended claim to be able to complete the proof in any way. A typical application of @{command "oops"} is to explain Isar proofs \<^emph>\within\ the system itself, in conjunction with the document preparation tools of Isabelle described in \chref{ch:document-prep}. Thus partial or even wrong proof attempts can be discussed in a logically sound manner. Note that the Isabelle {\LaTeX} macros can be easily adapted to print something like ``\\\'' instead of the keyword ``@{command "oops"}''. \ section \Statements\ subsection \Context elements \label{sec:proof-context}\ text \ \begin{matharray}{rcl} @{command_def "fix"} & : & \proof(state) \ proof(state)\ \\ @{command_def "assume"} & : & \proof(state) \ proof(state)\ \\ @{command_def "presume"} & : & \proof(state) \ proof(state)\ \\ @{command_def "define"} & : & \proof(state) \ proof(state)\ \\ \end{matharray} The logical proof context consists of fixed variables and assumptions. The former closely correspond to Skolem constants, or meta-level universal quantification as provided by the Isabelle/Pure logical framework. Introducing some \<^emph>\arbitrary, but fixed\ variable via ``@{command "fix"}~\x\'' results in a local value that may be used in the subsequent proof as any other variable or constant. Furthermore, any result \\ \[x]\ exported from the context will be universally closed wrt.\ \x\ at the outermost level: \\ \x. \[x]\ (this is expressed in normal form using Isabelle's meta-variables). Similarly, introducing some assumption \\\ has two effects. On the one hand, a local theorem is created that may be used as a fact in subsequent proof steps. On the other hand, any result \\ \ \\ exported from the context becomes conditional wrt.\ the assumption: \\ \ \ \\. Thus, solving an enclosing goal using such a result would basically introduce a new subgoal stemming from the assumption. How this situation is handled depends on the version of assumption command used: while @{command "assume"} insists on solving the subgoal by unification with some premise of the goal, @{command "presume"} leaves the subgoal unchanged in order to be proved later by the user. Local definitions, introduced by ``\<^theory_text>\define x where x = t\'', are achieved by combining ``@{command "fix"}~\x\'' with another version of assumption that causes any hypothetical equation \x \ t\ to be eliminated by the reflexivity rule. Thus, exporting some result \x \ t \ \[x]\ yields \\ \[t]\. \<^rail>\ @@{command fix} @{syntax vars} ; (@@{command assume} | @@{command presume}) concl prems @{syntax for_fixes} ; concl: (@{syntax props} + @'and') ; prems: (@'if' (@{syntax props'} + @'and'))? ; @@{command define} @{syntax vars} @'where' (@{syntax props} + @'and') @{syntax for_fixes} \ \<^descr> @{command "fix"}~\x\ introduces a local variable \x\ that is \<^emph>\arbitrary, but fixed\. \<^descr> @{command "assume"}~\a: \\ and @{command "presume"}~\a: \\ introduce a local fact \\ \ \\ by assumption. Subsequent results applied to an enclosing goal (e.g.\ by @{command_ref "show"}) are handled as follows: @{command "assume"} expects to be able to unify with existing premises in the goal, while @{command "presume"} leaves \\\ as new subgoals. Several lists of assumptions may be given (separated by @{keyword_ref "and"}; the resulting list of current facts consists of all of these concatenated. A structured assumption like \<^theory_text>\assume "B x" if "A x" for x\ is equivalent to \<^theory_text>\assume "\x. A x \ B x"\, but vacuous quantification is avoided: a for-context only effects propositions according to actual use of variables. \<^descr> \<^theory_text>\define x where "x = t"\ introduces a local (non-polymorphic) definition. In results that are exported from the context, \x\ is replaced by \t\. Internally, equational assumptions are added to the context in Pure form, using \x \ t\ instead of \x = t\ or \x \ t\ from the object-logic. When exporting results from the context, \x\ is generalized and the assumption discharged by reflexivity, causing the replacement by \t\. The default name for the definitional fact is \x_def\. Several simultaneous definitions may be given as well, with a collective default name. \<^medskip> It is also possible to abstract over local parameters as follows: \<^theory_text>\define f :: "'a \ 'b" where "f x = t" for x :: 'a\. \ subsection \Term abbreviations \label{sec:term-abbrev}\ text \ \begin{matharray}{rcl} @{command_def "let"} & : & \proof(state) \ proof(state)\ \\ @{keyword_def "is"} & : & syntax \\ \end{matharray} Abbreviations may be either bound by explicit @{command "let"}~\p \ t\ statements, or by annotating assumptions or goal statements with a list of patterns ``\<^theory_text>\(is p\<^sub>1 \ p\<^sub>n)\''. In both cases, higher-order matching is invoked to bind extra-logical term variables, which may be either named schematic variables of the form \?x\, or nameless dummies ``@{variable _}'' (underscore). Note that in the @{command "let"} form the patterns occur on the left-hand side, while the @{keyword "is"} patterns are in postfix position. Polymorphism of term bindings is handled in Hindley-Milner style, similar to ML. Type variables referring to local assumptions or open goal statements are \<^emph>\fixed\, while those of finished results or bound by @{command "let"} may occur in \<^emph>\arbitrary\ instances later. Even though actual polymorphism should be rarely used in practice, this mechanism is essential to achieve proper incremental type-inference, as the user proceeds to build up the Isar proof text from left to right. \<^medskip> Term abbreviations are quite different from local definitions as introduced via @{command "define"} (see \secref{sec:proof-context}). The latter are visible within the logic as actual equations, while abbreviations disappear during the input process just after type checking. Also note that @{command "define"} does not support polymorphism. \<^rail>\ @@{command let} ((@{syntax term} + @'and') '=' @{syntax term} + @'and') \ The syntax of @{keyword "is"} patterns follows @{syntax term_pat} or @{syntax prop_pat} (see \secref{sec:term-decls}). \<^descr> \<^theory_text>\let p\<^sub>1 = t\<^sub>1 and \ p\<^sub>n = t\<^sub>n\ binds any text variables in patterns \p\<^sub>1, \, p\<^sub>n\ by simultaneous higher-order matching against terms \t\<^sub>1, \, t\<^sub>n\. \<^descr> \<^theory_text>\(is p\<^sub>1 \ p\<^sub>n)\ resembles @{command "let"}, but matches \p\<^sub>1, \, p\<^sub>n\ against the preceding statement. Also note that @{keyword "is"} is not a separate command, but part of others (such as @{command "assume"}, @{command "have"} etc.). Some \<^emph>\implicit\ term abbreviations\index{term abbreviations} for goals and facts are available as well. For any open goal, @{variable_ref thesis} refers to its object-level statement, abstracted over any meta-level parameters (if present). Likewise, @{variable_ref this} is bound for fact statements resulting from assumptions or finished goals. In case @{variable this} refers to an object-logic statement that is an application \f t\, then \t\ is bound to the special text variable ``@{variable "\"}'' (three dots). The canonical application of this convenience are calculational proofs (see \secref{sec:calculation}). \ subsection \Facts and forward chaining \label{sec:proof-facts}\ text \ \begin{matharray}{rcl} @{command_def "note"} & : & \proof(state) \ proof(state)\ \\ @{command_def "then"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "from"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "with"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "using"} & : & \proof(prove) \ proof(prove)\ \\ @{command_def "unfolding"} & : & \proof(prove) \ proof(prove)\ \\ @{method_def "use"} & : & \method\ \\ @{fact_def "method_facts"} & : & \fact\ \\ \end{matharray} New facts are established either by assumption or proof of local statements. Any fact will usually be involved in further proofs, either as explicit arguments of proof methods, or when forward chaining towards the next goal via @{command "then"} (and variants); @{command "from"} and @{command "with"} are composite forms involving @{command "note"}. The @{command "using"} elements augments the collection of used facts \<^emph>\after\ a goal has been stated. Note that the special theorem name @{fact_ref this} refers to the most recently established facts, but only \<^emph>\before\ issuing a follow-up claim. \<^rail>\ @@{command note} (@{syntax thmdef}? @{syntax thms} + @'and') ; (@@{command from} | @@{command with} | @@{command using} | @@{command unfolding}) (@{syntax thms} + @'and') ; @{method use} @{syntax thms} @'in' @{syntax method} \ \<^descr> @{command "note"}~\a = b\<^sub>1 \ b\<^sub>n\ recalls existing facts \b\<^sub>1, \, b\<^sub>n\, binding the result as \a\. Note that attributes may be involved as well, both on the left and right hand sides. \<^descr> @{command "then"} indicates forward chaining by the current facts in order to establish the goal to be claimed next. The initial proof method invoked to refine that will be offered the facts to do ``anything appropriate'' (see also \secref{sec:proof-steps}). For example, method @{method (Pure) rule} (see \secref{sec:pure-meth-att}) would typically do an elimination rather than an introduction. Automatic methods usually insert the facts into the goal state before operation. This provides a simple scheme to control relevance of facts in automated proof search. \<^descr> @{command "from"}~\b\ abbreviates ``@{command "note"}~\b\~@{command "then"}''; thus @{command "then"} is equivalent to ``@{command "from"}~\this\''. \<^descr> @{command "with"}~\b\<^sub>1 \ b\<^sub>n\ abbreviates ``@{command "from"}~\b\<^sub>1 \ b\<^sub>n \ this\''; thus the forward chaining is from earlier facts together with the current ones. \<^descr> @{command "using"}~\b\<^sub>1 \ b\<^sub>n\ augments the facts to be used by a subsequent refinement step (such as @{command_ref "apply"} or @{command_ref "proof"}). \<^descr> @{command "unfolding"}~\b\<^sub>1 \ b\<^sub>n\ is structurally similar to @{command "using"}, but unfolds definitional equations \b\<^sub>1 \ b\<^sub>n\ throughout the goal state and facts. See also the proof method @{method_ref unfold}. \<^descr> \<^theory_text>\(use b\<^sub>1 \ b\<^sub>n in method)\ uses the facts in the given method expression. The facts provided by the proof state (via @{command "using"} etc.) are ignored, but it is possible to refer to @{fact method_facts} explicitly. \<^descr> @{fact method_facts} is a dynamic fact that refers to the currently used facts of the goal state. Forward chaining with an empty list of theorems is the same as not chaining at all. Thus ``@{command "from"}~\nothing\'' has no effect apart from entering \prove(chain)\ mode, since @{fact_ref nothing} is bound to the empty list of theorems. Basic proof methods (such as @{method_ref (Pure) rule}) expect multiple facts to be given in their proper order, corresponding to a prefix of the premises of the rule involved. Note that positions may be easily skipped using something like @{command "from"}~\_ \ a \ b\, for example. This involves the trivial rule \PROP \ \ PROP \\, which is bound in Isabelle/Pure as ``@{fact_ref "_"}'' (underscore). Automated methods (such as @{method simp} or @{method auto}) just insert any given facts before their usual operation. Depending on the kind of procedure involved, the order of facts is less significant here. \ subsection \Goals \label{sec:goals}\ text \ \begin{matharray}{rcl} @{command_def "lemma"} & : & \local_theory \ proof(prove)\ \\ @{command_def "theorem"} & : & \local_theory \ proof(prove)\ \\ @{command_def "corollary"} & : & \local_theory \ proof(prove)\ \\ @{command_def "proposition"} & : & \local_theory \ proof(prove)\ \\ @{command_def "schematic_goal"} & : & \local_theory \ proof(prove)\ \\ @{command_def "have"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "show"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "hence"} & : & \proof(state) \ proof(prove)\ \\ @{command_def "thus"} & : & \proof(state) \ proof(prove)\ \\ @{command_def "print_statement"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} From a theory context, proof mode is entered by an initial goal command such as @{command "lemma"}. Within a proof context, new claims may be introduced locally; there are variants to interact with the overall proof structure specifically, such as @{command have} or @{command show}. Goals may consist of multiple statements, resulting in a list of facts eventually. A pending multi-goal is internally represented as a meta-level conjunction (\&&&\), which is usually split into the corresponding number of sub-goals prior to an initial method application, via @{command_ref "proof"} (\secref{sec:proof-steps}) or @{command_ref "apply"} (\secref{sec:tactic-commands}). The @{method_ref induct} method covered in \secref{sec:cases-induct} acts on multiple claims simultaneously. Claims at the theory level may be either in short or long form. A short goal merely consists of several simultaneous propositions (often just one). A long goal includes an explicit context specification for the subsequent conclusion, involving local parameters and assumptions. Here the role of each part of the statement is explicitly marked by separate keywords (see also \secref{sec:locale}); the local assumptions being introduced here are available as @{fact_ref assms} in the proof. Moreover, there are two kinds of conclusions: @{element_def "shows"} states several simultaneous propositions (essentially a big conjunction), while @{element_def "obtains"} claims several simultaneous contexts --- essentially a big disjunction of eliminated parameters and assumptions (see \secref{sec:obtain}). \<^rail>\ (@@{command lemma} | @@{command theorem} | @@{command corollary} | @@{command proposition} | @@{command schematic_goal}) (long_statement | short_statement) ; (@@{command have} | @@{command show} | @@{command hence} | @@{command thus}) stmt cond_stmt @{syntax for_fixes} ; @@{command print_statement} @{syntax modes}? @{syntax thms} ; stmt: (@{syntax props} + @'and') ; cond_stmt: ((@'if' | @'when') stmt)? ; short_statement: stmt (@'if' stmt)? @{syntax for_fixes} ; long_statement: @{syntax thmdecl}? context conclusion ; context: (@{syntax_ref "includes"}?) (@{syntax context_elem} *) ; conclusion: @'shows' stmt | @'obtains' @{syntax obtain_clauses} ; @{syntax_def obtain_clauses}: (@{syntax par_name}? obtain_case + '|') ; @{syntax_def obtain_case}: @{syntax vars} @'where' (@{syntax thmdecl}? (@{syntax prop}+) + @'and') \ \<^descr> @{command "lemma"}~\a: \\ enters proof mode with \\\ as main goal, eventually resulting in some fact \\ \\ to be put back into the target context. A @{syntax long_statement} may build up an initial proof context for the subsequent claim, potentially including local definitions and syntax; see also @{syntax "includes"} in \secref{sec:bundle} and @{syntax context_elem} in \secref{sec:locale}. A @{syntax short_statement} consists of propositions as conclusion, with an option context of premises and parameters, via \<^verbatim>\if\/\<^verbatim>\for\ in postfix notation, corresponding to \<^verbatim>\assumes\/\<^verbatim>\fixes\ in the long prefix notation. Local premises (if present) are called ``\assms\'' for @{syntax long_statement}, and ``\that\'' for @{syntax short_statement}. \<^descr> @{command "theorem"}, @{command "corollary"}, and @{command "proposition"} are the same as @{command "lemma"}. The different command names merely serve as a formal comment in the theory source. \<^descr> @{command "schematic_goal"} is similar to @{command "theorem"}, but allows the statement to contain unbound schematic variables. Under normal circumstances, an Isar proof text needs to specify claims explicitly. Schematic goals are more like goals in Prolog, where certain results are synthesized in the course of reasoning. With schematic statements, the inherent compositionality of Isar proofs is lost, which also impacts performance, because proof checking is forced into sequential mode. \<^descr> @{command "have"}~\a: \\ claims a local goal, eventually resulting in a fact within the current logical context. This operation is completely independent of any pending sub-goals of an enclosing goal statements, so @{command "have"} may be freely used for experimental exploration of potential results within a proof body. \<^descr> @{command "show"}~\a: \\ is like @{command "have"}~\a: \\ plus a second stage to refine some pending sub-goal for each one of the finished result, after having been exported into the corresponding context (at the head of the sub-proof of this @{command "show"} command). To accommodate interactive debugging, resulting rules are printed before being applied internally. Even more, interactive execution of @{command "show"} predicts potential failure and displays the resulting error as a warning beforehand. Watch out for the following message: @{verbatim [display] \Local statement fails to refine any pending goal\} \<^descr> @{command "hence"} expands to ``@{command "then"}~@{command "have"}'' and @{command "thus"} expands to ``@{command "then"}~@{command "show"}''. These conflations are left-over from early history of Isar. The expanded syntax is more orthogonal and improves readability and maintainability of proofs. \<^descr> @{command "print_statement"}~\a\ prints facts from the current theory or proof context in long statement form, according to the syntax for @{command "lemma"} given above. Any goal statement causes some term abbreviations (such as @{variable_ref "?thesis"}) to be bound automatically, see also \secref{sec:term-abbrev}. Structured goal statements involving @{keyword_ref "if"} or @{keyword_ref "when"} define the special fact @{fact_ref that} to refer to these assumptions in the proof body. The user may provide separate names according to the syntax of the statement. \ section \Calculational reasoning \label{sec:calculation}\ text \ \begin{matharray}{rcl} @{command_def "also"} & : & \proof(state) \ proof(state)\ \\ @{command_def "finally"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "moreover"} & : & \proof(state) \ proof(state)\ \\ @{command_def "ultimately"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "print_trans_rules"}\\<^sup>*\ & : & \context \\ \\ @{attribute trans} & : & \attribute\ \\ @{attribute sym} & : & \attribute\ \\ @{attribute symmetric} & : & \attribute\ \\ \end{matharray} Calculational proof is forward reasoning with implicit application of transitivity rules (such those of \=\, \\\, \<\). Isabelle/Isar maintains an auxiliary fact register @{fact_ref calculation} for accumulating results obtained by transitivity composed with the current result. Command @{command "also"} updates @{fact calculation} involving @{fact this}, while @{command "finally"} exhibits the final @{fact calculation} by forward chaining towards the next goal statement. Both commands require valid current facts, i.e.\ may occur only after commands that produce theorems such as @{command "assume"}, @{command "note"}, or some finished proof of @{command "have"}, @{command "show"} etc. The @{command "moreover"} and @{command "ultimately"} commands are similar to @{command "also"} and @{command "finally"}, but only collect further results in @{fact calculation} without applying any rules yet. Also note that the implicit term abbreviation ``\\\'' has its canonical application with calculational proofs. It refers to the argument of the preceding statement. (The argument of a curried infix expression happens to be its right-hand side.) Isabelle/Isar calculations are implicitly subject to block structure in the sense that new threads of calculational reasoning are commenced for any new block (as opened by a local goal, for example). This means that, apart from being able to nest calculations, there is no separate \<^emph>\begin-calculation\ command required. \<^medskip> The Isar calculation proof commands may be defined as follows:\<^footnote>\We suppress internal bookkeeping such as proper handling of block-structure.\ \begin{matharray}{rcl} @{command "also"}\\<^sub>0\ & \equiv & @{command "note"}~\calculation = this\ \\ @{command "also"}\\<^sub>n\<^sub>+\<^sub>1\ & \equiv & @{command "note"}~\calculation = trans [OF calculation this]\ \\[0.5ex] @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~\calculation\ \\[0.5ex] @{command "moreover"} & \equiv & @{command "note"}~\calculation = calculation this\ \\ @{command "ultimately"} & \equiv & @{command "moreover"}~@{command "from"}~\calculation\ \\ \end{matharray} \<^rail>\ (@@{command also} | @@{command finally}) ('(' @{syntax thms} ')')? ; @@{attribute trans} (() | 'add' | 'del') \ \<^descr> @{command "also"}~\(a\<^sub>1 \ a\<^sub>n)\ maintains the auxiliary @{fact calculation} register as follows. The first occurrence of @{command "also"} in some calculational thread initializes @{fact calculation} by @{fact this}. Any subsequent @{command "also"} on the same level of block-structure updates @{fact calculation} by some transitivity rule applied to @{fact calculation} and @{fact this} (in that order). Transitivity rules are picked from the current context, unless alternative rules are given as explicit arguments. \<^descr> @{command "finally"}~\(a\<^sub>1 \ a\<^sub>n)\ maintains @{fact calculation} in the same way as @{command "also"} and then concludes the current calculational thread. The final result is exhibited as fact for forward chaining towards the next goal. Basically, @{command "finally"} abbreviates @{command "also"}~@{command "from"}~@{fact calculation}. Typical idioms for concluding calculational proofs are ``@{command "finally"}~@{command "show"}~\?thesis\~@{command "."}'' and ``@{command "finally"}~@{command "have"}~\\\~@{command "."}''. \<^descr> @{command "moreover"} and @{command "ultimately"} are analogous to @{command "also"} and @{command "finally"}, but collect results only, without applying rules. \<^descr> @{command "print_trans_rules"} prints the list of transitivity rules (for calculational commands @{command "also"} and @{command "finally"}) and symmetry rules (for the @{attribute symmetric} operation and single step elimination patters) of the current context. \<^descr> @{attribute trans} declares theorems as transitivity rules. \<^descr> @{attribute sym} declares symmetry rules, as well as @{attribute "Pure.elim"}\?\ rules. \<^descr> @{attribute symmetric} resolves a theorem with some rule declared as @{attribute sym} in the current context. For example, ``@{command "assume"}~\[symmetric]: x = y\'' produces a swapped fact derived from that assumption. In structured proof texts it is often more appropriate to use an explicit single-step elimination proof, such as ``@{command "assume"}~\x = y\~@{command "then"}~@{command "have"}~\y = x\~@{command ".."}''. \ section \Refinement steps\ subsection \Proof method expressions \label{sec:proof-meth}\ text \ Proof methods are either basic ones, or expressions composed of methods via ``\<^verbatim>\,\'' (sequential composition), ``\<^verbatim>\;\'' (structural composition), ``\<^verbatim>\|\'' (alternative choices), ``\<^verbatim>\?\'' (try), ``\<^verbatim>\+\'' (repeat at least once), ``\<^verbatim>\[\\n\\<^verbatim>\]\'' (restriction to first \n\ subgoals). In practice, proof methods are usually just a comma separated list of @{syntax name}~@{syntax args} specifications. Note that parentheses may be dropped for single method specifications (with no arguments). The syntactic precedence of method combinators is \<^verbatim>\|\ \<^verbatim>\;\ \<^verbatim>\,\ \<^verbatim>\[]\ \<^verbatim>\+\ \<^verbatim>\?\ (from low to high). \<^rail>\ @{syntax_def method}: (@{syntax name} | '(' methods ')') (() | '?' | '+' | '[' @{syntax nat}? ']') ; methods: (@{syntax name} @{syntax args} | @{syntax method}) + (',' | ';' | '|') \ Regular Isar proof methods do \<^emph>\not\ admit direct goal addressing, but refer to the first subgoal or to all subgoals uniformly. Nonetheless, the subsequent mechanisms allow to imitate the effect of subgoal addressing that is known from ML tactics. \<^medskip> Goal \<^emph>\restriction\ means the proof state is wrapped-up in a way that certain subgoals are exposed, and other subgoals are ``parked'' elsewhere. Thus a proof method has no other chance than to operate on the subgoals that are presently exposed. Structural composition ``\m\<^sub>1\\<^verbatim>\;\~\m\<^sub>2\'' means that method \m\<^sub>1\ is applied with restriction to the first subgoal, then \m\<^sub>2\ is applied consecutively with restriction to each subgoal that has newly emerged due to \m\<^sub>1\. This is analogous to the tactic combinator \<^ML_infix>\THEN_ALL_NEW\ in - Isabelle/ML, see also @{cite "isabelle-implementation"}. For example, \(rule + Isabelle/ML, see also \<^cite>\"isabelle-implementation"\. For example, \(rule r; blast)\ applies rule \r\ and then solves all new subgoals by \blast\. Moreover, the explicit goal restriction operator ``\[n]\'' exposes only the first \n\ subgoals (which need to exist), with default \n = 1\. For example, the method expression ``\simp_all[3]\'' simplifies the first three subgoals, while ``\(rule r, simp_all)[]\'' simplifies all new goals that emerge from applying rule \r\ to the originally first one. \<^medskip> Improper methods, notably tactic emulations, offer low-level goal addressing as explicit argument to the individual tactic being involved. Here ``\[!]\'' refers to all goals, and ``\[n-]\'' to all goals starting from \n\. \<^rail>\ @{syntax_def goal_spec}: '[' (@{syntax nat} '-' @{syntax nat} | @{syntax nat} '-' | @{syntax nat} | '!' ) ']' \ \ subsection \Initial and terminal proof steps \label{sec:proof-steps}\ text \ \begin{matharray}{rcl} @{command_def "proof"} & : & \proof(prove) \ proof(state)\ \\ @{command_def "qed"} & : & \proof(state) \ proof(state) | local_theory | theory\ \\ @{command_def "by"} & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{command_def ".."} & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{command_def "."} & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{command_def "sorry"} & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{method_def standard} & : & \method\ \\ \end{matharray} Arbitrary goal refinement via tactics is considered harmful. Structured proof composition in Isar admits proof methods to be invoked in two places only. \<^enum> An \<^emph>\initial\ refinement step @{command_ref "proof"}~\m\<^sub>1\ reduces a newly stated goal to a number of sub-goals that are to be solved later. Facts are passed to \m\<^sub>1\ for forward chaining, if so indicated by \proof(chain)\ mode. \<^enum> A \<^emph>\terminal\ conclusion step @{command_ref "qed"}~\m\<^sub>2\ is intended to solve remaining goals. No facts are passed to \m\<^sub>2\. The only other (proper) way to affect pending goals in a proof body is by @{command_ref "show"}, which involves an explicit statement of what is to be solved eventually. Thus we avoid the fundamental problem of unstructured tactic scripts that consist of numerous consecutive goal transformations, with invisible effects. \<^medskip> As a general rule of thumb for good proof style, initial proof methods should either solve the goal completely, or constitute some well-understood reduction to new sub-goals. Arbitrary automatic proof tools that are prone leave a large number of badly structured sub-goals are no help in continuing the proof document in an intelligible manner. Unless given explicitly by the user, the default initial method is @{method standard}, which subsumes at least @{method_ref (Pure) rule} or its classical variant @{method_ref (HOL) rule}. These methods apply a single standard elimination or introduction rule according to the topmost logical connective involved. There is no separate default terminal method. Any remaining goals are always solved by assumption in the very last step. \<^rail>\ @@{command proof} method? ; @@{command qed} method? ; @@{command "by"} method method? ; (@@{command "."} | @@{command ".."} | @@{command sorry}) \ \<^descr> @{command "proof"}~\m\<^sub>1\ refines the goal by proof method \m\<^sub>1\; facts for forward chaining are passed if so indicated by \proof(chain)\ mode. \<^descr> @{command "qed"}~\m\<^sub>2\ refines any remaining goals by proof method \m\<^sub>2\ and concludes the sub-proof by assumption. If the goal had been \show\, some pending sub-goal is solved as well by the rule resulting from the result \<^emph>\exported\ into the enclosing goal context. Thus \qed\ may fail for two reasons: either \m\<^sub>2\ fails, or the resulting rule does not fit to any pending goal\<^footnote>\This includes any additional ``strong'' assumptions as introduced by @{command "assume"}.\ of the enclosing context. Debugging such a situation might involve temporarily changing @{command "show"} into @{command "have"}, or weakening the local context by replacing occurrences of @{command "assume"} by @{command "presume"}. \<^descr> @{command "by"}~\m\<^sub>1 m\<^sub>2\ is a \<^emph>\terminal proof\\index{proof!terminal}; it abbreviates @{command "proof"}~\m\<^sub>1\~@{command "qed"}~\m\<^sub>2\, but with backtracking across both methods. Debugging an unsuccessful @{command "by"}~\m\<^sub>1 m\<^sub>2\ command can be done by expanding its definition; in many cases @{command "proof"}~\m\<^sub>1\ (or even \apply\~\m\<^sub>1\) is already sufficient to see the problem. \<^descr> ``@{command ".."}'' is a \<^emph>\standard proof\\index{proof!standard}; it abbreviates @{command "by"}~\standard\. \<^descr> ``@{command "."}'' is a \<^emph>\trivial proof\\index{proof!trivial}; it abbreviates @{command "by"}~\this\. \<^descr> @{command "sorry"} is a \<^emph>\fake proof\\index{proof!fake} pretending to solve the pending claim without further ado. This only works in interactive development, or if the @{attribute quick_and_dirty} is enabled. Facts emerging from fake proofs are not the real thing. Internally, the derivation object is tainted by an oracle invocation, which may be inspected via the command @{command "thm_oracles"} (\secref{sec:oracles}). The most important application of @{command "sorry"} is to support experimentation and top-down proof development. \<^descr> @{method standard} refers to the default refinement step of some Isar language elements (notably @{command proof} and ``@{command ".."}''). It is \<^emph>\dynamically scoped\, so the behaviour depends on the application environment. In Isabelle/Pure, @{method standard} performs elementary introduction~/ elimination steps (@{method_ref (Pure) rule}), introduction of type classes (@{method_ref intro_classes}) and locales (@{method_ref intro_locales}). In Isabelle/HOL, @{method standard} also takes classical rules into account (cf.\ \secref{sec:classical}). \ subsection \Fundamental methods and attributes \label{sec:pure-meth-att}\ text \ The following proof methods and attributes refer to basic logical operations of Isar. Further methods and attributes are provided by several generic and object-logic specific tools and packages (see \chref{ch:gen-tools} and \partref{part:hol}). \begin{matharray}{rcl} @{command_def "print_rules"}\\<^sup>*\ & : & \context \\ \\[0.5ex] @{method_def "-"} & : & \method\ \\ @{method_def "goal_cases"} & : & \method\ \\ @{method_def "subproofs"} & : & \method\ \\ @{method_def "fact"} & : & \method\ \\ @{method_def "assumption"} & : & \method\ \\ @{method_def "this"} & : & \method\ \\ @{method_def (Pure) "rule"} & : & \method\ \\ @{attribute_def (Pure) "intro"} & : & \attribute\ \\ @{attribute_def (Pure) "elim"} & : & \attribute\ \\ @{attribute_def (Pure) "dest"} & : & \attribute\ \\ @{attribute_def (Pure) "rule"} & : & \attribute\ \\[0.5ex] @{attribute_def "OF"} & : & \attribute\ \\ @{attribute_def "of"} & : & \attribute\ \\ @{attribute_def "where"} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{method goal_cases} (@{syntax name}*) ; @@{method subproofs} @{syntax method} ; @@{method fact} @{syntax thms}? ; @@{method (Pure) rule} @{syntax thms}? ; rulemod: ('intro' | 'elim' | 'dest') ((('!' | () | '?') @{syntax nat}?) | 'del') ':' @{syntax thms} ; (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? ; @@{attribute (Pure) rule} 'del' ; @@{attribute OF} @{syntax thms} ; @@{attribute of} @{syntax insts} ('concl' ':' @{syntax insts})? @{syntax for_fixes} ; @@{attribute "where"} @{syntax named_insts} @{syntax for_fixes} \ \<^descr> @{command "print_rules"} prints rules declared via attributes @{attribute (Pure) intro}, @{attribute (Pure) elim}, @{attribute (Pure) dest} of Isabelle/Pure. See also the analogous @{command "print_claset"} command for similar rule declarations of the classical reasoner (\secref{sec:classical}). \<^descr> ``@{method "-"}'' (minus) inserts the forward chaining facts as premises into the goal, and nothing else. Note that command @{command_ref "proof"} without any method actually performs a single reduction step using the @{method_ref (Pure) rule} method; thus a plain \<^emph>\do-nothing\ proof step would be ``@{command "proof"}~\-\'' rather than @{command "proof"} alone. \<^descr> @{method "goal_cases"}~\a\<^sub>1 \ a\<^sub>n\ turns the current subgoals into cases within the context (see also \secref{sec:cases-induct}). The specified case names are used if present; otherwise cases are numbered starting from 1. Invoking cases in the subsequent proof body via the @{command_ref case} command will @{command fix} goal parameters, @{command assume} goal premises, and @{command let} variable @{variable_ref ?case} refer to the conclusion. \<^descr> @{method "subproofs"}~\m\ applies the method expression \m\ consecutively to each subgoal, constructing individual subproofs internally (analogous to ``\<^theory_text>\show goal by m\'' for each subgoal of the proof state). Search alternatives of \m\ are truncated: the method is forced to be deterministic. This method combinator impacts the internal construction of proof terms: it makes a cascade of let-expressions within the derivation tree and may thus improve scalability. \<^descr> @{method "fact"}~\a\<^sub>1 \ a\<^sub>n\ composes some fact from \a\<^sub>1, \, a\<^sub>n\ (or implicitly from the current proof context) modulo unification of schematic type and term variables. The rule structure is not taken into account, i.e.\ meta-level implication is considered atomic. This is the same principle underlying literal facts (cf.\ \secref{sec:syn-att}): ``@{command "have"}~\\\~@{command "by"}~\fact\'' is equivalent to ``@{command "note"}~\<^verbatim>\`\\\\\<^verbatim>\`\'' provided that \\ \\ is an instance of some known \\ \\ in the proof context. \<^descr> @{method assumption} solves some goal by a single assumption step. All given facts are guaranteed to participate in the refinement; this means there may be only 0 or 1 in the first place. Recall that @{command "qed"} (\secref{sec:proof-steps}) already concludes any remaining sub-goals by assumption, so structured proofs usually need not quote the @{method assumption} method at all. \<^descr> @{method this} applies all of the current facts directly as rules. Recall that ``@{command "."}'' (dot) abbreviates ``@{command "by"}~\this\''. \<^descr> @{method (Pure) rule}~\a\<^sub>1 \ a\<^sub>n\ applies some rule given as argument in backward manner; facts are used to reduce the rule before applying it to the goal. Thus @{method (Pure) rule} without facts is plain introduction, while with facts it becomes elimination. When no arguments are given, the @{method (Pure) rule} method tries to pick appropriate rules automatically, as declared in the current context using the @{attribute (Pure) intro}, @{attribute (Pure) elim}, @{attribute (Pure) dest} attributes (see below). This is included in the standard behaviour of @{command "proof"} and ``@{command ".."}'' (double-dot) steps (see \secref{sec:proof-steps}). \<^descr> @{attribute (Pure) intro}, @{attribute (Pure) elim}, and @{attribute (Pure) dest} declare introduction, elimination, and destruct rules, to be used with method @{method (Pure) rule}, and similar tools. Note that the latter will ignore rules declared with ``\?\'', while ``\!\'' are used most aggressively. The classical reasoner (see \secref{sec:classical}) introduces its own variants of these attributes; use qualified names to access the present versions of Isabelle/Pure, i.e.\ @{attribute (Pure) "Pure.intro"}. \<^descr> @{attribute (Pure) rule}~\del\ undeclares introduction, elimination, or destruct rules. \<^descr> @{attribute OF}~\a\<^sub>1 \ a\<^sub>n\ applies some theorem to all of the given rules \a\<^sub>1, \, a\<^sub>n\ in canonical right-to-left order, which means that premises stemming from the \a\<^sub>i\ emerge in parallel in the result, without interfering with each other. In many practical situations, the \a\<^sub>i\ do not have premises themselves, so \rule [OF a\<^sub>1 \ a\<^sub>n]\ can be actually read as functional application (modulo unification). Argument positions may be effectively skipped by using ``\_\'' (underscore), which refers to the propositional identity rule in the Pure theory. \<^descr> @{attribute of}~\t\<^sub>1 \ t\<^sub>n\ performs positional instantiation of term variables. The terms \t\<^sub>1, \, t\<^sub>n\ are substituted for any schematic variables occurring in a theorem from left to right; ``\_\'' (underscore) indicates to skip a position. Arguments following a ``\concl:\'' specification refer to positions of the conclusion of a rule. An optional context of local variables \\ x\<^sub>1 \ x\<^sub>m\ may be specified: the instantiated theorem is exported, and these variables become schematic (usually with some shifting of indices). \<^descr> @{attribute "where"}~\x\<^sub>1 = t\<^sub>1 \ \ x\<^sub>n = t\<^sub>n\ performs named instantiation of schematic type and term variables occurring in a theorem. Schematic variables have to be specified on the left-hand side (e.g.\ \?x1.3\). The question mark may be omitted if the variable name is a plain identifier without index. As type instantiations are inferred from term instantiations, explicit type instantiations are seldom necessary. An optional context of local variables \\ x\<^sub>1 \ x\<^sub>m\ may be specified as for @{attribute "of"} above. \ subsection \Defining proof methods\ text \ \begin{matharray}{rcl} @{command_def "method_setup"} & : & \local_theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command method_setup} @{syntax name} '=' @{syntax text} @{syntax text}? \ \<^descr> @{command "method_setup"}~\name = text description\ defines a proof method in the current context. The given \text\ has to be an ML expression of type \<^ML_type>\(Proof.context -> Proof.method) context_parser\, cf.\ basic parsers defined in structure \<^ML_structure>\Args\ and \<^ML_structure>\Attrib\. There are also combinators like \<^ML>\METHOD\ and \<^ML>\SIMPLE_METHOD\ to turn certain tactic forms into official proof methods; the primed versions refer to tactics with explicit goal addressing. Here are some example method definitions: \ (*<*)experiment begin(*>*) method_setup my_method1 = \Scan.succeed (K (SIMPLE_METHOD' (fn i: int => no_tac)))\ "my first method (without any arguments)" method_setup my_method2 = \Scan.succeed (fn ctxt: Proof.context => SIMPLE_METHOD' (fn i: int => no_tac))\ "my second method (with context)" method_setup my_method3 = \Attrib.thms >> (fn thms: thm list => fn ctxt: Proof.context => SIMPLE_METHOD' (fn i: int => no_tac))\ "my third method (with theorem arguments and context)" (*<*)end(*>*) section \Proof by cases and induction \label{sec:cases-induct}\ subsection \Rule contexts\ text \ \begin{matharray}{rcl} @{command_def "case"} & : & \proof(state) \ proof(state)\ \\ @{command_def "print_cases"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def case_names} & : & \attribute\ \\ @{attribute_def case_conclusion} & : & \attribute\ \\ @{attribute_def params} & : & \attribute\ \\ @{attribute_def consumes} & : & \attribute\ \\ \end{matharray} The puristic way to build up Isar proof contexts is by explicit language elements like @{command "fix"}, @{command "assume"}, @{command "let"} (see \secref{sec:proof-context}). This is adequate for plain natural deduction, but easily becomes unwieldy in concrete verification tasks, which typically involve big induction rules with several cases. The @{command "case"} command provides a shorthand to refer to a local context symbolically: certain proof methods provide an environment of named ``cases'' of the form \c: x\<^sub>1, \, x\<^sub>m, \\<^sub>1, \, \\<^sub>n\; the effect of ``@{command "case"}~\c\'' is then equivalent to ``@{command "fix"}~\x\<^sub>1 \ x\<^sub>m\~@{command "assume"}~\c: \\<^sub>1 \ \\<^sub>n\''. Term bindings may be covered as well, notably @{variable ?case} for the main conclusion. By default, the ``terminology'' \x\<^sub>1, \, x\<^sub>m\ of a case value is marked as hidden, i.e.\ there is no way to refer to such parameters in the subsequent proof text. After all, original rule parameters stem from somewhere outside of the current proof text. By using the explicit form ``@{command "case"}~\(c y\<^sub>1 \ y\<^sub>m)\'' instead, the proof author is able to chose local names that fit nicely into the current context. \<^medskip> It is important to note that proper use of @{command "case"} does not provide means to peek at the current goal state, which is not directly observable in Isar! Nonetheless, goal refinement commands do provide named cases \goal\<^sub>i\ for each subgoal \i = 1, \, n\ of the resulting goal state. Using this extra feature requires great care, because some bits of the internal tactical machinery intrude the proof text. In particular, parameter names stemming from the left-over of automated reasoning tools are usually quite unpredictable. Under normal circumstances, the text of cases emerge from standard elimination or induction rules, which in turn are derived from previous theory specifications in a canonical way (say from @{command "inductive"} definitions). \<^medskip> Proper cases are only available if both the proof method and the rules involved support this. By using appropriate attributes, case names, conclusions, and parameters may be also declared by hand. Thus variant versions of rules that have been derived manually become ready to use in advanced case analysis later. \<^rail>\ @@{command case} @{syntax thmdecl}? (name | '(' name (('_' | @{syntax name}) *) ')') ; @@{attribute case_names} ((@{syntax name} ( '[' (('_' | @{syntax name}) *) ']' ) ? ) +) ; @@{attribute case_conclusion} @{syntax name} (@{syntax name} * ) ; @@{attribute params} ((@{syntax name} * ) + @'and') ; @@{attribute consumes} @{syntax int}? \ \<^descr> @{command "case"}~\a: (c x\<^sub>1 \ x\<^sub>m)\ invokes a named local context \c: x\<^sub>1, \, x\<^sub>m, \\<^sub>1, \, \\<^sub>m\, as provided by an appropriate proof method (such as @{method_ref cases} and @{method_ref induct}). The command ``@{command "case"}~\a: (c x\<^sub>1 \ x\<^sub>m)\'' abbreviates ``@{command "fix"}~\x\<^sub>1 \ x\<^sub>m\~@{command "assume"}~\a.c: \\<^sub>1 \ \\<^sub>n\''. Each local fact is qualified by the prefix \a\, and all such facts are collectively bound to the name \a\. The fact name is specification \a\ is optional, the default is to re-use \c\. So @{command "case"}~\(c x\<^sub>1 \ x\<^sub>m)\ is the same as @{command "case"}~\c: (c x\<^sub>1 \ x\<^sub>m)\. \<^descr> @{command "print_cases"} prints all local contexts of the current state, using Isar proof language notation. \<^descr> @{attribute case_names}~\c\<^sub>1 \ c\<^sub>k\ declares names for the local contexts of premises of a theorem; \c\<^sub>1, \, c\<^sub>k\ refers to the \<^emph>\prefix\ of the list of premises. Each of the cases \c\<^sub>i\ can be of the form \c[h\<^sub>1 \ h\<^sub>n]\ where the \h\<^sub>1 \ h\<^sub>n\ are the names of the hypotheses in case \c\<^sub>i\ from left to right. \<^descr> @{attribute case_conclusion}~\c d\<^sub>1 \ d\<^sub>k\ declares names for the conclusions of a named premise \c\; here \d\<^sub>1, \, d\<^sub>k\ refers to the prefix of arguments of a logical formula built by nesting a binary connective (e.g.\ \\\). Note that proof methods such as @{method induct} and @{method coinduct} already provide a default name for the conclusion as a whole. The need to name subformulas only arises with cases that split into several sub-cases, as in common co-induction rules. \<^descr> @{attribute params}~\p\<^sub>1 \ p\<^sub>m \ \ q\<^sub>1 \ q\<^sub>n\ renames the innermost parameters of premises \1, \, n\ of some theorem. An empty list of names may be given to skip positions, leaving the present parameters unchanged. Note that the default usage of case rules does \<^emph>\not\ directly expose parameters to the proof context. \<^descr> @{attribute consumes}~\n\ declares the number of ``major premises'' of a rule, i.e.\ the number of facts to be consumed when it is applied by an appropriate proof method. The default value of @{attribute consumes} is \n = 1\, which is appropriate for the usual kind of cases and induction rules for inductive sets (cf.\ \secref{sec:hol-inductive}). Rules without any @{attribute consumes} declaration given are treated as if @{attribute consumes}~\0\ had been specified. A negative \n\ is interpreted relatively to the total number of premises of the rule in the target context. Thus its absolute value specifies the remaining number of premises, after subtracting the prefix of major premises as indicated above. This form of declaration has the technical advantage of being stable under more morphisms, notably those that export the result from a nested @{command_ref context} with additional assumptions. Note that explicit @{attribute consumes} declarations are only rarely needed; this is already taken care of automatically by the higher-level @{attribute cases}, @{attribute induct}, and @{attribute coinduct} declarations. \ subsection \Proof methods\ text \ \begin{matharray}{rcl} @{method_def cases} & : & \method\ \\ @{method_def induct} & : & \method\ \\ @{method_def induction} & : & \method\ \\ @{method_def coinduct} & : & \method\ \\ \end{matharray} The @{method cases}, @{method induct}, @{method induction}, and @{method coinduct} methods provide a uniform interface to common proof techniques over datatypes, inductive predicates (or sets), recursive functions etc. The corresponding rules may be specified and instantiated in a casual manner. Furthermore, these methods provide named local contexts that may be invoked via the @{command "case"} proof command within the subsequent proof text. This accommodates compact proof texts even when reasoning about large specifications. The @{method induct} method also provides some infrastructure to work with structured statements (either using explicit meta-level connectives, or including facts and parameters separately). This avoids cumbersome encoding of ``strengthened'' inductive statements within the object-logic. Method @{method induction} differs from @{method induct} only in the names of the facts in the local context invoked by the @{command "case"} command. \<^rail>\ @@{method cases} ('(' 'no_simp' ')')? \ (@{syntax insts} * @'and') rule? ; (@@{method induct} | @@{method induction}) ('(' 'no_simp' ')')? (definsts * @'and') \ arbitrary? taking? rule? ; @@{method coinduct} @{syntax insts} taking rule? ; rule: ('type' | 'pred' | 'set') ':' (@{syntax name} +) | 'rule' ':' (@{syntax thm} +) ; definst: @{syntax name} ('==' | '\') @{syntax term} | '(' @{syntax term} ')' | @{syntax inst} ; definsts: ( definst * ) ; arbitrary: 'arbitrary' ':' ((@{syntax term} * ) @'and' +) ; taking: 'taking' ':' @{syntax insts} \ \<^descr> @{method cases}~\insts R\ applies method @{method rule} with an appropriate case distinction theorem, instantiated to the subjects \insts\. Symbolic case names are bound according to the rule's local contexts. The rule is determined as follows, according to the facts and arguments passed to the @{method cases} method: \<^medskip> \begin{tabular}{llll} facts & & arguments & rule \\\hline \\ R\ & @{method cases} & & implicit rule \R\ \\ & @{method cases} & & classical case split \\ & @{method cases} & \t\ & datatype exhaustion (type of \t\) \\ \\ A t\ & @{method cases} & \\\ & inductive predicate/set elimination (of \A\) \\ \\\ & @{method cases} & \\ rule: R\ & explicit rule \R\ \\ \end{tabular} \<^medskip> Several instantiations may be given, referring to the \<^emph>\suffix\ of premises of the case rule; within each premise, the \<^emph>\prefix\ of variables is instantiated. In most situations, only a single term needs to be specified; this refers to the first variable of the last premise (it is usually the same for all cases). The \(no_simp)\ option can be used to disable pre-simplification of cases (see the description of @{method induct} below for details). \<^descr> @{method induct}~\insts R\ and @{method induction}~\insts R\ are analogous to the @{method cases} method, but refer to induction rules, which are determined as follows: \<^medskip> \begin{tabular}{llll} facts & & arguments & rule \\\hline & @{method induct} & \P x\ & datatype induction (type of \x\) \\ \\ A x\ & @{method induct} & \\\ & predicate/set induction (of \A\) \\ \\\ & @{method induct} & \\ rule: R\ & explicit rule \R\ \\ \end{tabular} \<^medskip> Several instantiations may be given, each referring to some part of a mutual inductive definition or datatype --- only related partial induction rules may be used together, though. Any of the lists of terms \P, x, \\ refers to the \<^emph>\suffix\ of variables present in the induction rule. This enables the writer to specify only induction variables, or both predicates and variables, for example. Instantiations may be definitional: equations \x \ t\ introduce local definitions, which are inserted into the claim and discharged after applying the induction rule. Equalities reappear in the inductive cases, but have been transformed according to the induction principle being involved here. In order to achieve practically useful induction hypotheses, some variables occurring in \t\ need to generalized (see below). Instantiations of the form \t\, where \t\ is not a variable, are taken as a shorthand for \x \ t\, where \x\ is a fresh variable. If this is not intended, \t\ has to be enclosed in parentheses. By default, the equalities generated by definitional instantiations are pre-simplified using a specific set of rules, usually consisting of distinctness and injectivity theorems for datatypes. This pre-simplification may cause some of the parameters of an inductive case to disappear, or may even completely delete some of the inductive cases, if one of the equalities occurring in their premises can be simplified to \False\. The \(no_simp)\ option can be used to disable pre-simplification. Additional rules to be used in pre-simplification can be declared using the @{attribute_def induct_simp} attribute. The optional ``\arbitrary: x\<^sub>1 \ x\<^sub>m\'' specification generalizes variables \x\<^sub>1, \, x\<^sub>m\ of the original goal before applying induction. It is possible to separate variables by ``\and\'' to generalize in goals other than the first. Thus induction hypotheses may become sufficiently general to get the proof through. Together with definitional instantiations, one may effectively perform induction over expressions of a certain structure. The optional ``\taking: t\<^sub>1 \ t\<^sub>n\'' specification provides additional instantiations of a prefix of pending variables in the rule. Such schematic induction rules rarely occur in practice, though. \<^descr> @{method coinduct}~\inst R\ is analogous to the @{method induct} method, but refers to coinduction rules, which are determined as follows: \<^medskip> \begin{tabular}{llll} goal & & arguments & rule \\\hline & @{method coinduct} & \x\ & type coinduction (type of \x\) \\ \A x\ & @{method coinduct} & \\\ & predicate/set coinduction (of \A\) \\ \\\ & @{method coinduct} & \\ rule: R\ & explicit rule \R\ \\ \end{tabular} \<^medskip> Coinduction is the dual of induction. Induction essentially eliminates \A x\ towards a generic result \P x\, while coinduction introduces \A x\ starting with \B x\, for a suitable ``bisimulation'' \B\. The cases of a coinduct rule are typically named after the predicates or sets being covered, while the conclusions consist of several alternatives being named after the individual destructor patterns. The given instantiation refers to the \<^emph>\suffix\ of variables occurring in the rule's major premise, or conclusion if unavailable. An additional ``\taking: t\<^sub>1 \ t\<^sub>n\'' specification may be required in order to specify the bisimulation to be used in the coinduction step. Above methods produce named local contexts, as determined by the instantiated rule as given in the text. Beyond that, the @{method induct} and @{method coinduct} methods guess further instantiations from the goal specification itself. Any persisting unresolved schematic variables of the resulting rule will render the the corresponding case invalid. The term binding @{variable ?case} for the conclusion will be provided with each case, provided that term is fully specified. The @{command "print_cases"} command prints all named cases present in the current proof state. \<^medskip> Despite the additional infrastructure, both @{method cases} and @{method coinduct} merely apply a certain rule, after instantiation, while conforming due to the usual way of monotonic natural deduction: the context of a structured statement \\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ \\ reappears unchanged after the case split. The @{method induct} method is fundamentally different in this respect: the meta-level structure is passed through the ``recursive'' course involved in the induction. Thus the original statement is basically replaced by separate copies, corresponding to the induction hypotheses and conclusion; the original goal context is no longer available. Thus local assumptions, fixed parameters and definitions effectively participate in the inductive rephrasing of the original statement. In @{method induct} proofs, local assumptions introduced by cases are split into two different kinds: \hyps\ stemming from the rule and \prems\ from the goal statement. This is reflected in the extracted cases accordingly, so invoking ``@{command "case"}~\c\'' will provide separate facts \c.hyps\ and \c.prems\, as well as fact \c\ to hold the all-inclusive list. In @{method induction} proofs, local assumptions introduced by cases are split into three different kinds: \IH\, the induction hypotheses, \hyps\, the remaining hypotheses stemming from the rule, and \prems\, the assumptions from the goal statement. The names are \c.IH\, \c.hyps\ and \c.prems\, as above. \<^medskip> Facts presented to either method are consumed according to the number of ``major premises'' of the rule involved, which is usually 0 for plain cases and induction rules of datatypes etc.\ and 1 for rules of inductive predicates or sets and the like. The remaining facts are inserted into the goal verbatim before the actual \cases\, \induct\, or \coinduct\ rule is applied. \ subsection \Declaring rules\ text \ \begin{matharray}{rcl} @{command_def "print_induct_rules"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def cases} & : & \attribute\ \\ @{attribute_def induct} & : & \attribute\ \\ @{attribute_def coinduct} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute cases} spec ; @@{attribute induct} spec ; @@{attribute coinduct} spec ; spec: (('type' | 'pred' | 'set') ':' @{syntax name}) | 'del' \ \<^descr> @{command "print_induct_rules"} prints cases and induct rules for predicates (or sets) and types of the current context. \<^descr> @{attribute cases}, @{attribute induct}, and @{attribute coinduct} (as attributes) declare rules for reasoning about (co)inductive predicates (or sets) and types, using the corresponding methods of the same name. Certain definitional packages of object-logics usually declare emerging cases and induction rules as expected, so users rarely need to intervene. Rules may be deleted via the \del\ specification, which covers all of the \type\/\pred\/\set\ sub-categories simultaneously. For example, @{attribute cases}~\del\ removes any @{attribute cases} rules declared for some type, predicate, or set. Manual rule declarations usually refer to the @{attribute case_names} and @{attribute params} attributes to adjust names of cases and parameters of a rule; the @{attribute consumes} declaration is taken care of automatically: @{attribute consumes}~\0\ is specified for ``type'' rules and @{attribute consumes}~\1\ for ``predicate'' / ``set'' rules. \ section \Generalized elimination and case splitting \label{sec:obtain}\ text \ \begin{matharray}{rcl} @{command_def "consider"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "obtain"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ \end{matharray} Generalized elimination means that hypothetical parameters and premises may be introduced in the current context, potentially with a split into cases. This works by virtue of a locally proven rule that establishes the soundness of this temporary context extension. As representative examples, one may think of standard rules from Isabelle/HOL like this: \<^medskip> \begin{tabular}{ll} \\x. B x \ (\x. B x \ thesis) \ thesis\ \\ \A \ B \ (A \ B \ thesis) \ thesis\ \\ \A \ B \ (A \ thesis) \ (B \ thesis) \ thesis\ \\ \end{tabular} \<^medskip> In general, these particular rules and connectives need to get involved at all: this concept works directly in Isabelle/Pure via Isar commands defined below. In particular, the logic of elimination and case splitting is delegated to an Isar proof, which often involves automated tools. \<^rail>\ @@{command consider} @{syntax obtain_clauses} ; @@{command obtain} @{syntax par_name}? @{syntax vars} \ @'where' concl prems @{syntax for_fixes} ; concl: (@{syntax props} + @'and') ; prems: (@'if' (@{syntax props'} + @'and'))? \ \<^descr> @{command consider}~\(a) \<^vec>x \ \<^vec>A \<^vec>x | (b) \<^vec>y \ \<^vec>B \<^vec>y | \\ states a rule for case splitting into separate subgoals, such that each case involves new parameters and premises. After the proof is finished, the resulting rule may be used directly with the @{method cases} proof method (\secref{sec:cases-induct}), in order to perform actual case-splitting of the proof text via @{command case} and @{command next} as usual. Optional names in round parentheses refer to case names: in the proof of the rule this is a fact name, in the resulting rule it is used as annotation with the @{attribute_ref case_names} attribute. \<^medskip> Formally, the command @{command consider} is defined as derived Isar language element as follows: \begin{matharray}{l} @{command "consider"}~\(a) \<^vec>x \ \<^vec>A \<^vec>x | (b) \<^vec>y \ \<^vec>B \<^vec>y | \ \\ \\[1ex] \quad @{command "have"}~\[case_names a b \]: thesis\ \\ \qquad \\ a [Pure.intro?]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis\ \\ \qquad \\ b [Pure.intro?]: \\<^vec>y. \<^vec>B \<^vec>y \ thesis\ \\ \qquad \\ \\ \\ \qquad \\ thesis\ \\ \qquad @{command "apply"}~\(insert a b \)\ \\ \end{matharray} See also \secref{sec:goals} for @{keyword "obtains"} in toplevel goal statements, as well as @{command print_statement} to print existing rules in a similar format. \<^descr> @{command obtain}~\\<^vec>x \ \<^vec>A \<^vec>x\ states a generalized elimination rule with exactly one case. After the proof is finished, it is activated for the subsequent proof text: the context is augmented via @{command fix}~\\<^vec>x\ @{command assume}~\\<^vec>A \<^vec>x\, with special provisions to export later results by discharging these assumptions again. Note that according to the parameter scopes within the elimination rule, results \<^emph>\must not\ refer to hypothetical parameters; otherwise the export will fail! This restriction conforms to the usual manner of existential reasoning in Natural Deduction. \<^medskip> Formally, the command @{command obtain} is defined as derived Isar language element as follows, using an instrumented variant of @{command assume}: \begin{matharray}{l} @{command "obtain"}~\\<^vec>x \ a: \<^vec>A \<^vec>x \proof\ \\ \\[1ex] \quad @{command "have"}~\thesis\ \\ \qquad \\ that [Pure.intro?]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis\ \\ \qquad \\ thesis\ \\ \qquad @{command "apply"}~\(insert that)\ \\ \qquad \\proof\\ \\ \quad @{command "fix"}~\\<^vec>x\~@{command "assume"}\\<^sup>* a: \<^vec>A \<^vec>x\ \\ \end{matharray} In the proof of @{command consider} and @{command obtain} the local premises are always bound to the fact name @{fact_ref that}, according to structured Isar statements involving @{keyword_ref "if"} (\secref{sec:goals}). Facts that are established by @{command "obtain"} cannot be polymorphic: any type-variables occurring here are fixed in the present context. This is a natural consequence of the role of @{command fix} and @{command assume} in this construct. \ end diff --git a/src/Doc/Isar_Ref/Proof_Script.thy b/src/Doc/Isar_Ref/Proof_Script.thy --- a/src/Doc/Isar_Ref/Proof_Script.thy +++ b/src/Doc/Isar_Ref/Proof_Script.thy @@ -1,280 +1,279 @@ (*:maxLineLen=78:*) theory Proof_Script imports Main Base begin chapter \Proof scripts\ text \ Interactive theorem proving is traditionally associated with ``proof scripts'', but Isabelle/Isar is centered around structured \<^emph>\proof documents\ instead (see also \chref{ch:proofs}). Nonetheless, it is possible to emulate proof scripts by sequential refinements of a proof state in backwards mode, notably with the @{command apply} command (see \secref{sec:tactic-commands}). There are also various proof methods that allow to refer to implicit goal state information that is not accessible to structured Isar proofs (see \secref{sec:tactics}). Note that the @{command subgoal} (\secref{sec:subgoal}) command usually eliminates the need for implicit goal state references. \ section \Commands for step-wise refinement \label{sec:tactic-commands}\ text \ \begin{matharray}{rcl} @{command_def "supply"}\\<^sup>*\ & : & \proof(prove) \ proof(prove)\ \\ @{command_def "apply"}\\<^sup>*\ & : & \proof(prove) \ proof(prove)\ \\ @{command_def "apply_end"}\\<^sup>*\ & : & \proof(state) \ proof(state)\ \\ @{command_def "done"}\\<^sup>*\ & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{command_def "defer"}\\<^sup>*\ & : & \proof \ proof\ \\ @{command_def "prefer"}\\<^sup>*\ & : & \proof \ proof\ \\ @{command_def "back"}\\<^sup>*\ & : & \proof \ proof\ \\ \end{matharray} \<^rail>\ @@{command supply} (@{syntax thmdef}? @{syntax thms} + @'and') ; ( @@{command apply} | @@{command apply_end} ) @{syntax method} ; @@{command defer} @{syntax nat}? ; @@{command prefer} @{syntax nat} \ \<^descr> @{command "supply"} supports fact definitions during goal refinement: it is similar to @{command "note"}, but it operates in backwards mode and does not have any impact on chained facts. \<^descr> @{command "apply"}~\m\ applies proof method \m\ in initial position, but unlike @{command "proof"} it retains ``\proof(prove)\'' mode. Thus consecutive method applications may be given just as in tactic scripts. Facts are passed to \m\ as indicated by the goal's forward-chain mode, and are \<^emph>\consumed\ afterwards. Thus any further @{command "apply"} command would always work in a purely backward manner. \<^descr> @{command "apply_end"}~\m\ applies proof method \m\ as if in terminal position. Basically, this simulates a multi-step tactic script for @{command "qed"}, but may be given anywhere within the proof body. No facts are passed to \m\ here. Furthermore, the static context is that of the enclosing goal (as for actual @{command "qed"}). Thus the proof method may not refer to any assumptions introduced in the current body, for example. \<^descr> @{command "done"} completes a proof script, provided that the current goal state is solved completely. Note that actual structured proof commands (e.g.\ ``@{command "."}'' or @{command "sorry"}) may be used to conclude proof scripts as well. \<^descr> @{command "defer"}~\n\ and @{command "prefer"}~\n\ shuffle the list of pending goals: @{command "defer"} puts off sub-goal \n\ to the end of the list (\n = 1\ by default), while @{command "prefer"} brings sub-goal \n\ to the front. \<^descr> @{command "back"} does back-tracking over the result sequence of the latest proof command. Any proof command may return multiple results, and this command explores the possibilities step-by-step. It is mainly useful for experimentation and interactive exploration, and should be avoided in finished proofs. \ section \Explicit subgoal structure \label{sec:subgoal}\ text \ \begin{matharray}{rcl} @{command_def "subgoal"}\\<^sup>*\ & : & \proof \ proof\ \\ \end{matharray} \<^rail>\ @@{command subgoal} @{syntax thmbind}? prems? params? ; prems: @'premises' @{syntax thmbind}? ; params: @'for' '\'? (('_' | @{syntax name})+) \ \<^descr> @{command "subgoal"} allows to impose some structure on backward refinements, to avoid proof scripts degenerating into long of @{command apply} sequences. The current goal state, which is essentially a hidden part of the Isar/VM configuration, is turned into a proof context and remaining conclusion. This corresponds to @{command fix}~/ @{command assume}~/ @{command show} in structured proofs, but the text of the parameters, premises and conclusion is not given explicitly. Goal parameters may be specified separately, in order to allow referring to them in the proof body: ``@{command subgoal}~@{keyword "for"}~\x y z\'' names a \<^emph>\prefix\, and ``@{command subgoal}~@{keyword "for"}~\\ x y z\'' names a \<^emph>\suffix\ of goal parameters. The latter uses a literal \<^verbatim>\\\ symbol as notation. Parameter positions may be skipped via dummies (underscore). Unspecified names remain internal, and thus inaccessible in the proof text. ``@{command subgoal}~@{keyword "premises"}~\prems\'' indicates that goal premises should be turned into assumptions of the context (otherwise the remaining conclusion is a Pure implication). The fact name and attributes are optional; the particular name ``\prems\'' is a common convention for the premises of an arbitrary goal context in proof scripts. ``@{command subgoal}~\result\'' indicates a fact name for the result of a proven subgoal. Thus it may be re-used in further reasoning, similar to the result of @{command show} in structured Isar proofs. Here are some abstract examples: \ lemma "\x y z. A x \ B y \ C z" and "\u v. X u \ Y v" subgoal \ subgoal \ done lemma "\x y z. A x \ B y \ C z" and "\u v. X u \ Y v" subgoal for x y z \ subgoal for u v \ done lemma "\x y z. A x \ B y \ C z" and "\u v. X u \ Y v" subgoal premises for x y z using \A x\ \B y\ \ subgoal premises for u v using \X u\ \ done lemma "\x y z. A x \ B y \ C z" and "\u v. X u \ Y v" subgoal r premises prems for x y z proof - have "A x" by (fact prems) moreover have "B y" by (fact prems) ultimately show ?thesis \ qed subgoal premises prems for u v proof - have "\x y z. A x \ B y \ C z" by (fact r) moreover have "X u" by (fact prems) ultimately show ?thesis \ qed done lemma "\x y z. A x \ B y \ C z" subgoal premises prems for \ z proof - from prems show "C z" \ qed done section \Tactics: improper proof methods \label{sec:tactics}\ text \ The following improper proof methods emulate traditional tactics. These admit direct access to the goal state, which is normally considered harmful! In particular, this may involve both numbered goal addressing (default 1), and dynamic instantiation within the scope of some subgoal. \begin{warn} Dynamic instantiations refer to universally quantified parameters of a subgoal (the dynamic context) rather than fixed variables and term abbreviations of a (static) Isar context. \end{warn} Tactic emulation methods, unlike their ML counterparts, admit simultaneous instantiation from both dynamic and static contexts. If names occur in both contexts goal parameters hide locally fixed variables. Likewise, schematic variables refer to term abbreviations, if present in the static context. Otherwise the schematic variable is interpreted as a schematic variable and left to be solved by unification with certain parts of the subgoal. Note that the tactic emulation proof methods in Isabelle/Isar are consistently named \foo_tac\. Note also that variable names occurring on left hand sides of instantiations must be preceded by a question mark if they coincide with a keyword or contain dots. This is consistent with the attribute @{attribute "where"} (see \secref{sec:pure-meth-att}). \begin{matharray}{rcl} @{method_def rule_tac}\\<^sup>*\ & : & \method\ \\ @{method_def erule_tac}\\<^sup>*\ & : & \method\ \\ @{method_def drule_tac}\\<^sup>*\ & : & \method\ \\ @{method_def frule_tac}\\<^sup>*\ & : & \method\ \\ @{method_def cut_tac}\\<^sup>*\ & : & \method\ \\ @{method_def thin_tac}\\<^sup>*\ & : & \method\ \\ @{method_def subgoal_tac}\\<^sup>*\ & : & \method\ \\ @{method_def rename_tac}\\<^sup>*\ & : & \method\ \\ @{method_def rotate_tac}\\<^sup>*\ & : & \method\ \\ @{method_def tactic}\\<^sup>*\ & : & \method\ \\ @{method_def raw_tactic}\\<^sup>*\ & : & \method\ \\ \end{matharray} \<^rail>\ (@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} | @@{method frule_tac} | @@{method cut_tac}) @{syntax goal_spec}? \ (@{syntax named_insts} @{syntax for_fixes} @'in' @{syntax thm} | @{syntax thms} ) ; @@{method thin_tac} @{syntax goal_spec}? @{syntax prop} @{syntax for_fixes} ; @@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +) @{syntax for_fixes} ; @@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +) ; @@{method rotate_tac} @{syntax goal_spec}? @{syntax int}? ; (@@{method tactic} | @@{method raw_tactic}) @{syntax text} \ \<^descr> @{method rule_tac} etc. do resolution of rules with explicit - instantiation. This works the same way as the ML tactics \<^ML>\Rule_Insts.res_inst_tac\ etc.\ (see @{cite "isabelle-implementation"}). + instantiation. This works the same way as the ML tactics \<^ML>\Rule_Insts.res_inst_tac\ etc.\ (see \<^cite>\"isabelle-implementation"\). Multiple rules may be only given if there is no instantiation; then @{method - rule_tac} is the same as \<^ML>\resolve_tac\ in ML (see @{cite - "isabelle-implementation"}). + rule_tac} is the same as \<^ML>\resolve_tac\ in ML (see \<^cite>\"isabelle-implementation"\). \<^descr> @{method cut_tac} inserts facts into the proof state as assumption of a subgoal; instantiations may be given as well. Note that the scope of schematic variables is spread over the main goal statement and rule premises are turned into new subgoals. This is in contrast to the regular method @{method insert} which inserts closed rule statements. \<^descr> @{method thin_tac}~\\\ deletes the specified premise from a subgoal. Note that \\\ may contain schematic variables, to abbreviate the intended proposition; the first matching subgoal premise will be deleted. Removing useless premises from a subgoal increases its readability and can make search tactics run faster. \<^descr> @{method subgoal_tac}~\\\<^sub>1 \ \\<^sub>n\ adds the propositions \\\<^sub>1 \ \\<^sub>n\ as local premises to a subgoal, and poses the same as new subgoals (in the original context). \<^descr> @{method rename_tac}~\x\<^sub>1 \ x\<^sub>n\ renames parameters of a goal according to the list \x\<^sub>1, \, x\<^sub>n\, which refers to the \<^emph>\suffix\ of variables. \<^descr> @{method rotate_tac}~\n\ rotates the premises of a subgoal by \n\ positions: from right to left if \n\ is positive, and from left to right if \n\ is negative; the default value is 1. \<^descr> @{method tactic}~\text\ produces a proof method from any ML text of type \<^ML_type>\tactic\. Apart from the usual ML environment and the current proof context, the ML code may refer to the locally bound values \<^ML_text>\facts\, which indicates any current facts used for forward-chaining. \<^descr> @{method raw_tactic} is similar to @{method tactic}, but presents the goal state in its raw internal form, where simultaneous subgoals appear as conjunction of the logical framework instead of the usual split into several subgoals. While feature this is useful for debugging of complex method definitions, it should not never appear in production theories. \ end diff --git a/src/Doc/Isar_Ref/Spec.thy b/src/Doc/Isar_Ref/Spec.thy --- a/src/Doc/Isar_Ref/Spec.thy +++ b/src/Doc/Isar_Ref/Spec.thy @@ -1,1549 +1,1548 @@ (*:maxLineLen=78:*) theory Spec imports Main Base begin chapter \Specifications\ text \ The Isabelle/Isar theory format integrates specifications and proofs, with support for interactive development by continuous document editing. There is a separate document preparation system (see \chref{ch:document-prep}), for typesetting formal developments together with informal text. The resulting hyper-linked PDF documents can be used both for WWW presentation and printed copies. The Isar proof language (see \chref{ch:proofs}) is embedded into the theory language as a proper sub-language. Proof mode is entered by stating some \<^theory_text>\theorem\ or \<^theory_text>\lemma\ at the theory level, and left again with the final conclusion (e.g.\ via \<^theory_text>\qed\). \ section \Defining theories \label{sec:begin-thy}\ text \ \begin{matharray}{rcl} @{command_def "theory"} & : & \toplevel \ theory\ \\ @{command_def (global) "end"} & : & \theory \ toplevel\ \\ @{command_def "thy_deps"}\\<^sup>*\ & : & \theory \\ \\ \end{matharray} Isabelle/Isar theories are defined via theory files, which consist of an outermost sequence of definition--statement--proof elements. Some definitions are self-sufficient (e.g.\ \<^theory_text>\fun\ in Isabelle/HOL), with foundational proofs performed internally. Other definitions require an explicit proof as justification (e.g.\ \<^theory_text>\function\ and \<^theory_text>\termination\ in Isabelle/HOL). Plain statements like \<^theory_text>\theorem\ or \<^theory_text>\lemma\ are merely a special case of that, defining a theorem from a given proposition and its proof. The theory body may be sub-structured by means of \<^emph>\local theory targets\, such as \<^theory_text>\locale\ and \<^theory_text>\class\. It is also possible to use \<^theory_text>\context begin \ end\ blocks to delimited a local theory context: a \<^emph>\named context\ to augment a locale or class specification, or an \<^emph>\unnamed context\ to refer to local parameters and assumptions that are discharged later. See \secref{sec:target} for more details. \<^medskip> A theory is commenced by the \<^theory_text>\theory\ command, which indicates imports of previous theories, according to an acyclic foundational order. Before the initial \<^theory_text>\theory\ command, there may be optional document header material (like \<^theory_text>\section\ or \<^theory_text>\text\, see \secref{sec:markup}). The document header is outside of the formal theory context, though. A theory is concluded by a final @{command (global) "end"} command, one that does not belong to a local theory target. No further commands may follow such a global @{command (global) "end"}. \<^rail>\ @@{command theory} @{syntax system_name} @'imports' (@{syntax system_name} +) \ keywords? abbrevs? @'begin' ; keywords: @'keywords' (keyword_decls + @'and') ; keyword_decls: (@{syntax string} +) ('::' @{syntax name} @{syntax tags})? ; abbrevs: @'abbrevs' (((text+) '=' (text+)) + @'and') ; @@{command thy_deps} (thy_bounds thy_bounds?)? ; thy_bounds: @{syntax name} | '(' (@{syntax name} + @'|') ')' \ \<^descr> \<^theory_text>\theory A imports B\<^sub>1 \ B\<^sub>n begin\ starts a new theory \A\ based on the merge of existing theories \B\<^sub>1 \ B\<^sub>n\. Due to the possibility to import more than one ancestor, the resulting theory structure of an Isabelle session forms a directed acyclic graph (DAG). Isabelle takes care that sources contributing to the development graph are always up-to-date: changed files are automatically rechecked whenever a theory header specification is processed. Empty imports are only allowed in the bootstrap process of the special theory \<^theory>\Pure\, which is the start of any other formal development based on Isabelle. Regular user theories usually refer to some more complex entry point, such as theory \<^theory>\Main\ in Isabelle/HOL. The @{keyword_def "keywords"} specification declares outer syntax (\chref{ch:outer-syntax}) that is introduced in this theory later on (rare in end-user applications). Both minor keywords and major keywords of the Isar command language need to be specified, in order to make parsing of proof documents work properly. Command keywords need to be classified according to their structural role in the formal text. Examples may be seen in Isabelle/HOL sources itself, such as @{keyword "keywords"}~\<^verbatim>\"typedef"\ \:: thy_goal_defn\ or @{keyword "keywords"}~\<^verbatim>\"datatype"\ \:: thy_defn\ for theory-level definitions with and without proof, respectively. Additional @{syntax tags} provide defaults for document preparation (\secref{sec:document-markers}). The @{keyword_def "abbrevs"} specification declares additional abbreviations for syntactic completion. The default for a new keyword is just its name, but completion may be avoided by defining @{keyword "abbrevs"} with empty text. \<^descr> @{command (global) "end"} concludes the current theory definition. Note that some other commands, e.g.\ local theory targets \<^theory_text>\locale\ or \<^theory_text>\class\ may involve a \<^theory_text>\begin\ that needs to be matched by @{command (local) "end"}, according to the usual rules for nested blocks. \<^descr> \<^theory_text>\thy_deps\ visualizes the theory hierarchy as a directed acyclic graph. By default, all imported theories are shown. This may be restricted by specifying bounds wrt. the theory inclusion relation. \ section \Local theory targets \label{sec:target}\ text \ \begin{matharray}{rcll} @{command_def "context"} & : & \theory \ local_theory\ \\ @{command_def (local) "end"} & : & \local_theory \ theory\ \\ @{keyword_def "private"} \\ @{keyword_def "qualified"} \\ \end{matharray} A local theory target is a specification context that is managed separately within the enclosing theory. Contexts may introduce parameters (fixed variables) and assumptions (hypotheses). Definitions and theorems depending on the context may be added incrementally later on. \<^emph>\Named contexts\ refer to locales (cf.\ \secref{sec:locale}) or type classes (cf.\ \secref{sec:class}); the name ``\-\'' signifies the global theory context. \<^emph>\Unnamed contexts\ may introduce additional parameters and assumptions, and results produced in the context are generalized accordingly. Such auxiliary contexts may be nested within other targets, like \<^theory_text>\locale\, \<^theory_text>\class\, \<^theory_text>\instantiation\, \<^theory_text>\overloading\. \<^rail>\ @@{command context} @{syntax name} @{syntax_ref "opening"}? @'begin' ; @@{command context} @{syntax_ref "includes"}? (@{syntax context_elem} * ) @'begin' ; @{syntax_def target}: '(' @'in' @{syntax name} ')' \ \<^descr> \<^theory_text>\context c bundles begin\ opens a named context, by recommencing an existing locale or class \c\. Note that locale and class definitions allow to include the \<^theory_text>\begin\ keyword as well, in order to continue the local theory immediately after the initial specification. Optionally given \bundles\ only take effect in the surface context within the \<^theory_text>\begin\ / \<^theory_text>\end\ block. \<^descr> \<^theory_text>\context bundles elements begin\ opens an unnamed context, by extending the enclosing global or local theory target by the given declaration bundles (\secref{sec:bundle}) and context elements (\<^theory_text>\fixes\, \<^theory_text>\assumes\ etc.). This means any results stemming from definitions and proofs in the extended context will be exported into the enclosing target by lifting over extra parameters and premises. \<^descr> @{command (local) "end"} concludes the current local theory, according to the nesting of contexts. Note that a global @{command (global) "end"} has a different meaning: it concludes the theory itself (\secref{sec:begin-thy}). \<^descr> \<^theory_text>\private\ or \<^theory_text>\qualified\ may be given as modifiers before any local theory command. This restricts name space accesses to the local scope, as determined by the enclosing \<^theory_text>\context begin \ end\ block. Outside its scope, a \<^theory_text>\private\ name is inaccessible, and a \<^theory_text>\qualified\ name is only accessible with some qualification. Neither a global \<^theory_text>\theory\ nor a \<^theory_text>\locale\ target provides a local scope by itself: an extra unnamed context is required to use \<^theory_text>\private\ or \<^theory_text>\qualified\ here. \<^descr> \(\@{keyword_def "in"}~\c)\ given after any local theory command specifies an immediate target, e.g.\ ``\<^theory_text>\definition (in c)\'' or ``\<^theory_text>\theorem (in c)\''. This works both in a local or global theory context; the current target context will be suspended for this command only. Note that ``\<^theory_text>\(in -)\'' will always produce a global result independently of the current target context. Any specification element that operates on \local_theory\ according to this manual implicitly allows the above target syntax \<^theory_text>\(in c)\, but individual syntax diagrams omit that aspect for clarity. \<^medskip> The exact meaning of results produced within a local theory context depends on the underlying target infrastructure (locale, type class etc.). The general idea is as follows, considering a context named \c\ with parameter \x\ and assumption \A[x]\. Definitions are exported by introducing a global version with additional arguments; a syntactic abbreviation links the long form with the abstract version of the target context. For example, \a \ t[x]\ becomes \c.a ?x \ t[?x]\ at the theory level (for arbitrary \?x\), together with a local abbreviation \a \ c.a x\ in the target context (for the fixed parameter \x\). Theorems are exported by discharging the assumptions and generalizing the parameters of the context. For example, \a: B[x]\ becomes \c.a: A[?x] \ B[?x]\, again for arbitrary \?x\. \ section \Bundled declarations \label{sec:bundle}\ text \ \begin{matharray}{rcl} @{command_def "bundle"} & : & \local_theory \ local_theory\ \\ @{command "bundle"} & : & \theory \ local_theory\ \\ @{command_def "print_bundles"}\\<^sup>*\ & : & \context \\ \\ @{command_def "include"} & : & \proof(state) \ proof(state)\ \\ @{command_def "including"} & : & \proof(prove) \ proof(prove)\ \\ @{keyword_def "includes"} & : & syntax \\ \end{matharray} The outer syntax of fact expressions (\secref{sec:syn-att}) involves theorems and attributes, which are evaluated in the context and applied to it. Attributes may declare theorems to the context, as in \this_rule [intro] that_rule [elim]\ for example. Configuration options (\secref{sec:config}) are special declaration attributes that operate on the context without a theorem, as in \[[show_types = false]]\ for example. Expressions of this form may be defined as \<^emph>\bundled declarations\ in the context, and included in other situations later on. Including declaration bundles augments a local context casually without logical dependencies, which is in contrast to locales and locale interpretation (\secref{sec:locale}). \<^rail>\ @@{command bundle} @{syntax name} ( '=' @{syntax thms} @{syntax for_fixes} | @'begin') ; @@{command print_bundles} ('!'?) ; (@@{command include} | @@{command including}) (@{syntax name}+) ; @{syntax_def "includes"}: @'includes' (@{syntax name}+) ; @{syntax_def "opening"}: @'opening' (@{syntax name}+) ; @@{command unbundle} (@{syntax name}+) \ \<^descr> \<^theory_text>\bundle b = decls\ defines a bundle of declarations in the current context. The RHS is similar to the one of the \<^theory_text>\declare\ command. Bundles defined in local theory targets are subject to transformations via morphisms, when moved into different application contexts; this works analogously to any other local theory specification. \<^descr> \<^theory_text>\bundle b begin body end\ defines a bundle of declarations from the \body\ of local theory specifications. It may consist of commands that are technically equivalent to \<^theory_text>\declare\ or \<^theory_text>\declaration\, which also includes \<^theory_text>\notation\, for example. Named fact declarations like ``\<^theory_text>\lemmas a [simp] = b\'' or ``\<^theory_text>\lemma a [simp]: B \\'' are also admitted, but the name bindings are not recorded in the bundle. \<^descr> \<^theory_text>\print_bundles\ prints the named bundles that are available in the current context; the ``\!\'' option indicates extra verbosity. \<^descr> \<^theory_text>\include b\<^sub>1 \ b\<^sub>n\ activates the declarations from the given bundles in a proof body (forward mode). This is analogous to \<^theory_text>\note\ (\secref{sec:proof-facts}) with the expanded bundles. \<^descr> \<^theory_text>\including b\<^sub>1 \ b\<^sub>n\ is similar to \<^theory_text>\include\, but works in proof refinement (backward mode). This is analogous to \<^theory_text>\using\ (\secref{sec:proof-facts}) with the expanded bundles. \<^descr> \<^theory_text>\includes b\<^sub>1 \ b\<^sub>n\ is similar to \<^theory_text>\include\, but applies to a confined specification context: unnamed \<^theory_text>\context\s and long statements of \<^theory_text>\theorem\. \<^descr> \<^theory_text>\opening b\<^sub>1 \ b\<^sub>n\ is similar to \<^theory_text>\includes\, but applies to a named specification context: \<^theory_text>\locale\s, \<^theory_text>\class\es and named \<^theory_text>\context\s. The effect is confined to the surface context within the specification block itself and the corresponding \<^theory_text>\begin\ / \<^theory_text>\end\ block. \<^descr> \<^theory_text>\unbundle b\<^sub>1 \ b\<^sub>n\ activates the declarations from the given bundles in the current local theory context. This is analogous to \<^theory_text>\lemmas\ (\secref{sec:theorems}) with the expanded bundles. Here is an artificial example of bundling various configuration options: \ (*<*)experiment begin(*>*) bundle trace = [[simp_trace, linarith_trace, metis_trace, smt_trace]] lemma "x = x" including trace by metis (*<*)end(*>*) section \Term definitions \label{sec:term-definitions}\ text \ \begin{matharray}{rcll} @{command_def "definition"} & : & \local_theory \ local_theory\ \\ @{attribute_def "defn"} & : & \attribute\ \\ @{command_def "print_defn_rules"}\\<^sup>*\ & : & \context \\ \\ @{command_def "abbreviation"} & : & \local_theory \ local_theory\ \\ @{command_def "print_abbrevs"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} Term definitions may either happen within the logic (as equational axioms of a certain form (see also \secref{sec:overloading}), or outside of it as rewrite system on abstract syntax. The second form is called ``abbreviation''. \<^rail>\ @@{command definition} decl? definition ; @@{command abbreviation} @{syntax mode}? decl? abbreviation ; @@{command print_abbrevs} ('!'?) ; decl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? @'where' ; definition: @{syntax thmdecl}? @{syntax prop} @{syntax spec_prems} @{syntax for_fixes} ; abbreviation: @{syntax prop} @{syntax for_fixes} \ \<^descr> \<^theory_text>\definition c where eq\ produces an internal definition \c \ t\ according to the specification given as \eq\, which is then turned into a proven fact. The given proposition may deviate from internal meta-level equality according to the rewrite rules declared as @{attribute defn} by the object-logic. This usually covers object-level equality \x = y\ and equivalence \A \ B\. End-users normally need not change the @{attribute defn} setup. Definitions may be presented with explicit arguments on the LHS, as well as additional conditions, e.g.\ \f x y = t\ instead of \f \ \x y. t\ and \y \ 0 \ g x y = u\ instead of an unrestricted \g \ \x y. u\. \<^descr> \<^theory_text>\print_defn_rules\ prints the definitional rewrite rules declared via @{attribute defn} in the current context. \<^descr> \<^theory_text>\abbreviation c where eq\ introduces a syntactic constant which is associated with a certain term according to the meta-level equality \eq\. Abbreviations participate in the usual type-inference process, but are expanded before the logic ever sees them. Pretty printing of terms involves higher-order rewriting with rules stemming from reverted abbreviations. This needs some care to avoid overlapping or looping syntactic replacements! The optional \mode\ specification restricts output to a particular print mode; using ``\input\'' here achieves the effect of one-way abbreviations. The mode may also include an ``\<^theory_text>\output\'' qualifier that affects the concrete syntax declared for abbreviations, cf.\ \<^theory_text>\syntax\ in \secref{sec:syn-trans}. \<^descr> \<^theory_text>\print_abbrevs\ prints all constant abbreviations of the current context; the ``\!\'' option indicates extra verbosity. \ section \Axiomatizations \label{sec:axiomatizations}\ text \ \begin{matharray}{rcll} @{command_def "axiomatization"} & : & \theory \ theory\ & (axiomatic!) \\ \end{matharray} \<^rail>\ @@{command axiomatization} @{syntax vars}? (@'where' axiomatization)? ; axiomatization: (@{syntax thmdecl} @{syntax prop} + @'and') @{syntax spec_prems} @{syntax for_fixes} \ \<^descr> \<^theory_text>\axiomatization c\<^sub>1 \ c\<^sub>m where \\<^sub>1 \ \\<^sub>n\ introduces several constants simultaneously and states axiomatic properties for these. The constants are marked as being specified once and for all, which prevents additional specifications for the same constants later on, but it is always possible to emit axiomatizations without referring to particular constants. Note that lack of precise dependency tracking of axiomatizations may disrupt the well-formedness of an otherwise definitional theory. Axiomatization is restricted to a global theory context: support for local theory targets \secref{sec:target} would introduce an extra dimension of uncertainty what the written specifications really are, and make it infeasible to argue why they are correct. Axiomatic specifications are required when declaring a new logical system within Isabelle/Pure, but in an application environment like Isabelle/HOL the user normally stays within definitional mechanisms provided by the logic and its libraries. \ section \Generic declarations\ text \ \begin{matharray}{rcl} @{command_def "declaration"} & : & \local_theory \ local_theory\ \\ @{command_def "syntax_declaration"} & : & \local_theory \ local_theory\ \\ @{command_def "declare"} & : & \local_theory \ local_theory\ \\ \end{matharray} Arbitrary operations on the background context may be wrapped-up as generic declaration elements. Since the underlying concept of local theories may be subject to later re-interpretation, there is an additional dependency on a morphism that tells the difference of the original declaration context wrt.\ the application context encountered later on. A fact declaration is an important special case: it consists of a theorem which is applied to the context by means of an attribute. \<^rail>\ (@@{command declaration} | @@{command syntax_declaration}) ('(' @'pervasive' ')')? \ @{syntax text} ; @@{command declare} (@{syntax thms} + @'and') \ \<^descr> \<^theory_text>\declaration d\ adds the declaration function \d\ of ML type \<^ML_type>\declaration\, to the current local theory under construction. In later application contexts, the function is transformed according to the morphisms being involved in the interpretation hierarchy. If the \<^theory_text>\(pervasive)\ option is given, the corresponding declaration is applied to all possible contexts involved, including the global background theory. \<^descr> \<^theory_text>\syntax_declaration\ is similar to \<^theory_text>\declaration\, but is meant to affect only ``syntactic'' tools by convention (such as notation and type-checking information). \<^descr> \<^theory_text>\declare thms\ declares theorems to the current local theory context. No theorem binding is involved here, unlike \<^theory_text>\lemmas\ (cf.\ \secref{sec:theorems}), so \<^theory_text>\declare\ only has the effect of applying attributes as included in the theorem specification. \ section \Locales \label{sec:locale}\ text \ A locale is a functor that maps parameters (including implicit type parameters) and a specification to a list of declarations. The syntax of locales is modeled after the Isar proof context commands (cf.\ \secref{sec:proof-context}). Locale hierarchies are supported by maintaining a graph of dependencies between locale instances in the global theory. Dependencies may be introduced through import (where a locale is defined as sublocale of the imported instances) or by proving that an existing locale is a sublocale of one or several locale instances. A locale may be opened with the purpose of appending to its list of declarations (cf.\ \secref{sec:target}). When opening a locale declarations from all dependencies are collected and are presented as a local theory. In this process, which is called \<^emph>\roundup\, redundant locale instances are omitted. A locale instance is redundant if it is subsumed by an instance encountered earlier. A more detailed description of this process is - available elsewhere @{cite Ballarin2014}. + available elsewhere \<^cite>\Ballarin2014\. \ subsection \Locale expressions \label{sec:locale-expr}\ text \ A \<^emph>\locale expression\ denotes a context composed of instances of existing locales. The context consists of the declaration elements from the locale instances. Redundant locale instances are omitted according to roundup. \<^rail>\ @{syntax_def locale_expr}: (instance + '+') @{syntax for_fixes} ; instance: (qualifier ':')? @{syntax name} (pos_insts | named_insts) \ rewrites? ; qualifier: @{syntax name} ('?')? ; pos_insts: ('_' | @{syntax term})* ; named_insts: @'where' (@{syntax name} '=' @{syntax term} + @'and') ; rewrites: @'rewrites' (@{syntax thmdecl}? @{syntax prop} + @'and') \ A locale instance consists of a reference to a locale and either positional or named parameter instantiations optionally followed by rewrites clauses. Identical instantiations (that is, those that instantiate a parameter by itself) may be omitted. The notation ``\_\'' enables to omit the instantiation for a parameter inside a positional instantiation. Terms in instantiations are from the context the locale expressions is declared in. Local names may be added to this context with the optional \<^theory_text>\for\ clause. This is useful for shadowing names bound in outer contexts, and for declaring syntax. In addition, syntax declarations from one instance are effective when parsing subsequent instances of the same expression. Instances have an optional qualifier which applies to names in declarations. Names include local definitions and theorem names. If present, the qualifier itself is either mandatory (default) or non-mandatory (when followed by ``\<^verbatim>\?\''). Non-mandatory means that the qualifier may be omitted on input. Qualifiers only affect name spaces; they play no role in determining whether one locale instance subsumes another. Rewrite clauses amend instances with equations that act as rewrite rules. This is particularly useful for changing concepts introduced through definitions. Rewrite clauses are available only in interpretation commands (see \secref{sec:locale-interpretation} below) and must be proved the user. \ subsection \Locale declarations\ text \ \begin{tabular}{rcl} @{command_def "locale"} & : & \theory \ local_theory\ \\ @{command_def "experiment"} & : & \theory \ local_theory\ \\ @{command_def "print_locale"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_locales"}\\<^sup>*\ & : & \context \\ \\ @{command_def "locale_deps"}\\<^sup>*\ & : & \context \\ \\ \end{tabular} @{index_ref \\<^theory_text>\fixes\ (element)\} @{index_ref \\<^theory_text>\constrains\ (element)\} @{index_ref \\<^theory_text>\assumes\ (element)\} @{index_ref \\<^theory_text>\defines\ (element)\} @{index_ref \\<^theory_text>\notes\ (element)\} \<^rail>\ @@{command locale} @{syntax name} ('=' @{syntax locale})? @'begin'? ; @@{command experiment} (@{syntax context_elem}*) @'begin' ; @@{command print_locale} '!'? @{syntax name} ; @@{command print_locales} ('!'?) ; @{syntax_def locale}: @{syntax context_elem}+ | @{syntax_ref "opening"} ('+' (@{syntax context_elem}+))? | @{syntax locale_expr} @{syntax_ref "opening"}? ('+' (@{syntax context_elem}+))? ; @{syntax_def context_elem}: @'fixes' @{syntax vars} | @'constrains' (@{syntax name} '::' @{syntax type} + @'and') | @'assumes' (@{syntax props} + @'and') | @'defines' (@{syntax thmdecl}? @{syntax prop} @{syntax prop_pat}? + @'and') | @'notes' (@{syntax thmdef}? @{syntax thms} + @'and') \ \<^descr> \<^theory_text>\locale loc = import opening bundles + body\ defines a new locale \loc\ as a context consisting of a certain view of existing locales (\import\) plus some additional elements (\body\) with declaration \bundles\ enriching the context of the command itself. All \import\, \bundles\ and \body\ are optional; the degenerate form \<^theory_text>\locale loc\ defines an empty locale, which may still be useful to collect declarations of facts later on. Type-inference on locale expressions automatically takes care of the most general typing that the combined context elements may acquire. The \import\ consists of a locale expression; see \secref{sec:locale-expr} above. Its \<^theory_text>\for\ clause defines the parameters of \import\. These are parameters of the defined locale. Locale parameters whose instantiation is omitted automatically extend the (possibly empty) \<^theory_text>\for\ clause: they are inserted at its beginning. This means that these parameters may be referred to from within the expression and also in the subsequent context elements and provides a notational convenience for the inheritance of parameters in locale declarations. Declarations from \bundles\, see \secref{sec:bundle}, are effective in the entire command including a subsequent \<^theory_text>\begin\ / \<^theory_text>\end\ block, but they do not contribute to the declarations stored in the locale. The \body\ consists of context elements: \<^descr> @{element "fixes"}~\x :: \ (mx)\ declares a local parameter of type \\\ and mixfix annotation \mx\ (both are optional). The special syntax declaration ``\(\@{keyword_ref "structure"}\)\'' means that \x\ may be referenced implicitly in this context. \<^descr> @{element "constrains"}~\x :: \\ introduces a type constraint \\\ on the local parameter \x\. This element is deprecated. The type constraint should be introduced in the \<^theory_text>\for\ clause or the relevant @{element "fixes"} element. \<^descr> @{element "assumes"}~\a: \\<^sub>1 \ \\<^sub>n\ introduces local premises, similar to \<^theory_text>\assume\ within a proof (cf.\ \secref{sec:proof-context}). \<^descr> @{element "defines"}~\a: x \ t\ defines a previously declared parameter. This is similar to \<^theory_text>\define\ within a proof (cf.\ \secref{sec:proof-context}), but @{element "defines"} is restricted to Pure equalities and the defined variable needs to be declared beforehand via @{element "fixes"}. The left-hand side of the equation may have additional arguments, e.g.\ ``@{element "defines"}~\f x\<^sub>1 \ x\<^sub>n \ t\'', which need to be free in the context. \<^descr> @{element "notes"}~\a = b\<^sub>1 \ b\<^sub>n\ reconsiders facts within a local context. Most notably, this may include arbitrary declarations in any attribute specifications included here, e.g.\ a local @{attribute simp} rule. Both @{element "assumes"} and @{element "defines"} elements contribute to the locale specification. When defining an operation derived from the parameters, \<^theory_text>\definition\ (\secref{sec:term-definitions}) is usually more appropriate. Note that ``\<^theory_text>\(is p\<^sub>1 \ p\<^sub>n)\'' patterns given in the syntax of @{element "assumes"} and @{element "defines"} above are illegal in locale definitions. In the long goal format of \secref{sec:goals}, term bindings may be included as expected, though. \<^medskip> Locale specifications are ``closed up'' by turning the given text into a predicate definition \loc_axioms\ and deriving the original assumptions as local lemmas (modulo local definitions). The predicate statement covers only the newly specified assumptions, omitting the content of included locale expressions. The full cumulative view is only provided on export, involving another predicate \loc\ that refers to the complete specification text. In any case, the predicate arguments are those locale parameters that actually occur in the respective piece of text. Also these predicates operate at the meta-level in theory, but the locale packages attempts to internalize statements according to the object-logic setup (e.g.\ replacing \\\ by \\\, and \\\ by \\\ in HOL; see also \secref{sec:object-logic}). Separate introduction rules \loc_axioms.intro\ and \loc.intro\ are provided as well. \<^descr> \<^theory_text>\experiment body begin\ opens an anonymous locale context with private naming policy. Specifications in its body are inaccessible from outside. This is useful to perform experiments, without polluting the name space. \<^descr> \<^theory_text>\print_locale "locale"\ prints the contents of the named locale. The command omits @{element "notes"} elements by default. Use \<^theory_text>\print_locale!\ to have them included. \<^descr> \<^theory_text>\print_locales\ prints the names of all locales of the current theory; the ``\!\'' option indicates extra verbosity. \<^descr> \<^theory_text>\locale_deps\ visualizes all locales and their relations as a Hasse diagram. This includes locales defined as type classes (\secref{sec:class}). \ subsection \Locale interpretation \label{sec:locale-interpretation}\ text \ \begin{matharray}{rcl} @{command "interpretation"} & : & \local_theory \ proof(prove)\ \\ @{command_def "interpret"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "global_interpretation"} & : & \theory | local_theory \ proof(prove)\ \\ @{command_def "sublocale"} & : & \theory | local_theory \ proof(prove)\ \\ @{command_def "print_interps"}\\<^sup>*\ & : & \context \\ \\ @{method_def intro_locales} & : & \method\ \\ @{method_def unfold_locales} & : & \method\ \\ @{attribute_def trace_locales} & : & \mbox{\attribute\ \quad default \false\} \\ \end{matharray} Locales may be instantiated, and the resulting instantiated declarations added to the current context. This requires proof (of the instantiated specification) and is called \<^emph>\locale interpretation\. Interpretation is possible within arbitrary local theories (\<^theory_text>\interpretation\), within proof bodies (\<^theory_text>\interpret\), into global theories (\<^theory_text>\global_interpretation\) and into locales (\<^theory_text>\sublocale\). \<^rail>\ @@{command interpretation} @{syntax locale_expr} ; @@{command interpret} @{syntax locale_expr} ; @@{command global_interpretation} @{syntax locale_expr} definitions? ; @@{command sublocale} (@{syntax name} ('<' | '\'))? @{syntax locale_expr} \ definitions? ; @@{command print_interps} @{syntax name} ; definitions: @'defines' (@{syntax thmdecl}? @{syntax name} \ @{syntax mixfix}? '=' @{syntax term} + @'and'); \ The core of each interpretation command is a locale expression \expr\; the command generates proof obligations for the instantiated specifications. Once these are discharged by the user, instantiated declarations (in particular, facts) are added to the context in a post-processing phase, in a manner specific to each command. Interpretation commands are aware of interpretations that are already active: post-processing is achieved through a variant of roundup that takes interpretations of the current global or local theory into account. In order to simplify the proof obligations according to existing interpretations use methods @{method intro_locales} or @{method unfold_locales}. Rewrites clauses \<^theory_text>\rewrites eqns\ occur within expressions. They amend the morphism through which a locale instance is interpreted with rewrite rules, also called rewrite morphisms. This is particularly useful for interpreting concepts introduced through definitions. The equations must be proved the user. To enable syntax of the instantiated locale within the equation, while reading a locale expression, equations of a locale instance are read in a temporary context where the instance is already activated. If activation fails, typically due to duplicate constant declarations, processing falls back to reading the equation first. Given definitions \defs\ produce corresponding definitions in the local theory's underlying target \<^emph>\and\ amend the morphism with rewrite rules stemming from the symmetric of those definitions. Hence these need not be proved explicitly the user. Such rewrite definitions are a even more useful device for interpreting concepts introduced through definitions, but they are only supported for interpretation commands operating in a local theory whose implementing target actually supports this. Note that despite the suggestive \<^theory_text>\and\ connective, \defs\ are processed sequentially without mutual recursion. \<^descr> \<^theory_text>\interpretation expr\ interprets \expr\ into a local theory such that its lifetime is limited to the current context block (e.g. a locale or unnamed context). At the closing @{command end} of the block the interpretation and its declarations disappear. Hence facts based on interpretation can be established without creating permanent links to the interpreted locale instances, as would be the case with @{command sublocale}. When used on the level of a global theory, there is no end of a current context block, hence \<^theory_text>\interpretation\ behaves identically to \<^theory_text>\global_interpretation\ then. \<^descr> \<^theory_text>\interpret expr\ interprets \expr\ into a proof context: the interpretation and its declarations disappear when closing the current proof block. Note that for \<^theory_text>\interpret\ the \eqns\ should be explicitly universally quantified. \<^descr> \<^theory_text>\global_interpretation expr defines defs\ interprets \expr\ into a global theory. When adding declarations to locales, interpreted versions of these declarations are added to the global theory for all interpretations in the global theory as well. That is, interpretations into global theories dynamically participate in any declarations added to locales. Free variables in the interpreted expression are allowed. They are turned into schematic variables in the generated declarations. In order to use a free variable whose name is already bound in the context --- for example, because a constant of that name exists --- add it to the \<^theory_text>\for\ clause. When used in a nested target, resulting declarations are propagated through the whole target stack. \<^descr> \<^theory_text>\sublocale name \ expr defines defs\ interprets \expr\ into the locale \name\. A proof that the specification of \name\ implies the specification of \expr\ is required. As in the localized version of the theorem command, the proof is in the context of \name\. After the proof obligation has been discharged, the locale hierarchy is changed as if \name\ imported \expr\ (hence the name \<^theory_text>\sublocale\). When the context of \name\ is subsequently entered, traversing the locale hierarchy will involve the locale instances of \expr\, and their declarations will be added to the context. This makes \<^theory_text>\sublocale\ dynamic: extensions of a locale that is instantiated in \expr\ may take place after the \<^theory_text>\sublocale\ declaration and still become available in the context. Circular \<^theory_text>\sublocale\ declarations are allowed as long as they do not lead to infinite chains. If interpretations of \name\ exist in the current global theory, the command adds interpretations for \expr\ as well, with the same qualifier, although only for fragments of \expr\ that are not interpreted in the theory already. Rewrites clauses in the expression or rewrite definitions \defs\ can help break infinite chains induced by circular \<^theory_text>\sublocale\ declarations. In a named context block the \<^theory_text>\sublocale\ command may also be used, but the locale argument must be omitted. The command then refers to the locale (or class) target of the context block. \<^descr> \<^theory_text>\print_interps name\ lists all interpretations of locale \name\ in the current theory or proof context, including those due to a combination of an \<^theory_text>\interpretation\ or \<^theory_text>\interpret\ and one or several \<^theory_text>\sublocale\ declarations. \<^descr> @{method intro_locales} and @{method unfold_locales} repeatedly expand all introduction rules of locale predicates of the theory. While @{method intro_locales} only applies the \loc.intro\ introduction rules and therefore does not descend to assumptions, @{method unfold_locales} is more aggressive and applies \loc_axioms.intro\ as well. Both methods are aware of locale specifications entailed by the context, both from target statements, and from interpretations (see below). New goals that are entailed by the current context are discharged automatically. While @{method unfold_locales} is part of the default method for \<^theory_text>\proof\ and often invoked ``behind the scenes'', @{method intro_locales} helps understand which proof obligations originated from which locale instances. The latter method is useful while developing proofs but rare in finished developments. \<^descr> @{attribute trace_locales}, when set to \true\, prints the locale instances activated during roundup. Use this when locale commands yield obscure errors or for understanding local theories created by complex locale hierarchies. \begin{warn} If a global theory inherits declarations (body elements) for a locale from one parent and an interpretation of that locale from another parent, the interpretation will not be applied to the declarations. \end{warn} \begin{warn} Since attributes are applied to interpreted theorems, interpretation may modify the context of common proof tools, e.g.\ the Simplifier or Classical Reasoner. As the behaviour of such tools is \<^emph>\not\ stable under interpretation morphisms, manual declarations might have to be added to the target context of the interpretation to revert such declarations. \end{warn} \begin{warn} An interpretation in a local theory or proof context may subsume previous interpretations. This happens if the same specification fragment is interpreted twice and the instantiation of the second interpretation is more general than the interpretation of the first. The locale package does not attempt to remove subsumed interpretations. \end{warn} \begin{warn} While \<^theory_text>\interpretation (in c) \\ is admissible, it is not useful since its result is discarded immediately. \end{warn} \ section \Classes \label{sec:class}\ text \ \begin{matharray}{rcl} @{command_def "class"} & : & \theory \ local_theory\ \\ @{command_def "instantiation"} & : & \theory \ local_theory\ \\ @{command_def "instance"} & : & \local_theory \ local_theory\ \\ @{command "instance"} & : & \theory \ proof(prove)\ \\ @{command_def "subclass"} & : & \local_theory \ local_theory\ \\ @{command_def "print_classes"}\\<^sup>*\ & : & \context \\ \\ @{command_def "class_deps"}\\<^sup>*\ & : & \context \\ \\ @{method_def intro_classes} & : & \method\ \\ \end{matharray} A class is a particular locale with \<^emph>\exactly one\ type variable \\\. Beyond the underlying locale, a corresponding type class is established which is - interpreted logically as axiomatic type class @{cite "Wenzel:1997:TPHOL"} + interpreted logically as axiomatic type class \<^cite>\"Wenzel:1997:TPHOL"\ whose logical content are the assumptions of the locale. Thus, classes provide the full generality of locales combined with the commodity of type - classes (notably type-inference). See @{cite "isabelle-classes"} for a short + classes (notably type-inference). See \<^cite>\"isabelle-classes"\ for a short tutorial. \<^rail>\ @@{command class} class_spec @'begin'? ; class_spec: @{syntax name} '=' ((@{syntax name} @{syntax_ref "opening"}? '+' (@{syntax context_elem}+)) | @{syntax name} @{syntax_ref "opening"}? | @{syntax_ref "opening"}? '+' (@{syntax context_elem}+)) ; @@{command instantiation} (@{syntax name} + @'and') '::' @{syntax arity} @'begin' ; @@{command instance} (() | (@{syntax name} + @'and') '::' @{syntax arity} | @{syntax name} ('<' | '\') @{syntax name} ) ; @@{command subclass} @{syntax name} ; @@{command class_deps} (class_bounds class_bounds?)? ; class_bounds: @{syntax sort} | '(' (@{syntax sort} + @'|') ')' \ \<^descr> \<^theory_text>\class c = superclasses bundles + body\ defines a new class \c\, inheriting from \superclasses\. This introduces a locale \c\ with import of all locales \superclasses\. Any @{element "fixes"} in \body\ are lifted to the global theory level (\<^emph>\class operations\ \f\<^sub>1, \, f\<^sub>n\ of class \c\), mapping the local type parameter \\\ to a schematic type variable \?\ :: c\. Likewise, @{element "assumes"} in \body\ are also lifted, mapping each local parameter \f :: \[\]\ to its corresponding global constant \f :: \[?\ :: c]\. The corresponding introduction rule is provided as \c_class_axioms.intro\. This rule should be rarely needed directly --- the @{method intro_classes} method takes care of the details of class membership proofs. Optionally given \bundles\ take effect in the surface context within the \body\ and the potentially following \<^theory_text>\begin\ / \<^theory_text>\end\ block. \<^descr> \<^theory_text>\instantiation t :: (s\<^sub>1, \, s\<^sub>n)s begin\ opens a target (cf.\ \secref{sec:target}) which allows to specify class operations \f\<^sub>1, \, f\<^sub>n\ corresponding to sort \s\ at the particular type instance \(\\<^sub>1 :: s\<^sub>1, \, \\<^sub>n :: s\<^sub>n) t\. A plain \<^theory_text>\instance\ command in the target body poses a goal stating these type arities. The target is concluded by an @{command_ref (local) "end"} command. Note that a list of simultaneous type constructors may be given; this corresponds nicely to mutually recursive type definitions, e.g.\ in Isabelle/HOL. \<^descr> \<^theory_text>\instance\ in an instantiation target body sets up a goal stating the type arities claimed at the opening \<^theory_text>\instantiation\. The proof would usually proceed by @{method intro_classes}, and then establish the characteristic theorems of the type classes involved. After finishing the proof, the background theory will be augmented by the proven type arities. On the theory level, \<^theory_text>\instance t :: (s\<^sub>1, \, s\<^sub>n)s\ provides a convenient way to instantiate a type class with no need to specify operations: one can continue with the instantiation proof immediately. \<^descr> \<^theory_text>\subclass c\ in a class context for class \d\ sets up a goal stating that class \c\ is logically contained in class \d\. After finishing the proof, class \d\ is proven to be subclass \c\ and the locale \c\ is interpreted into \d\ simultaneously. A weakened form of this is available through a further variant of @{command instance}: \<^theory_text>\instance c\<^sub>1 \ c\<^sub>2\ opens a proof that class \c\<^sub>2\ implies \c\<^sub>1\ without reference to the underlying locales; this is useful if the properties to prove the logical connection are not sufficient on the locale level but on the theory level. \<^descr> \<^theory_text>\print_classes\ prints all classes in the current theory. \<^descr> \<^theory_text>\class_deps\ visualizes classes and their subclass relations as a directed acyclic graph. By default, all classes from the current theory context are show. This may be restricted by optional bounds as follows: \<^theory_text>\class_deps upper\ or \<^theory_text>\class_deps upper lower\. A class is visualized, iff it is a subclass of some sort from \upper\ and a superclass of some sort from \lower\. \<^descr> @{method intro_classes} repeatedly expands all class introduction rules of this theory. Note that this method usually needs not be named explicitly, as it is already included in the default proof step (e.g.\ of \<^theory_text>\proof\). In particular, instantiation of trivial (syntactic) classes may be performed by a single ``\<^theory_text>\..\'' proof step. \ subsection \The class target\ text \ %FIXME check A named context may refer to a locale (cf.\ \secref{sec:target}). If this locale is also a class \c\, apart from the common locale target behaviour the following happens. \<^item> Local constant declarations \g[\]\ referring to the local type parameter \\\ and local parameters \f[\]\ are accompanied by theory-level constants \g[?\ :: c]\ referring to theory-level class operations \f[?\ :: c]\. \<^item> Local theorem bindings are lifted as are assumptions. \<^item> Local syntax refers to local operations \g[\]\ and global operations \g[?\ :: c]\ uniformly. Type inference resolves ambiguities. In rare cases, manual type annotations are needed. \ subsection \Co-regularity of type classes and arities\ text \ The class relation together with the collection of type-constructor arities must obey the principle of \<^emph>\co-regularity\ as defined below. \<^medskip> For the subsequent formulation of co-regularity we assume that the class relation is closed by transitivity and reflexivity. Moreover the collection of arities \t :: (\<^vec>s)c\ is completed such that \t :: (\<^vec>s)c\ and \c \ c'\ implies \t :: (\<^vec>s)c'\ for all such declarations. Treating sorts as finite sets of classes (meaning the intersection), the class relation \c\<^sub>1 \ c\<^sub>2\ is extended to sorts as follows: \[ \s\<^sub>1 \ s\<^sub>2 \ \c\<^sub>2 \ s\<^sub>2. \c\<^sub>1 \ s\<^sub>1. c\<^sub>1 \ c\<^sub>2\ \] This relation on sorts is further extended to tuples of sorts (of the same length) in the component-wise way. \<^medskip> Co-regularity of the class relation together with the arities relation means: \[ \t :: (\<^vec>s\<^sub>1)c\<^sub>1 \ t :: (\<^vec>s\<^sub>2)c\<^sub>2 \ c\<^sub>1 \ c\<^sub>2 \ \<^vec>s\<^sub>1 \ \<^vec>s\<^sub>2\ \] for all such arities. In other words, whenever the result classes of some type-constructor arities are related, then the argument sorts need to be related in the same way. \<^medskip> Co-regularity is a very fundamental property of the order-sorted algebra of types. For example, it entails principal types and most general unifiers, - e.g.\ see @{cite "nipkow-prehofer"}. + e.g.\ see \<^cite>\"nipkow-prehofer"\. \ section \Overloaded constant definitions \label{sec:overloading}\ text \ Definitions essentially express abbreviations within the logic. The simplest form of a definition is \c :: \ \ t\, where \c\ is a new constant and \t\ is a closed term that does not mention \c\. Moreover, so-called \<^emph>\hidden polymorphism\ is excluded: all type variables in \t\ need to occur in its type \\\. \<^emph>\Overloading\ means that a constant being declared as \c :: \ decl\ may be defined separately on type instances \c :: (\\<^sub>1, \, \\<^sub>n)\ decl\ for each type constructor \\\. At most occasions overloading will be used in a Haskell-like fashion together with type classes by means of \<^theory_text>\instantiation\ (see \secref{sec:class}). Sometimes low-level overloading is desirable; this is supported by \<^theory_text>\consts\ and \<^theory_text>\overloading\ explained below. The right-hand side of overloaded definitions may mention overloaded constants recursively at type instances corresponding to the immediate argument types \\\<^sub>1, \, \\<^sub>n\. Incomplete specification patterns impose global constraints on all occurrences. E.g.\ \d :: \ \ \\ on the left-hand side means that all corresponding occurrences on some right-hand side need to be an instance of this, and general \d :: \ \ \\ will be disallowed. Full - details are given by Kun\v{c}ar @{cite "Kuncar:2015"}. + details are given by Kun\v{c}ar \<^cite>\"Kuncar:2015"\. \<^medskip> The \<^theory_text>\consts\ command and the \<^theory_text>\overloading\ target provide a convenient interface for end-users. Regular specification elements such as @{command definition}, @{command inductive}, @{command function} may be used in the body. It is also possible to use \<^theory_text>\consts c :: \\ with later \<^theory_text>\overloading c \ c :: \\ to keep the declaration and definition of a constant separate. \begin{matharray}{rcl} @{command_def "consts"} & : & \theory \ theory\ \\ @{command_def "overloading"} & : & \theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command consts} ((@{syntax name} '::' @{syntax type} @{syntax mixfix}?) +) ; @@{command overloading} ( spec + ) @'begin' ; spec: @{syntax name} ( '\' | '==' ) @{syntax term} ( '(' @'unchecked' ')' )? \ \<^descr> \<^theory_text>\consts c :: \\ declares constant \c\ to have any instance of type scheme \\\. The optional mixfix annotations may attach concrete syntax to the constants declared. \<^descr> \<^theory_text>\overloading x\<^sub>1 \ c\<^sub>1 :: \\<^sub>1 \ x\<^sub>n \ c\<^sub>n :: \\<^sub>n begin \ end\ defines a theory target (cf.\ \secref{sec:target}) which allows to specify already declared constants via definitions in the body. These are identified by an explicitly given mapping from variable names \x\<^sub>i\ to constants \c\<^sub>i\ at particular type instances. The definitions themselves are established using common specification tools, using the names \x\<^sub>i\ as reference to the corresponding constants. Option \<^theory_text>\(unchecked)\ disables global dependency checks for the corresponding definition, which is occasionally useful for exotic overloading; this is a form of axiomatic specification. It is at the discretion of the user to avoid malformed theory specifications! \ subsubsection \Example\ consts Length :: "'a \ nat" overloading Length\<^sub>0 \ "Length :: unit \ nat" Length\<^sub>1 \ "Length :: 'a \ unit \ nat" Length\<^sub>2 \ "Length :: 'a \ 'b \ unit \ nat" Length\<^sub>3 \ "Length :: 'a \ 'b \ 'c \ unit \ nat" begin fun Length\<^sub>0 :: "unit \ nat" where "Length\<^sub>0 () = 0" fun Length\<^sub>1 :: "'a \ unit \ nat" where "Length\<^sub>1 (a, ()) = 1" fun Length\<^sub>2 :: "'a \ 'b \ unit \ nat" where "Length\<^sub>2 (a, b, ()) = 2" fun Length\<^sub>3 :: "'a \ 'b \ 'c \ unit \ nat" where "Length\<^sub>3 (a, b, c, ()) = 3" end lemma "Length (a, b, c, ()) = 3" by simp lemma "Length ((a, b), (c, d), ()) = 2" by simp lemma "Length ((a, b, c, d, e), ()) = 1" by simp section \Incorporating ML code \label{sec:ML}\ text \ \begin{matharray}{rcl} @{command_def "SML_file"} & : & \local_theory \ local_theory\ \\ @{command_def "SML_file_debug"} & : & \local_theory \ local_theory\ \\ @{command_def "SML_file_no_debug"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_file"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_file_debug"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_file_no_debug"} & : & \local_theory \ local_theory\ \\ @{command_def "ML"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_export"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_prf"} & : & \proof \ proof\ \\ @{command_def "ML_val"} & : & \any \\ \\ @{command_def "ML_command"} & : & \any \\ \\ @{command_def "setup"} & : & \theory \ theory\ \\ @{command_def "local_setup"} & : & \local_theory \ local_theory\ \\ @{command_def "attribute_setup"} & : & \local_theory \ local_theory\ \\ \end{matharray} \begin{tabular}{rcll} @{attribute_def ML_print_depth} & : & \attribute\ & default 10 \\ @{attribute_def ML_source_trace} & : & \attribute\ & default \false\ \\ @{attribute_def ML_debugger} & : & \attribute\ & default \false\ \\ @{attribute_def ML_exception_trace} & : & \attribute\ & default \false\ \\ @{attribute_def ML_exception_debugger} & : & \attribute\ & default \false\ \\ @{attribute_def ML_environment} & : & \attribute\ & default \Isabelle\ \\ \end{tabular} \<^rail>\ (@@{command SML_file} | @@{command SML_file_debug} | @@{command SML_file_no_debug} | @@{command ML_file} | @@{command ML_file_debug} | @@{command ML_file_no_debug}) @{syntax name} ';'? ; (@@{command ML} | @@{command ML_export} | @@{command ML_prf} | @@{command ML_val} | @@{command ML_command} | @@{command setup} | @@{command local_setup}) @{syntax text} ; @@{command attribute_setup} @{syntax name} '=' @{syntax text} @{syntax text}? \ \<^descr> \<^theory_text>\SML_file name\ reads and evaluates the given Standard ML file. Top-level SML bindings are stored within the (global or local) theory context; the initial environment is restricted to the Standard ML implementation of Poly/ML, without the many add-ons of Isabelle/ML. Multiple \<^theory_text>\SML_file\ commands may be used to build larger Standard ML projects, independently of the regular Isabelle/ML environment. \<^descr> \<^theory_text>\ML_file name\ reads and evaluates the given ML file. The current theory context is passed down to the ML toplevel and may be modified, using \<^ML>\Context.>>\ or derived ML commands. Top-level ML bindings are stored within the (global or local) theory context. \<^descr> \<^theory_text>\SML_file_debug\, \<^theory_text>\SML_file_no_debug\, \<^theory_text>\ML_file_debug\, and \<^theory_text>\ML_file_no_debug\ change the @{attribute ML_debugger} option locally while the given file is compiled. \<^descr> \<^theory_text>\ML\ is similar to \<^theory_text>\ML_file\, but evaluates directly the given \text\. Top-level ML bindings are stored within the (global or local) theory context. \<^descr> \<^theory_text>\ML_export\ is similar to \<^theory_text>\ML\, but the resulting toplevel bindings are exported to the global bootstrap environment of the ML process --- it has a lasting effect that cannot be retracted. This allows ML evaluation without a formal theory context, e.g. for command-line tools via @{tool - process} @{cite "isabelle-system"}. + process} \<^cite>\"isabelle-system"\. \<^descr> \<^theory_text>\ML_prf\ is analogous to \<^theory_text>\ML\ but works within a proof context. Top-level ML bindings are stored within the proof context in a purely sequential fashion, disregarding the nested proof structure. ML bindings introduced by \<^theory_text>\ML_prf\ are discarded at the end of the proof. \<^descr> \<^theory_text>\ML_val\ and \<^theory_text>\ML_command\ are diagnostic versions of \<^theory_text>\ML\, which means that the context may not be updated. \<^theory_text>\ML_val\ echos the bindings produced at the ML toplevel, but \<^theory_text>\ML_command\ is silent. \<^descr> \<^theory_text>\setup "text"\ changes the current theory context by applying \text\, which refers to an ML expression of type \<^ML_type>\theory -> theory\. This enables to initialize any object-logic specific tools and packages written in ML, for example. \<^descr> \<^theory_text>\local_setup\ is similar to \<^theory_text>\setup\ for a local theory context, and an ML expression of type \<^ML_type>\local_theory -> local_theory\. This allows to invoke local theory specification packages without going through concrete outer syntax, for example. \<^descr> \<^theory_text>\attribute_setup name = "text" description\ defines an attribute in the current context. The given \text\ has to be an ML expression of type \<^ML_type>\attribute context_parser\, cf.\ basic parsers defined in structure \<^ML_structure>\Args\ and \<^ML_structure>\Attrib\. In principle, attributes can operate both on a given theorem and the implicit context, although in practice only one is modified and the other serves as parameter. Here are examples for these two cases: \ (*<*)experiment begin(*>*) attribute_setup my_rule = \Attrib.thms >> (fn ths => Thm.rule_attribute ths (fn context: Context.generic => fn th: thm => let val th' = th OF ths in th' end))\ attribute_setup my_declaration = \Attrib.thms >> (fn ths => Thm.declaration_attribute (fn th: thm => fn context: Context.generic => let val context' = context in context' end))\ (*<*)end(*>*) text \ \<^descr> @{attribute ML_print_depth} controls the printing depth of the ML toplevel pretty printer. Typically the limit should be less than 10. Bigger values such as 100--1000 are occasionally useful for debugging. \<^descr> @{attribute ML_source_trace} indicates whether the source text that is given to the ML compiler should be output: it shows the raw Standard ML after expansion of Isabelle/ML antiquotations. \<^descr> @{attribute ML_debugger} controls compilation of sources with or without debugging information. The global system option @{system_option_ref ML_debugger} does the same when building a session image. It is also possible use commands like \<^theory_text>\ML_file_debug\ etc. The ML debugger is - explained further in @{cite "isabelle-jedit"}. + explained further in \<^cite>\"isabelle-jedit"\. \<^descr> @{attribute ML_exception_trace} indicates whether the ML run-time system should print a detailed stack trace on exceptions. The result is dependent on various ML compiler optimizations. The boundary for the exception trace is the current Isar command transactions: it is occasionally better to insert the combinator \<^ML>\Runtime.exn_trace\ into ML code for debugging - @{cite "isabelle-implementation"}, closer to the point where it actually + \<^cite>\"isabelle-implementation"\, closer to the point where it actually happens. \<^descr> @{attribute ML_exception_debugger} controls detailed exception trace via the Poly/ML debugger, at the cost of extra compile-time and run-time overhead. Relevant ML modules need to be compiled beforehand with debugging enabled, see @{attribute ML_debugger} above. \<^descr> @{attribute ML_environment} determines the named ML environment for toplevel declarations, e.g.\ in command \<^theory_text>\ML\ or \<^theory_text>\ML_file\. The following ML environments are predefined in Isabelle/Pure: \<^item> \Isabelle\ for Isabelle/ML. It contains all modules of Isabelle/Pure and further add-ons, e.g. material from Isabelle/HOL. \<^item> \SML\ for official Standard ML. It contains only the initial basis according to \<^url>\http://sml-family.org/Basis/overview.html\. The Isabelle/ML function \<^ML>\ML_Env.setup\ defines a new ML environment. This is useful to incorporate big SML projects in an isolated name space, possibly with variations on ML syntax; the existing setup of \<^ML>\ML_Env.SML_operations\ follows the official standard. It is also possible to move toplevel bindings between ML environments, using a notation with ``\>\'' as separator. For example: \ (*<*)experiment begin(*>*) declare [[ML_environment = "Isabelle>SML"]] ML \val println = writeln\ declare [[ML_environment = "SML"]] ML \println "test"\ declare [[ML_environment = "Isabelle"]] ML \ML \println\ (*bad*) handle ERROR msg => warning msg\ (*<*)end(*>*) section \Generated files and exported files\ text \ Write access to the physical file-system is incompatible with the stateless model of processing Isabelle documents. To avoid bad effects, the following concepts for abstract file-management are provided by Isabelle: \<^descr>[Generated files] are stored within the theory context in Isabelle/ML. This allows to operate on the content in Isabelle/ML, e.g. via the command @{command compile_generated_files}. \<^descr>[Exported files] are stored within the session database in Isabelle/Scala. This allows to deliver artefacts to external tools, see - also @{cite "isabelle-system"} for session \<^verbatim>\ROOT\ declaration + also \<^cite>\"isabelle-system"\ for session \<^verbatim>\ROOT\ declaration \<^theory_text>\export_files\, and @{tool build} option \<^verbatim>\-e\. A notable example is the command @{command_ref export_code} (\chref{ch:export-code}): it uses both concepts simultaneously. File names are hierarchically structured, using a slash as separator. The (long) theory name is used as a prefix: the resulting name needs to be globally unique. \begin{matharray}{rcll} @{command_def "generate_file"} & : & \local_theory \ local_theory\ \\ @{command_def "export_generated_files"} & : & \context \\ \\ @{command_def "compile_generated_files"} & : & \context \\ \\ @{command_def "external_file"} & : & \any \ any\ \\ \end{matharray} \<^rail>\ @@{command generate_file} path '=' content ; path: @{syntax embedded} ; content: @{syntax embedded} ; @@{command export_generated_files} (files_in_theory + @'and') ; files_in_theory: (@'_' | (path+)) (('(' @'in' @{syntax name} ')')?) ; @@{command compile_generated_files} (files_in_theory + @'and') \ (@'external_files' (external_files + @'and'))? \ (@'export_files' (export_files + @'and'))? \ (@'export_prefix' path)? ; external_files: (path+) (('(' @'in' path ')')?) ; export_files: (path+) (executable?) ; executable: '(' ('exe' | 'executable') ')' ; @@{command external_file} @{syntax name} ';'? \ \<^descr> \<^theory_text>\generate_file path = content\ augments the table of generated files within the current theory by a new entry: duplicates are not allowed. The name extension determines a pre-existent file-type; the \content\ is a string that is preprocessed according to rules of this file-type. For example, Isabelle/Pure supports \<^verbatim>\.hs\ as file-type for Haskell: embedded cartouches are evaluated as Isabelle/ML expressions of type \<^ML_type>\string\, the result is inlined in Haskell string syntax. \<^descr> \<^theory_text>\export_generated_files paths (in thy)\ retrieves named generated files from the given theory (that needs be reachable via imports of the current one). By default, the current theory node is used. Using ``\<^verbatim>\_\'' (underscore) instead of explicit path names refers to \emph{all} files of a theory node. The overall list of files is prefixed with the respective (long) theory name and exported to the session database. In Isabelle/jEdit the result can be browsed via the virtual file-system with prefix ``\<^verbatim>\isabelle-export:\'' (using the regular file-browser). \<^descr> \<^theory_text>\scala_build_generated_files paths (in thy)\ retrieves named generated files as for \<^theory_text>\export_generated_files\ and writes them into a temporary directory, which is taken as starting point for build process of - Isabelle/Scala/Java modules (see @{cite "isabelle-system"}). The + Isabelle/Scala/Java modules (see \<^cite>\"isabelle-system"\). The corresponding @{path \build.props\} file is expected directly in the toplevel directory, instead of @{path \etc/build.props\} for Isabelle system components. These properties need to specify sources, resources, services etc. as usual. The resulting JAR module becomes an export artefact of the session database, with a name of the form ``\theory\\<^verbatim>\:classpath/\\module\\<^verbatim>\.jar\''. \<^descr> \<^theory_text>\compile_generated_files paths (in thy) where compile_body\ retrieves named generated files as for \<^theory_text>\export_generated_files\ and writes them into a temporary directory, such that the \compile_body\ may operate on them as an ML function of type \<^ML_type>\Path.T -> unit\. This may create further files, e.g.\ executables produced by a compiler that is invoked as external process (e.g.\ via \<^ML>\Isabelle_System.bash\), or any other files. The option ``\<^theory_text>\external_files paths (in base_dir)\'' copies files from the physical file-system into the temporary directory, \emph{before} invoking \compile_body\. The \base_dir\ prefix is removed from each of the \paths\, but the remaining sub-directory structure is reconstructed in the target directory. The option ``\<^theory_text>\export_files paths\'' exports the specified files from the temporary directory to the session database, \emph{after} invoking \compile_body\. Entries may be decorated with ``\<^theory_text>\(exe)\'' to say that it is a platform-specific executable program: the executable file-attribute will be set, and on Windows the \<^verbatim>\.exe\ file-extension will be included; ``\<^theory_text>\(executable)\'' only refers to the file-attribute, without special treatment of the \<^verbatim>\.exe\ extension. The option ``\<^theory_text>\export_prefix path\'' specifies an extra path prefix for all exports of \<^theory_text>\export_files\ above. \<^descr> \<^theory_text>\external_file name\ declares the formal dependency on the given file - name, such that the Isabelle build process knows about it (see also @{cite - "isabelle-system"}). This is required for any files mentioned in + name, such that the Isabelle build process knows about it (see also \<^cite>\"isabelle-system"\). This is required for any files mentioned in \<^theory_text>\compile_generated_files / external_files\ above, in order to document source dependencies properly. It is also possible to use \<^theory_text>\external_file\ alone, e.g.\ when other Isabelle/ML tools use \<^ML>\File.read\, without specific management of content by the Prover IDE. \ section \Primitive specification elements\ subsection \Sorts\ text \ \begin{matharray}{rcll} @{command_def "default_sort"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command default_sort} @{syntax sort} \ \<^descr> \<^theory_text>\default_sort s\ makes sort \s\ the new default sort for any type variable that is given explicitly in the text, but lacks a sort constraint (wrt.\ the current context). Type variables generated by type inference are not affected. Usually the default sort is only changed when defining a new object-logic. For example, the default sort in Isabelle/HOL is \<^class>\type\, the class of all HOL types. When merging theories, the default sorts of the parents are logically intersected, i.e.\ the representations as lists of classes are joined. \ subsection \Types \label{sec:types-pure}\ text \ \begin{matharray}{rcll} @{command_def "type_synonym"} & : & \local_theory \ local_theory\ \\ @{command_def "typedecl"} & : & \local_theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command type_synonym} (@{syntax typespec} '=' @{syntax type} @{syntax mixfix}?) ; @@{command typedecl} @{syntax typespec} @{syntax mixfix}? \ \<^descr> \<^theory_text>\type_synonym (\\<^sub>1, \, \\<^sub>n) t = \\ introduces a \<^emph>\type synonym\ \(\\<^sub>1, \, \\<^sub>n) t\ for the existing type \\\. Unlike the semantic type definitions in Isabelle/HOL, type synonyms are merely syntactic abbreviations without any logical significance. Internally, type synonyms are fully expanded. \<^descr> \<^theory_text>\typedecl (\\<^sub>1, \, \\<^sub>n) t\ declares a new type constructor \t\. If the object-logic defines a base sort \s\, then the constructor is declared to operate on that, via the axiomatic type-class instance \t :: (s, \, s)s\. \begin{warn} If you introduce a new type axiomatically, i.e.\ via @{command_ref typedecl} and @{command_ref axiomatization} (\secref{sec:axiomatizations}), the minimum requirement is that it has a non-empty model, to avoid immediate collapse of the logical environment. Moreover, one needs to demonstrate that the interpretation of such free-form axiomatizations can coexist with other axiomatization schemes for types, notably @{command_def typedef} in Isabelle/HOL (\secref{sec:hol-typedef}), or any other extension that people might have introduced elsewhere. \end{warn} \ section \Naming existing theorems \label{sec:theorems}\ text \ \begin{matharray}{rcll} @{command_def "lemmas"} & : & \local_theory \ local_theory\ \\ @{command_def "named_theorems"} & : & \local_theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command lemmas} (@{syntax thmdef}? @{syntax thms} + @'and') @{syntax for_fixes} ; @@{command named_theorems} (@{syntax name} @{syntax text}? + @'and') \ \<^descr> \<^theory_text>\lemmas a = b\<^sub>1 \ b\<^sub>n\~@{keyword_def "for"}~\x\<^sub>1 \ x\<^sub>m\ evaluates given facts (with attributes) in the current context, which may be augmented by local variables. Results are standardized before being stored, i.e.\ schematic variables are renamed to enforce index \0\ uniformly. \<^descr> \<^theory_text>\named_theorems name description\ declares a dynamic fact within the context. The same \name\ is used to define an attribute with the usual \add\/\del\ syntax (e.g.\ see \secref{sec:simp-rules}) to maintain the content incrementally, in canonical declaration order of the text structure. \ section \Oracles \label{sec:oracles}\ text \ \begin{matharray}{rcll} @{command_def "oracle"} & : & \theory \ theory\ & (axiomatic!) \\ @{command_def "thm_oracles"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} Oracles allow Isabelle to take advantage of external reasoners such as arithmetic decision procedures, model checkers, fast tautology checkers or computer algebra systems. Invoked as an oracle, an external reasoner can create arbitrary Isabelle theorems. It is the responsibility of the user to ensure that the external reasoner is as trustworthy as the application requires. Another typical source of errors is the linkup between Isabelle and the external tool, not just its concrete implementation, but also the required translation between two different logical environments. Isabelle merely guarantees well-formedness of the propositions being asserted, and records within the internal derivation object how presumed theorems depend on unproven suppositions. This also includes implicit type-class reasoning via the order-sorted algebra of class relations and type arities (see also @{command_ref instantiation} and @{command_ref instance}). \<^rail>\ @@{command oracle} @{syntax name} '=' @{syntax text} ; @@{command thm_oracles} @{syntax thms} \ \<^descr> \<^theory_text>\oracle name = "text"\ turns the given ML expression \text\ of type \<^ML_text>\'a -> cterm\ into an ML function of type \<^ML_text>\'a -> thm\, which is bound to the global identifier \<^ML_text>\name\. This acts like an infinitary specification of axioms! Invoking the oracle only works within the scope of the resulting theory. See \<^file>\~~/src/HOL/Examples/Iff_Oracle.thy\ for a worked example of defining a new primitive rule as oracle, and turning it into a proof method. \<^descr> \<^theory_text>\thm_oracles thms\ displays all oracles used in the internal derivation of the given theorems; this covers the full graph of transitive dependencies. \ section \Name spaces\ text \ \begin{matharray}{rcl} @{command_def "alias"} & : & \local_theory \ local_theory\ \\ @{command_def "type_alias"} & : & \local_theory \ local_theory\ \\ @{command_def "hide_class"} & : & \theory \ theory\ \\ @{command_def "hide_type"} & : & \theory \ theory\ \\ @{command_def "hide_const"} & : & \theory \ theory\ \\ @{command_def "hide_fact"} & : & \theory \ theory\ \\ \end{matharray} \<^rail>\ (@{command alias} | @{command type_alias}) @{syntax name} '=' @{syntax name} ; (@{command hide_class} | @{command hide_type} | @{command hide_const} | @{command hide_fact}) ('(' @'open' ')')? (@{syntax name} + ) \ Isabelle organizes any kind of name declarations (of types, constants, theorems etc.) by separate hierarchically structured name spaces. Normally the user does not have to control the behaviour of name spaces by hand, yet the following commands provide some way to do so. \<^descr> \<^theory_text>\alias\ and \<^theory_text>\type_alias\ introduce aliases for constants and type constructors, respectively. This allows adhoc changes to name-space accesses. \<^descr> \<^theory_text>\type_alias b = c\ introduces an alias for an existing type constructor. \<^descr> \<^theory_text>\hide_class names\ fully removes class declarations from a given name space; with the \(open)\ option, only the unqualified base name is hidden. Note that hiding name space accesses has no impact on logical declarations --- they remain valid internally. Entities that are no longer accessible to the user are printed with the special qualifier ``\??\'' prefixed to the full internal name. \<^descr> \<^theory_text>\hide_type\, \<^theory_text>\hide_const\, and \<^theory_text>\hide_fact\ are similar to \<^theory_text>\hide_class\, but hide types, constants, and facts, respectively. \ end diff --git a/src/Doc/JEdit/JEdit.thy b/src/Doc/JEdit/JEdit.thy --- a/src/Doc/JEdit/JEdit.thy +++ b/src/Doc/JEdit/JEdit.thy @@ -1,2246 +1,2236 @@ (*:maxLineLen=78:*) theory JEdit imports Base begin chapter \Introduction\ section \Concepts and terminology\ text \ Isabelle/jEdit is a Prover IDE that integrates \<^emph>\parallel proof checking\ - @{cite "Wenzel:2009" and "Wenzel:2013:ITP"} with \<^emph>\asynchronous user - interaction\ @{cite "Wenzel:2010" and "Wenzel:2012:UITP-EPTCS" and - "Wenzel:2014:ITP-PIDE" and "Wenzel:2014:UITP"}, based on a document-oriented - approach to \<^emph>\continuous proof processing\ @{cite "Wenzel:2011:CICM" and - "Wenzel:2012" and "Wenzel:2018:FIDE" and "Wenzel:2019:MKM"}. Many concepts + \<^cite>\"Wenzel:2009" and "Wenzel:2013:ITP"\ with \<^emph>\asynchronous user + interaction\ \<^cite>\"Wenzel:2010" and "Wenzel:2012:UITP-EPTCS" and + "Wenzel:2014:ITP-PIDE" and "Wenzel:2014:UITP"\, based on a document-oriented + approach to \<^emph>\continuous proof processing\ \<^cite>\"Wenzel:2011:CICM" and + "Wenzel:2012" and "Wenzel:2018:FIDE" and "Wenzel:2019:MKM"\. Many concepts and system components are fit together in order to make this work. The main building blocks are as follows. \<^descr>[Isabelle/ML] is the implementation and extension language of Isabelle, - see also @{cite "isabelle-implementation"}. It is integrated into the + see also \<^cite>\"isabelle-implementation"\. It is integrated into the logical context of Isabelle/Isar and allows to manipulate logical entities directly. Arbitrary add-on tools may be implemented for object-logics such as Isabelle/HOL. \<^descr>[Isabelle/Scala] is the system programming language of Isabelle. It extends the pure logical environment of Isabelle/ML towards the outer world of graphical user interfaces, text editors, IDE frameworks, web services, SSH servers, SQL databases etc. Both Scala and ML provide library modules to support formatted text with formal markup, and to encode/decode algebraic datatypes. Scala communicates with ML via asynchronous protocol commands; from the ML perspective this is wrapped up as synchronous function call (RPC). \<^descr>[PIDE] is a general framework for Prover IDEs based on Isabelle/Scala. It is built around a concept of parallel and asynchronous document processing, which is supported natively by the parallel proof engine that is implemented in Isabelle/ML. The traditional prover command loop is given up; instead there is direct support for editing of source text, with rich formal markup for GUI rendering. \<^descr>[jEdit] is a sophisticated text editor\<^footnote>\\<^url>\http://www.jedit.org\\ implemented in Java\<^footnote>\\<^url>\https://openjdk.java.net\\. The editor is easily extensible by plugins written in any language that works on the JVM. In the context of Isabelle this is usually Scala\<^footnote>\\<^url>\https://www.scala-lang.org\\. \<^descr>[Isabelle/jEdit] is the main application of the PIDE framework and the default user-interface for Isabelle. It targets both beginners and experts. Technically, Isabelle/jEdit consists of the original jEdit code base with minimal patches and a special plugin for Isabelle. This is integrated as a desktop application for the main operating system families: Linux, Windows, macOS. End-users of Isabelle download and run a standalone application that exposes jEdit as a text editor on the surface. Thus there is occasionally a tendency to apply the name ``jEdit'' to any of the Isabelle Prover IDE aspects, without proper differentiation. When discussing these PIDE building blocks in public forums, mailing lists, or even scientific publications, it is particularly important to distinguish Isabelle/ML versus Standard ML, Isabelle/Scala versus Scala, Isabelle/jEdit versus jEdit. \ section \The Isabelle/jEdit Prover IDE\ text \ \begin{figure}[!htb] \begin{center} \includegraphics[width=\textwidth]{isabelle-jedit} \end{center} \caption{The Isabelle/jEdit Prover IDE} \label{fig:isabelle-jedit} \end{figure} Isabelle/jEdit (\figref{fig:isabelle-jedit}) consists of some plugins for the jEdit text editor, while preserving its overall look-and-feel. The main plugin is called ``Isabelle'' and has its own menu \<^emph>\Plugins~/ Isabelle\ with access to several actions and add-on panels (see also \secref{sec:dockables}), as well as \<^emph>\Plugins~/ Plugin Options~/ Isabelle\ (see also \secref{sec:options}). The options allow to specify a logic session name, but the same selector is also accessible in the \<^emph>\Theories\ panel (\secref{sec:theories}). After startup of the Isabelle plugin, the selected logic session image is provided - automatically by the Isabelle build tool @{cite "isabelle-system"}: if it is + automatically by the Isabelle build tool \<^cite>\"isabelle-system"\: if it is absent or outdated wrt.\ its sources, the build process updates it within the running text editor. Prover IDE functionality is fully activated after successful termination of the build process. A failure may require changing some options and restart of the Isabelle plugin or application. Changing the logic session requires a restart of the whole application to take effect. \<^medskip> The main job of the Prover IDE is to manage sources and their changes, taking the logical structure as a formal document into account (see also \secref{sec:document-model}). The editor and the prover are connected asynchronously without locking. The prover is free to organize the checking of the formal text in parallel on multiple cores, and provides feedback via markup, which is rendered in the editor via colors, boxes, squiggly underlines, hyperlinks, popup windows, icons, clickable output etc. Using the mouse together with the modifier key \<^verbatim>\CONTROL\ (Linux, Windows) or \<^verbatim>\COMMAND\ (macOS) exposes formal content via tooltips and/or hyperlinks (see also \secref{sec:tooltips-hyperlinks}). Output (in popups etc.) may be explored recursively, using the same techniques as in the editor source buffer. Thus the Prover IDE gives an impression of direct access to formal content of the prover within the editor, but in reality only certain aspects are exposed, according to the possibilities of the prover and its add-on tools. \ subsection \Documentation\ text \ The \<^emph>\Documentation\ panel of Isabelle/jEdit provides access to some example theory files and the standard Isabelle documentation. PDF files are opened by regular desktop operations of the underlying platform. The section ``Original jEdit Documentation'' contains the original \<^emph>\User's Guide\ of this sophisticated text editor. The same is accessible via the \<^verbatim>\Help\ menu or \<^verbatim>\F1\ keyboard shortcut, using the built-in HTML viewer of Java/Swing. The latter also includes \<^emph>\Frequently Asked Questions\ and documentation of individual plugins. Most of the information about jEdit is relevant for Isabelle/jEdit as well, but users need to keep in mind that defaults sometimes differ, and the official jEdit documentation does not know about the Isabelle plugin with its support for continuous checking of formal source text: jEdit is a plain text editor, but Isabelle/jEdit is a Prover IDE. \ subsection \Plugins\ text \ The \<^emph>\Plugin Manager\ of jEdit allows to augment editor functionality by JVM modules (jars) that are provided by the central plugin repository, which is accessible via various mirror sites. Connecting to the plugin server-infrastructure of the jEdit project allows to update bundled plugins or to add further functionality. This needs to be done with the usual care for such an open bazaar of contributions. Arbitrary combinations of add-on features are apt to cause problems. It is advisable to start with the default configuration of Isabelle/jEdit and develop a sense how it is meant to work, before loading too many other plugins. \<^medskip> The \<^emph>\Isabelle\ plugin is responsible for the main Prover IDE functionality of Isabelle/jEdit: it manages the prover session in the background. A few additional plugins are bundled with Isabelle/jEdit for convenience or out of necessity, notably \<^emph>\Console\ with its \<^emph>\Scala\ sub-plugin (\secref{sec:scala-console}) and \<^emph>\SideKick\ with some Isabelle-specific parsers for document tree structure (\secref{sec:sidekick}). The \<^emph>\Navigator\ plugin is particularly important for hyperlinks within the formal document-model (\secref{sec:tooltips-hyperlinks}). Further plugins (e.g.\ \<^emph>\ErrorList\, \<^emph>\Code2HTML\) are included to saturate the dependencies of bundled plugins, but have no particular use in Isabelle/jEdit. \ subsection \Options \label{sec:options}\ text \ Both jEdit and Isabelle have distinctive management of persistent options. Regular jEdit options are accessible via the dialogs \<^emph>\Utilities~/ Global Options\ or \<^emph>\Plugins~/ Plugin Options\, with a second chance to flip the two within the central options dialog. Changes are stored in \<^path>\$JEDIT_SETTINGS/properties\ and \<^path>\$JEDIT_SETTINGS/keymaps\. Isabelle system options are managed by Isabelle/Scala and changes are stored in \<^path>\$ISABELLE_HOME_USER/etc/preferences\, independently of - other jEdit properties. See also @{cite "isabelle-system"}, especially the + other jEdit properties. See also \<^cite>\"isabelle-system"\, especially the coverage of sessions and command-line tools like @{tool build} or @{tool options}. Those Isabelle options that are declared as \<^verbatim>\public\ are configurable in Isabelle/jEdit via \<^emph>\Plugin Options~/ Isabelle~/ General\. Moreover, there are various options for rendering document content, which are configurable via \<^emph>\Plugin Options~/ Isabelle~/ Rendering\. Thus \<^emph>\Plugin Options~/ Isabelle\ in jEdit provides a view on a subset of Isabelle system options. Note that some of these options affect general parameters that are relevant outside Isabelle/jEdit as well, e.g.\ @{system_option threads} or - @{system_option parallel_proofs} for the Isabelle build tool @{cite - "isabelle-system"}, but it is possible to use the settings variable + @{system_option parallel_proofs} for the Isabelle build tool \<^cite>\"isabelle-system"\, but it is possible to use the settings variable @{setting ISABELLE_BUILD_OPTIONS} to change defaults for batch builds without affecting the Prover IDE. The jEdit action @{action_def isabelle.options} opens the options dialog for the Isabelle plugin; it can be mapped to editor GUI elements as usual. \<^medskip> Options are usually loaded on startup and saved on shutdown of Isabelle/jEdit. Editing the generated \<^path>\$JEDIT_SETTINGS/properties\ or \<^path>\$ISABELLE_HOME_USER/etc/preferences\ manually while the application is running may cause lost updates! \ subsection \Keymaps\ text \ Keyboard shortcuts are managed as a separate concept of \<^emph>\keymap\ that is configurable via \<^emph>\Global Options~/ Shortcuts\. The \<^verbatim>\imported\ keymap is derived from the initial environment of properties that is available at the first start of the editor; afterwards the keymap file takes precedence and is no longer affected by change of default properties. Users may modify their keymap later, but this can lead to conflicts with \<^verbatim>\shortcut\ properties in \<^file>\$JEDIT_HOME/properties/jEdit.props\. The action @{action_def "isabelle.keymap-merge"} helps to resolve pending Isabelle keymap changes wrt. the current jEdit keymap; non-conflicting changes are applied implicitly. This action is automatically invoked on Isabelle/jEdit startup. \ section \Command-line invocation \label{sec:command-line}\ text \ Isabelle/jEdit is normally invoked as a single-instance desktop application, based on platform-specific executables for Linux, Windows, macOS. It is also possible to invoke the Prover IDE on the command-line, with some extra options and environment settings. The command-line usage of @{tool_def jedit} is as follows: @{verbatim [display] \Usage: isabelle jedit [OPTIONS] [FILES ...] Options are: -A NAME ancestor session for option -R (default: parent) -D NAME=X set JVM system property -J OPTION add JVM runtime option (default $JEDIT_JAVA_SYSTEM_OPTIONS $JEDIT_JAVA_OPTIONS) -R NAME build image with requirements from other sessions -b build only -d DIR include session directory -f fresh build -i NAME include session in name-space of theories -j OPTION add jEdit runtime option (default $JEDIT_OPTIONS) -l NAME logic image name -m MODE add print mode for output -n no build of session image on startup -p CMD ML process command prefix (process policy) -s system build mode for session image (system_heaps=true) -u user build mode for session image (system_heaps=false) Start jEdit with Isabelle plugin setup and open FILES (default "$USER_HOME/Scratch.thy" or ":" for empty buffer).\} The \<^verbatim>\-l\ option specifies the session name of the logic image to be used for proof processing. Additional session root directories may be included - via option \<^verbatim>\-d\ to augment the session name space (see also @{cite - "isabelle-system"}). By default, the specified image is checked and built on + via option \<^verbatim>\-d\ to augment the session name space (see also \<^cite>\"isabelle-system"\). By default, the specified image is checked and built on demand, but option \<^verbatim>\-n\ bypasses the implicit build process for the selected session image. Options \<^verbatim>\-s\ and \<^verbatim>\-u\ override the default system option @{system_option system_heaps}: this determines where to store the session image of @{tool build}. The \<^verbatim>\-R\ option builds an auxiliary logic image with all theories from other sessions that are not already present in its parent; it also opens the session \<^verbatim>\ROOT\ entry in the editor to facilitate editing of the main session. The \<^verbatim>\-A\ option specifies and alternative ancestor session for option \<^verbatim>\-R\: this allows to restructure the hierarchy of session images on the spot. Options \<^verbatim>\-R\ and \<^verbatim>\-l\ are mutually exclusive. The \<^verbatim>\-i\ option includes additional sessions into the name-space of theories: multiple occurrences are possible. The \<^verbatim>\-m\ option specifies additional print modes for the prover process. Note that the system option @{system_option_ref jedit_print_mode} allows to do the same persistently (e.g.\ via the \<^emph>\Plugin Options\ dialog of Isabelle/jEdit), without requiring command-line invocation. The \<^verbatim>\-J\ and \<^verbatim>\-j\ options pass additional low-level options to the JVM or jEdit, respectively. The defaults are provided by the Isabelle settings - environment @{cite "isabelle-system"}, but note that these only work for the + environment \<^cite>\"isabelle-system"\, but note that these only work for the command-line tool described here, and not the desktop application. The \<^verbatim>\-D\ option allows to define JVM system properties; this is passed directly to the underlying \<^verbatim>\java\ process. The \<^verbatim>\-b\ and \<^verbatim>\-f\ options control the self-build mechanism of Isabelle/Scala/PIDE/jEdit. This is only relevant for building from sources, the official Isabelle release already includes a pre-built version of Isabelle/jEdit. \<^bigskip> It is also possible to connect to an already running Isabelle/jEdit process via @{tool_def jedit_client}: @{verbatim [display] \Usage: isabelle jedit_client [OPTIONS] [FILES ...] Options are: -c only check presence of server -n only report server name -s NAME server name (default "Isabelle") Connect to already running Isabelle/jEdit instance and open FILES\} The \<^verbatim>\-c\ option merely checks the presence of the server, producing a process return-code. The \<^verbatim>\-n\ option reports the server name, and the \<^verbatim>\-s\ option provides a different server name. The default server name is the official distribution name (e.g.\ \<^verbatim>\Isabelle2022\). Thus @{tool jedit_client} can connect to the Isabelle desktop application without further options. The \<^verbatim>\-p\ option allows to override the implicit default of the system option @{system_option_ref ML_process_policy} for ML processes started by the Prover IDE, e.g. to control CPU affinity on multiprocessor systems. The JVM system property \<^verbatim>\isabelle.jedit_server\ provides a different server name, e.g.\ use \<^verbatim>\isabelle jedit -Disabelle.jedit_server=\\name\ and \<^verbatim>\isabelle jedit_client -s\~\name\ to connect later on. \ section \GUI rendering\ text \ Isabelle/jEdit is a classic Java/AWT/Swing application: its GUI rendering usually works well, but there are technical side-conditions on the Java window system and graphics engine. When researching problems and solutions on the Web, it often helps to include other well-known Swing applications, notably IntelliJ IDEA and Netbeans. \ subsection \Portable and scalable look-and-feel\ text \ In the past, \<^emph>\system look-and-feels\ tried hard to imitate native GUI elements on specific platforms (Windows, macOS/Aqua, Linux/GTK+), but many technical problems have accumulated in recent years (e.g.\ see \secref{sec:problems}). In 2021, we are de-facto back to \<^emph>\portable look-and-feels\, which also happen to be \emph{scalable} on high-resolution displays: \<^item> \<^verbatim>\FlatLaf Light\ is the default for Isabelle/jEdit on all platforms. It generally looks good and adapts itself pretty well to high-resolution displays. \<^item> \<^verbatim>\FlatLaf Dark\ is an alternative, but it requires further changes of editor colors by the user (or by the jEdit plugin \<^verbatim>\Editor Scheme\). Also note that Isabelle/PIDE has its own extensive set of rendering options that need to be revisited. \<^item> \<^verbatim>\Metal\ still works smoothly, although it is stylistically outdated. It can accommodate high-resolution displays via font properties (see below). Changing the look-and-feel in \<^emph>\Global Options~/ Appearance\ often updates the GUI only partially: a full restart of Isabelle/jEdit is required to see the true effect. \ subsection \Adjusting fonts\ text \ The preferred font family for Isabelle/jEdit is \<^verbatim>\Isabelle DejaVu\: it is used by default for the main text area and various GUI elements. The default font sizes attempt to deliver a usable application for common display types, such as ``Full HD'' at $1920 \times 1080$ and ``Ultra HD'' at $3840 \times 2160$. \<^medskip> Isabelle/jEdit provides various options to adjust font sizes in particular GUI elements. Here is a summary of all relevant font properties: \<^item> \<^emph>\Global Options / Text Area / Text font\: the main text area font, which is also used as reference point for various derived font sizes, e.g.\ the \<^emph>\Output\ (\secref{sec:output}) and \<^emph>\State\ (\secref{sec:state-output}) panels. \<^item> \<^emph>\Global Options / Gutter / Gutter font\: the font for the gutter area left of the main text area, e.g.\ relevant for display of line numbers (disabled by default). \<^item> \<^emph>\Global Options / Appearance / Button, menu and label font\ as well as \<^emph>\List and text field font\: this specifies the primary and secondary font for the \<^emph>\Metal\ look-and-feel. \<^item> \<^emph>\Plugin Options / Isabelle / General / Reset Font Size\: the main text area font size for action @{action_ref "isabelle.reset-font-size"}, e.g.\ relevant for quick scaling like in common web browsers. \<^item> \<^emph>\Plugin Options / Console / General / Font\: the console window font, e.g.\ relevant for Isabelle/Scala command-line. \ chapter \Augmented jEdit functionality\ section \Dockable windows \label{sec:dockables}\ text \ In jEdit terminology, a \<^emph>\view\ is an editor window with one or more \<^emph>\text areas\ that show the content of one or more \<^emph>\buffers\. A regular view may be surrounded by \<^emph>\dockable windows\ that show additional information in arbitrary format, not just text; a \<^emph>\plain view\ does not allow dockables. The \<^emph>\dockable window manager\ of jEdit organizes these dockable windows, either as \<^emph>\floating\ windows, or \<^emph>\docked\ panels within one of the four margins of the view. There may be any number of floating instances of some dockable window, but at most one docked instance; jEdit actions that address \<^emph>\the\ dockable window of a particular kind refer to the unique docked instance. Dockables are used routinely in jEdit for important functionality like \<^emph>\HyperSearch Results\ or the \<^emph>\File System Browser\. Plugins often provide a central dockable to access their main functionality, which may be opened by the user on demand. The Isabelle/jEdit plugin takes this approach to the extreme: its plugin menu provides the entry-points to many panels that are managed as dockable windows. Some important panels are docked by default, e.g.\ \<^emph>\Documentation\, \<^emph>\State\, \<^emph>\Theories\ \<^emph>\Output\, \<^emph>\Query\. The user can change this arrangement easily and persistently. Compared to plain jEdit, dockable window management in Isabelle/jEdit is slightly augmented according to the the following principles: \<^item> Floating windows are dependent on the main window as \<^emph>\dialog\ in the sense of Java/AWT/Swing. Dialog windows always stay on top of the view, which is particularly important in full-screen mode. The desktop environment of the underlying platform may impose further policies on such dependent dialogs, in contrast to fully independent windows, e.g.\ some window management functions may be missing. \<^item> Keyboard focus of the main view vs.\ a dockable window is carefully managed according to the intended semantics, as a panel mainly for output or input. For example, activating the \<^emph>\Output\ (\secref{sec:output}) or State (\secref{sec:state-output}) panel via the dockable window manager returns keyboard focus to the main text area, but for \<^emph>\Query\ (\secref{sec:query}) or \<^emph>\Sledgehammer\ \secref{sec:sledgehammer} the focus is given to the main input field of that panel. \<^item> Panels that provide their own text area for output have an additional dockable menu item \<^emph>\Detach\. This produces an independent copy of the current output as a floating \<^emph>\Info\ window, which displays that content independently of ongoing changes of the PIDE document-model. Note that Isabelle/jEdit popup windows (\secref{sec:tooltips-hyperlinks}) provide a similar \<^emph>\Detach\ operation as an icon. \ section \Isabelle symbols \label{sec:symbols}\ text \ Isabelle sources consist of \<^emph>\symbols\ that extend plain ASCII to allow infinitely many mathematical symbols within the formal sources. This works without depending on particular encodings and varying Unicode standards.\<^footnote>\Raw Unicode characters within formal sources compromise portability and reliability in the face of changing interpretation of special features of Unicode, such as Combining Characters or Bi-directional - Text.\ See @{cite "Wenzel:2011:CICM"}. + Text.\ See \<^cite>\"Wenzel:2011:CICM"\. For the prover back-end, formal text consists of ASCII characters that are grouped according to some simple rules, e.g.\ as plain ``\<^verbatim>\a\'' or symbolic ``\<^verbatim>\\\''. For the editor front-end, a certain subset of symbols is rendered physically via Unicode glyphs, to show ``\<^verbatim>\\\'' as ``\\\'', for example. This symbol interpretation is specified by the Isabelle system distribution in \<^file>\$ISABELLE_HOME/etc/symbols\ and may be augmented by the user in \<^path>\$ISABELLE_HOME_USER/etc/symbols\. - The appendix of @{cite "isabelle-isar-ref"} gives an overview of the + The appendix of \<^cite>\"isabelle-isar-ref"\ gives an overview of the standard interpretation of finitely many symbols from the infinite collection. Uninterpreted symbols are displayed literally, e.g.\ ``\<^verbatim>\\\''. Overlap of Unicode characters used in symbol interpretation with informal ones (which might appear e.g.\ in comments) needs to be avoided. Raw Unicode characters within prover source files should be restricted to informal parts, e.g.\ to write text in non-latin alphabets in comments. \ paragraph \Encoding.\ text \Technically, the Unicode interpretation of Isabelle symbols is an \<^emph>\encoding\ called \<^verbatim>\UTF-8-Isabelle\ in jEdit (\<^emph>\not\ in the underlying JVM). It is provided by the Isabelle Base plugin and enabled by default for all source files in Isabelle/jEdit. Sometimes such defaults are reset accidentally, or malformed UTF-8 sequences in the text force jEdit to fall back on a different encoding like \<^verbatim>\ISO-8859-15\. In that case, verbatim ``\<^verbatim>\\\'' will be shown in the text buffer instead of its Unicode rendering ``\\\''. The jEdit menu operation \<^emph>\File~/ Reload with Encoding~/ UTF-8-Isabelle\ helps to resolve such problems (after repairing malformed parts of the text). If the loaded text already contains Unicode sequences that are in conflict with the Isabelle symbol encoding, the fallback-encoding UTF-8 is used and Isabelle symbols remain in literal \<^verbatim>\\\ form. The jEdit menu operation \<^emph>\Utilities~/ Buffer Options~/ Character encoding\ allows to enforce \<^verbatim>\UTF-8-Isabelle\, but this will also change original Unicode text into Isabelle symbols when saving the file! \ paragraph \Font.\ text \Correct rendering via Unicode requires a font that contains glyphs for the corresponding codepoints. There are also various unusual symbols with particular purpose in Isabelle, e.g.\ control symbols and very long arrows. Isabelle/jEdit prefers its own font collection \<^verbatim>\Isabelle DejaVu\, with families \<^verbatim>\Serif\ (default for help texts), \<^verbatim>\Sans\ (default for GUI elements), \<^verbatim>\Mono Sans\ (default for text area). This ensures that all standard Isabelle symbols are shown on the screen (or printer) as expected. Note that a Java/AWT/Swing application can load additional fonts only if they are not installed on the operating system already! Outdated versions of Isabelle fonts that happen to be provided by the operating system prevent Isabelle/jEdit to use its bundled version. This could lead to missing glyphs (black rectangles), when the system version of a font is older than the application version. This problem can be avoided by refraining to ``install'' a copy of the Isabelle fonts in the first place, although it might be tempting to use the same font in other applications. HTML pages generated by Isabelle refer to the same Isabelle fonts as a server-side resource. Thus a web-browser can use that without requiring a locally installed copy. \ paragraph \Input methods.\ text \In principle, Isabelle/jEdit could delegate the problem to produce Isabelle symbols in their Unicode rendering to the underlying operating system and its \<^emph>\input methods\. Regular jEdit also provides various ways to work with \<^emph>\abbreviations\ to produce certain non-ASCII characters. Since none of these standard input methods work satisfactorily for the mathematical characters required for Isabelle, various specific Isabelle/jEdit mechanisms are provided. This is a summary for practically relevant input methods for Isabelle symbols. \<^enum> The \<^emph>\Symbols\ panel: some GUI buttons allow to insert certain symbols in the text buffer. There are also tooltips to reveal the official Isabelle representation with some additional information about \<^emph>\symbol abbreviations\ (see below). \<^enum> Copy/paste from decoded source files: text that is already rendered as Unicode can be re-used for other text. This also works between different applications, e.g.\ Isabelle/jEdit and some web browser or mail client, as long as the same Unicode interpretation of Isabelle symbols is used. \<^enum> Copy/paste from prover output within Isabelle/jEdit. The same principles as for text buffers apply, but note that \<^emph>\copy\ in secondary Isabelle/jEdit windows works via the keyboard shortcuts \<^verbatim>\C+c\ or \<^verbatim>\C+INSERT\, while jEdit menu actions always refer to the primary text area! \<^enum> Completion provided by the Isabelle plugin (see \secref{sec:completion}). Isabelle symbols have a canonical name and optional abbreviations. This can be used with the text completion mechanism of Isabelle/jEdit, to replace a prefix of the actual symbol like \<^verbatim>\\\, or its name preceded by backslash \<^verbatim>\\lambda\, or its ASCII abbreviation \<^verbatim>\%\ by the Unicode rendering. The following table is an extract of the information provided by the standard \<^file>\$ISABELLE_HOME/etc/symbols\ file: \<^medskip> \begin{tabular}{lll} \<^bold>\symbol\ & \<^bold>\name with backslash\ & \<^bold>\abbreviation\ \\\hline \\\ & \<^verbatim>\\lambda\ & \<^verbatim>\%\ \\ \\\ & \<^verbatim>\\Rightarrow\ & \<^verbatim>\=>\ \\ \\\ & \<^verbatim>\\Longrightarrow\ & \<^verbatim>\==>\ \\[0.5ex] \\\ & \<^verbatim>\\And\ & \<^verbatim>\!!\ \\ \\\ & \<^verbatim>\\equiv\ & \<^verbatim>\==\ \\[0.5ex] \\\ & \<^verbatim>\\forall\ & \<^verbatim>\!\ \\ \\\ & \<^verbatim>\\exists\ & \<^verbatim>\?\ \\ \\\ & \<^verbatim>\\longrightarrow\ & \<^verbatim>\-->\ \\ \\\ & \<^verbatim>\\and\ & \<^verbatim>\&\ \\ \\\ & \<^verbatim>\\or\ & \<^verbatim>\|\ \\ \\\ & \<^verbatim>\\not\ & \<^verbatim>\~\ \\ \\\ & \<^verbatim>\\noteq\ & \<^verbatim>\~=\ \\ \\\ & \<^verbatim>\\in\ & \<^verbatim>\:\ \\ \\\ & \<^verbatim>\\notin\ & \<^verbatim>\~:\ \\ \end{tabular} \<^medskip> Note that the above abbreviations refer to the input method. The logical notation provides ASCII alternatives that often coincide, but sometimes deviate. This occasionally causes user confusion with old-fashioned Isabelle source that use ASCII replacement notation like \<^verbatim>\!\ or \<^verbatim>\ALL\ directly in the text. On the other hand, coincidence of symbol abbreviations with ASCII replacement syntax syntax helps to update old theory sources via explicit completion (see also \<^verbatim>\C+b\ explained in \secref{sec:completion}). \ paragraph \Control symbols.\ text \There are some special control symbols to modify the display style of a single symbol (without nesting). Control symbols may be applied to a region of selected text, either using the \<^emph>\Symbols\ panel or keyboard shortcuts or jEdit actions. These editor operations produce a separate control symbol for each symbol in the text, in order to make the whole text appear in a certain style. \<^medskip> \begin{tabular}{llll} \<^bold>\style\ & \<^bold>\symbol\ & \<^bold>\shortcut\ & \<^bold>\action\ \\\hline superscript & \<^verbatim>\\<^sup>\ & \<^verbatim>\C+e UP\ & @{action_ref "isabelle.control-sup"} \\ subscript & \<^verbatim>\\<^sub>\ & \<^verbatim>\C+e DOWN\ & @{action_ref "isabelle.control-sub"} \\ bold face & \<^verbatim>\\<^bold>\ & \<^verbatim>\C+e RIGHT\ & @{action_ref "isabelle.control-bold"} \\ emphasized & \<^verbatim>\\<^emph>\ & \<^verbatim>\C+e LEFT\ & @{action_ref "isabelle.control-emph"} \\ reset & & \<^verbatim>\C+e BACK_SPACE\ & @{action_ref "isabelle.control-reset"} \\ \end{tabular} \<^medskip> To produce a single control symbol, it is also possible to complete on \<^verbatim>\\sup\, \<^verbatim>\\sub\, \<^verbatim>\\bold\, \<^verbatim>\\emph\ as for regular symbols. The emphasized style only takes effect in document output (when used with a cartouche), but not in the editor. \ section \Scala console \label{sec:scala-console}\ text \ The \<^emph>\Console\ plugin manages various shells (command interpreters), e.g.\ \<^emph>\BeanShell\, which is the official jEdit scripting language, and the cross-platform \<^emph>\System\ shell. Thus the console provides similar functionality than the Emacs buffers \<^verbatim>\*scratch*\ and \<^verbatim>\*shell*\. Isabelle/jEdit extends the repertoire of the console by \<^emph>\Scala\, which is the regular Scala toplevel loop running inside the same JVM process as Isabelle/jEdit itself. This means the Scala command interpreter has access to the JVM name space and state of the running Prover IDE application. The default environment imports the full content of packages \<^verbatim>\isabelle\ and \<^verbatim>\isabelle.jedit\. For example, \<^verbatim>\PIDE\ refers to the Isabelle/jEdit plugin object, and \<^verbatim>\view\ to the current editor view of jEdit. The Scala expression \<^verbatim>\PIDE.snapshot(view)\ makes a PIDE document snapshot of the current buffer within the current editor view: it allows to retrieve document markup in a timeless~/ stateless manner, while the prover continues its processing. This helps to explore Isabelle/Scala functionality interactively. Some care is required to avoid interference with the internals of the running application. \ section \Physical and logical files \label{sec:files}\ text \ File specifications in jEdit follow various formats and conventions according to \<^emph>\Virtual File Systems\, which may be also provided by plugins. This allows to access remote files via the \<^verbatim>\https:\ protocol prefix, for example. Isabelle/jEdit attempts to work with the file-system model of jEdit as far as possible. In particular, theory sources are passed from the editor to the prover, without indirection via the file-system. Thus files don't need to be saved: the editor buffer content is used directly. \ subsection \Local files and environment variables \label{sec:local-files}\ text \ Local files (without URL notation) are particularly important. The file path notation is that of the Java Virtual Machine on the underlying platform. On Windows the preferred form uses backslashes, but happens to accept forward slashes like Unix/POSIX as well. Further differences arise due to Windows drive letters and network shares: thus relative paths (with forward slashes) are portable, but absolute paths are not. File paths in Java are distinct from Isabelle; the latter uses POSIX notation with forward slashes on \<^emph>\all\ platforms. Isabelle/ML on Windows uses Unix-style path notation, with drive letters according to Cygwin (e.g.\ \<^verbatim>\/cygdrive/c\). Environment variables from the Isabelle process may be used freely, e.g.\ \<^file>\$ISABELLE_HOME/etc/symbols\ or \<^file>\$POLYML_HOME/README\. There are special shortcuts: \<^dir>\~\ for \<^dir>\$USER_HOME\ and \<^dir>\~~\ for \<^dir>\$ISABELLE_HOME\. \<^medskip> Since jEdit happens to support environment variables within file specifications as well, it is natural to use similar notation within the editor, e.g.\ in the file-browser. This does not work in full generality, though, due to the bias of jEdit towards platform-specific notation and of Isabelle towards POSIX. Moreover, the Isabelle settings environment is not accessible when starting Isabelle/jEdit via the desktop application wrapper, in contrast to @{tool jedit} run from the command line (\secref{sec:command-line}). Isabelle/jEdit imitates important system settings within the Java process environment, in order to allow easy access to these important places from the editor: \<^verbatim>\$ISABELLE_HOME\, \<^verbatim>\$ISABELLE_HOME_USER\, \<^verbatim>\$JEDIT_HOME\, \<^verbatim>\$JEDIT_SETTINGS\. The file browser of jEdit also includes \<^emph>\Favorites\ for these locations. \<^medskip> Path specifications in prover input or output usually include formal markup that turns it into a hyperlink (see also \secref{sec:tooltips-hyperlinks}). This allows to open the corresponding file in the text editor, independently of the path notation. If the path refers to a directory, it is opened in the jEdit file browser. Formally checked paths in prover input are subject to completion (\secref{sec:completion}): partial specifications are resolved via directory content and possible completions are offered in a popup. \ subsection \PIDE resources via virtual file-systems\ text \ The jEdit file browser is docked by default. It provides immediate access to the local file-system, as well as important Isabelle resources via the \<^emph>\Favorites\ menu. Environment variables like \<^verbatim>\$ISABELLE_HOME\ are discussed in \secref{sec:local-files}. Virtual file-systems are more special: the idea is to present structured information like a directory tree. The following URLs are offered in the \<^emph>\Favorites\ menu, or by corresponding jEdit actions. \<^item> URL \<^verbatim>\isabelle-export:\ or action @{action_def "isabelle-export-browser"} shows a toplevel directory with theory names: each may provide its own tree structure of session exports. Exports are like a logical file-system for the current prover session, maintained as Isabelle/Scala data structures and written to the session database eventually. The \<^verbatim>\isabelle-export:\ URL exposes the current content according to a snapshot of the document model. The file browser is \<^emph>\not\ updated continuously when the PIDE document changes: the reload operation needs to be used explicitly. A notable example for exports is the command - @{command_ref export_code} @{cite "isabelle-isar-ref"}. + @{command_ref export_code} \<^cite>\"isabelle-isar-ref"\. \<^item> URL \<^verbatim>\isabelle-session:\ or action @{action_def "isabelle-session-browser"} show the structure of session chapters and sessions within them. What looks like a file-entry is actually a reference to the session definition in its corresponding \<^verbatim>\ROOT\ file. The latter is subject to Prover IDE markup, so the session theories and other files may be browsed quickly by following hyperlinks in the text. \ section \Indentation\ text \ Isabelle/jEdit augments the existing indentation facilities of jEdit to take the structure of theory and proof texts into account. There is also special support for unstructured proof scripts (\<^theory_text>\apply\ etc.). \<^descr>[Syntactic indentation] follows the outer syntax of Isabelle/Isar. Action @{action "indent-lines"} (shortcut \<^verbatim>\C+i\) indents the current line according to command keywords and some command substructure: this approximation may need further manual tuning. Action @{action "isabelle.newline"} (shortcut \<^verbatim>\ENTER\) indents the old and the new line according to command keywords only: leading to precise alignment of the main Isar language elements. This depends on option @{system_option_def "jedit_indent_newline"} (enabled by default). Regular input (via keyboard or completion) indents the current line whenever an new keyword is emerging at the start of the line. This depends on option @{system_option_def "jedit_indent_input"} (enabled by default). \<^descr>[Semantic indentation] adds additional white space to unstructured proof scripts via the number of subgoals. This requires information of ongoing document processing and may thus lag behind when the user is editing too quickly; see also option @{system_option_def "jedit_script_indent"} and @{system_option_def "jedit_script_indent_limit"}. The above options are accessible in the menu \<^emph>\Plugins / Plugin Options / Isabelle / General\. A prerequisite for advanced indentation is \<^emph>\Utilities / Buffer Options / Automatic indentation\: it needs to be set to \<^verbatim>\full\ (default). \ section \SideKick parsers \label{sec:sidekick}\ text \ The \<^emph>\SideKick\ plugin provides some general services to display buffer structure in a tree view. Isabelle/jEdit provides SideKick parsers for its main mode for theory files, ML files, as well as some minor modes for the \<^verbatim>\NEWS\ file (see \figref{fig:sidekick}), session \<^verbatim>\ROOT\ files, system \<^verbatim>\options\, and Bib{\TeX} files (\secref{sec:bibtex}). \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{sidekick} \end{center} \caption{The Isabelle NEWS file with SideKick tree view} \label{fig:sidekick} \end{figure} The default SideKick parser for theory files is \<^verbatim>\isabelle\: it provides a tree-view on the formal document structure, with section headings at the top and formal specification elements at the bottom. The alternative parser \<^verbatim>\isabelle-context\ shows nesting of context blocks according to \<^theory_text>\begin \ end\ structure. \<^medskip> Isabelle/ML files are structured according to semi-formal comments that are - explained in @{cite "isabelle-implementation"}. This outline is turned into + explained in \<^cite>\"isabelle-implementation"\. This outline is turned into a tree-view by default, by using the \<^verbatim>\isabelle-ml\ parser. There is also a folding mode of the same name, for hierarchic text folds within ML files. \<^medskip> The special SideKick parser \<^verbatim>\isabelle-markup\ exposes the uninterpreted markup tree of the PIDE document model of the current buffer. This is occasionally useful for informative purposes, but the amount of displayed information might cause problems for large buffers. \ chapter \Prover IDE functionality \label{sec:document-model}\ section \Document model \label{sec:document-model}\ text \ The document model is central to the PIDE architecture: the editor and the prover have a common notion of structured source text with markup, which is produced by formal processing. The editor is responsible for edits of document source, as produced by the user. The prover is responsible for reports of document markup, as produced by its processing in the background. Isabelle/jEdit handles classic editor events of jEdit, in order to connect the physical world of the GUI (with its singleton state) to the mathematical world of multiple document versions (with timeless and stateless updates). \ subsection \Editor buffers and document nodes \label{sec:buffer-node}\ text \ As a regular text editor, jEdit maintains a collection of \<^emph>\buffers\ to store text files; each buffer may be associated with any number of visible \<^emph>\text areas\. Buffers are subject to an \<^emph>\edit mode\ that is determined from the file name extension. The following modes are treated specifically in Isabelle/jEdit: \<^medskip> \begin{tabular}{lll} \<^bold>\mode\ & \<^bold>\file name\ & \<^bold>\content\ \\\hline \<^verbatim>\isabelle\ & \<^verbatim>\*.thy\ & theory source \\ \<^verbatim>\isabelle-ml\ & \<^verbatim>\*.ML\ & Isabelle/ML source \\ \<^verbatim>\sml\ & \<^verbatim>\*.sml\ or \<^verbatim>\*.sig\ & Standard ML source \\ \<^verbatim>\isabelle-root\ & \<^verbatim>\ROOT\ & session root \\ \<^verbatim>\isabelle-options\ & & Isabelle options \\ \<^verbatim>\isabelle-news\ & & Isabelle NEWS \\ \end{tabular} \<^medskip> All jEdit buffers are automatically added to the PIDE document-model as \<^emph>\document nodes\. The overall document structure is defined by the theory nodes in two dimensions: \<^enum> via \<^bold>\theory imports\ that are specified in the \<^emph>\theory header\ using - concrete syntax of the @{command_ref theory} command @{cite - "isabelle-isar-ref"}; + concrete syntax of the @{command_ref theory} command \<^cite>\"isabelle-isar-ref"\; \<^enum> via \<^bold>\auxiliary files\ that are included into a theory by \<^emph>\load commands\, notably @{command_ref ML_file} and @{command_ref SML_file} - @{cite "isabelle-isar-ref"}. + \<^cite>\"isabelle-isar-ref"\. In any case, source files are managed by the PIDE infrastructure: the physical file-system only plays a subordinate role. The relevant version of source text is passed directly from the editor to the prover, using internal communication channels. \ subsection \Theories \label{sec:theories}\ text \ The \<^emph>\Theories\ panel (see also \figref{fig:theories}) provides an overview of the status of continuous checking of theory nodes within the document model. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{theories} \end{center} \caption{Theories panel with an overview of the document-model, and jEdit text areas as editable views on some of the document nodes} \label{fig:theories} \end{figure} Theory imports are resolved automatically by the PIDE document model: all required files are loaded and stored internally, without the need to open corresponding jEdit buffers. Opening or closing editor buffers later on has no direct impact on the formal document content: it only affects visibility. In contrast, auxiliary files (e.g.\ from @{command ML_file} commands) are \<^emph>\not\ resolved within the editor by default, but the prover process takes care of that. This may be changed by enabling the system option @{system_option jedit_auto_resolve}: it ensures that all files are uniformly provided by the editor. \<^medskip> The visible \<^emph>\perspective\ of Isabelle/jEdit is defined by the collective view on theory buffers via open text areas. The perspective is taken as a hint for document processing: the prover ensures that those parts of a theory where the user is looking are checked, while other parts that are presently not required are ignored. The perspective is changed by opening or closing text area windows, or scrolling within a window. The \<^emph>\Theories\ panel provides some further options to influence the process of continuous checking: it may be switched off globally to restrict the prover to superficial processing of command syntax. It is also possible to indicate theory nodes as \<^emph>\required\ for continuous checking: this means such nodes and all their imports are always processed independently of the visibility status (if continuous checking is enabled). Big theory libraries that are marked as required can have significant impact on performance! The \<^emph>\Purge\ button restricts the document model to theories that are required for open editor buffers: inaccessible theories are removed and will be rechecked when opened or imported later. \<^medskip> Formal markup of checked theory content is turned into GUI rendering, based on a standard repertoire known from mainstream IDEs for programming languages: colors, icons, highlighting, squiggly underlines, tooltips, hyperlinks etc. For outer syntax of Isabelle/Isar there is some traditional syntax-highlighting via static keywords and tokenization within the editor; this buffer syntax is determined from theory imports. In contrast, the painting of inner syntax (term language etc.)\ uses semantic information that is reported dynamically from the logical context. Thus the prover can provide additional markup to help the user to understand the meaning of formal text, and to produce more text with some add-on tools (e.g.\ information messages with \<^emph>\sendback\ markup by automated provers or disprovers in the background). \ subsection \Auxiliary files \label{sec:aux-files}\ text \ Special load commands like @{command_ref ML_file} and @{command_ref - SML_file} @{cite "isabelle-isar-ref"} refer to auxiliary files within some + SML_file} \<^cite>\"isabelle-isar-ref"\ refer to auxiliary files within some theory. Conceptually, the file argument of the command extends the theory source by the content of the file, but its editor buffer may be loaded~/ changed~/ saved separately. The PIDE document model propagates changes of auxiliary file content to the corresponding load command in the theory, to update and process it accordingly: changes of auxiliary file content are treated as changes of the corresponding load command. \<^medskip> As a concession to the massive amount of ML files in Isabelle/HOL itself, the content of auxiliary files is only added to the PIDE document-model on demand, the first time when opened explicitly in the editor. There are further tricks to manage markup of ML files, such that Isabelle/HOL may be edited conveniently in the Prover IDE on small machines with only 8\,GB of main memory. Using \<^verbatim>\Pure\ as logic session image, the exploration may start at the top \<^file>\$ISABELLE_HOME/src/HOL/Main.thy\ or the bottom \<^file>\$ISABELLE_HOME/src/HOL/HOL.thy\, for example. It is also possible to explore the Isabelle/Pure bootstrap process (a virtual copy) by opening \<^file>\$ISABELLE_HOME/src/Pure/ROOT.ML\ like a theory in the Prover IDE. Initially, before an auxiliary file is opened in the editor, the prover reads its content from the physical file-system. After the file is opened for the first time in the editor, e.g.\ by following the hyperlink (\secref{sec:tooltips-hyperlinks}) for the argument of its @{command ML_file} command, the content is taken from the jEdit buffer. The change of responsibility from prover to editor counts as an update of the document content, so subsequent theory sources need to be re-checked. When the buffer is closed, the responsibility remains to the editor: the file may be opened again without causing another document update. A file that is opened in the editor, but its theory with the load command is not, is presently inactive in the document model. A file that is loaded via multiple load commands is associated to an arbitrary one: this situation is morally unsupported and might lead to confusion. \<^medskip> Output that refers to an auxiliary file is combined with that of the corresponding load command, and shown whenever the file or the command are active (see also \secref{sec:output}). Warnings, errors, and other useful markup is attached directly to the positions in the auxiliary file buffer, in the manner of standard IDEs. By using the load command @{command SML_file} as explained in \<^file>\$ISABELLE_HOME/src/Tools/SML/Examples.thy\, Isabelle/jEdit may be used as fully-featured IDE for Standard ML, independently of theory or proof development: the required theory merely serves as some kind of project file for a collection of SML source modules. \ section \Output \label{sec:output}\ text \ Prover output consists of \<^emph>\markup\ and \<^emph>\messages\. Both are directly attached to the corresponding positions in the original source text, and visualized in the text area, e.g.\ as text colours for free and bound variables, or as squiggly underlines for warnings, errors etc.\ (see also \figref{fig:output}). In the latter case, the corresponding messages are shown by hovering with the mouse over the highlighted text --- although in many situations the user should already get some clue by looking at the position of the text highlighting, without seeing the message body itself. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{output} \end{center} \caption{Multiple views on prover output: gutter with icon, text area with popup, text overview column, \<^emph>\Theories\ panel, \<^emph>\Output\ panel} \label{fig:output} \end{figure} The ``gutter'' on the left-hand-side of the text area uses icons to provide a summary of the messages within the adjacent text line. Message priorities are used to prefer errors over warnings, warnings over information messages; other output is ignored. The ``text overview column'' on the right-hand-side of the text area uses similar information to paint small rectangles for the overall status of the whole text buffer. The graphics is scaled to fit the logical buffer length into the given window height. Mouse clicks on the overview area move the cursor approximately to the corresponding text line in the buffer. The \<^emph>\Theories\ panel provides another course-grained overview, but without direct correspondence to text positions. The coloured rectangles represent the amount of messages of a certain kind (warnings, errors, etc.) and the execution status of commands. The border of each rectangle indicates the overall status of processing: a thick border means it is \<^emph>\finished\ or \<^emph>\failed\ (with color for errors). A double-click on one of the theory entries with their status overview opens the corresponding text buffer, without moving the cursor to a specific point. \<^medskip> The \<^emph>\Output\ panel displays prover messages that correspond to a given command, within a separate window. The cursor position in the presently active text area determines the prover command whose cumulative message output is appended and shown in that window (in canonical order according to the internal execution of the command). There are also control elements to modify the update policy of the output wrt.\ continued editor movements: \<^emph>\Auto update\ and \<^emph>\Update\. This is particularly useful for multiple instances of the \<^emph>\Output\ panel to look at different situations. Alternatively, the panel can be turned into a passive \<^emph>\Info\ window via the \<^emph>\Detach\ menu item. Proof state is handled separately (\secref{sec:state-output}), but it is also possible to tick the corresponding checkbox to append it to regular output (\figref{fig:output-including-state}). This is a globally persistent option: it affects all open panels and future editor sessions. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{output-including-state} \end{center} \caption{Proof state display within the regular output panel} \label{fig:output-including-state} \end{figure} \<^medskip> Following the IDE principle, regular messages are attached to the original source in the proper place and may be inspected on demand via popups. This excludes messages that are somehow internal to the machinery of proof checking, notably \<^emph>\proof state\ and \<^emph>\tracing\. In any case, the same display technology is used for small popups and big output windows. The formal text contains markup that may be explored recursively via further popups and hyperlinks (see \secref{sec:tooltips-hyperlinks}), or clicked directly to initiate certain actions (see \secref{sec:auto-tools} and \secref{sec:sledgehammer}). \<^medskip> Alternatively, the subsequent actions (with keyboard shortcuts) allow to show tooltip messages or navigate error positions: \<^medskip> \begin{tabular}[t]{l} @{action_ref "isabelle.tooltip"} (\<^verbatim>\CS+b\) \\ @{action_ref "isabelle.message"} (\<^verbatim>\CS+m\) \\ \end{tabular}\quad \begin{tabular}[t]{l} @{action_ref "isabelle.first-error"} (\<^verbatim>\CS+a\) \\ @{action_ref "isabelle.last-error"} (\<^verbatim>\CS+z\) \\ @{action_ref "isabelle.next-error"} (\<^verbatim>\CS+n\) \\ @{action_ref "isabelle.prev-error"} (\<^verbatim>\CS+p\) \\ \end{tabular} \<^medskip> \ section \Proof state \label{sec:state-output}\ text \ The main purpose of the Prover IDE is to help the user editing proof documents, with ongoing formal checking by the prover in the background. This can be done to some extent in the main text area alone, especially for well-structured Isar proofs. Nonetheless, internal proof state needs to be inspected in many situations of exploration and ``debugging''. The \<^emph>\State\ panel shows exclusively such proof state messages without further distraction, while all other messages are displayed in \<^emph>\Output\ (\secref{sec:output}). \Figref{fig:output-and-state} shows a typical GUI layout where both panels are open. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{output-and-state} \end{center} \caption{Separate proof state display (right) and other output (bottom).} \label{fig:output-and-state} \end{figure} Another typical arrangement has more than one \<^emph>\State\ panel open (as floating windows), with \<^emph>\Auto update\ disabled to look at an old situation while the proof text in the vicinity is changed. The \<^emph>\Update\ button triggers an explicit one-shot update; this operation is also available via the action @{action "isabelle.update-state"} (keyboard shortcut \<^verbatim>\S+ENTER\). On small screens, it is occasionally useful to have all messages concatenated in the regular \<^emph>\Output\ panel, e.g.\ see \figref{fig:output-including-state}. \<^medskip> The mechanics of \<^emph>\Output\ versus \<^emph>\State\ are slightly different: \<^item> \<^emph>\Output\ shows information that is continuously produced and already present when the GUI wants to show it. This is implicitly controlled by the visible perspective on the text. \<^item> \<^emph>\State\ initiates a real-time query on demand, with a full round trip including a fresh print operation on the prover side. This is controlled explicitly when the cursor is moved to the next command (\<^emph>\Auto update\) or the \<^emph>\Update\ operation is triggered. This can make a difference in GUI responsibility and resource usage within the prover process. Applications with very big proof states that are only inspected in isolation work better with the \<^emph>\State\ panel. \ section \Query \label{sec:query}\ text \ The \<^emph>\Query\ panel provides various GUI forms to request extra information from the prover, as a replacement of old-style diagnostic commands like @{command find_theorems}. There are input fields and buttons for a particular query command, with output in a dedicated text area. The main query modes are presented as separate tabs: \<^emph>\Find Theorems\, \<^emph>\Find Constants\, \<^emph>\Print Context\, e.g.\ see \figref{fig:query}. As usual in jEdit, multiple \<^emph>\Query\ windows may be active at the same time: any number of floating instances, but at most one docked instance (which is used by default). \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{query} \end{center} \caption{An instance of the Query panel: find theorems} \label{fig:query} \end{figure} \<^medskip> The following GUI elements are common to all query modes: \<^item> The spinning wheel provides feedback about the status of a pending query wrt.\ the evaluation of its context and its own operation. \<^item> The \<^emph>\Apply\ button attaches a fresh query invocation to the current context of the command where the cursor is pointing in the text. \<^item> The \<^emph>\Search\ field allows to highlight query output according to some regular expression, in the notation that is commonly used on the Java platform.\<^footnote>\\<^url>\https://docs.oracle.com/en/java/javase/17/docs/api/java.base/java/util/regex/Pattern.html\\ This may serve as an additional visual filter of the result. \<^item> The \<^emph>\Zoom\ box controls the font size of the output area. All query operations are asynchronous: there is no need to wait for the evaluation of the document for the query context, nor for the query operation itself. Query output may be detached as independent \<^emph>\Info\ window, using a menu operation of the dockable window manager. The printed result usually provides sufficient clues about the original query, with some hyperlink to its context (via markup of its head line). \ subsection \Find theorems\ text \ The \<^emph>\Query\ panel in \<^emph>\Find Theorems\ mode retrieves facts from the theory or proof context matching all of given criteria in the \<^emph>\Find\ text field. A single criterion has the following syntax: \<^rail>\ ('-'?) ('name' ':' @{syntax name} | 'intro' | 'elim' | 'dest' | 'solves' | 'simp' ':' @{syntax term} | @{syntax term}) \ - See also the Isar command @{command_ref find_theorems} in @{cite - "isabelle-isar-ref"}. + See also the Isar command @{command_ref find_theorems} in \<^cite>\"isabelle-isar-ref"\. \ subsection \Find constants\ text \ The \<^emph>\Query\ panel in \<^emph>\Find Constants\ mode prints all constants whose type meets all of the given criteria in the \<^emph>\Find\ text field. A single criterion has the following syntax: \<^rail>\ ('-'?) ('name' ':' @{syntax name} | 'strict' ':' @{syntax type} | @{syntax type}) \ - See also the Isar command @{command_ref find_consts} in @{cite - "isabelle-isar-ref"}. + See also the Isar command @{command_ref find_consts} in \<^cite>\"isabelle-isar-ref"\. \ subsection \Print context\ text \ The \<^emph>\Query\ panel in \<^emph>\Print Context\ mode prints information from the theory or proof context, or proof state. See also the Isar commands @{command_ref print_context}, @{command_ref print_cases}, @{command_ref - print_term_bindings}, @{command_ref print_theorems}, described in @{cite - "isabelle-isar-ref"}. + print_term_bindings}, @{command_ref print_theorems}, described in \<^cite>\"isabelle-isar-ref"\. \ section \Tooltips and hyperlinks \label{sec:tooltips-hyperlinks}\ text \ Formally processed text (prover input or output) contains rich markup that can be explored by using the \<^verbatim>\CONTROL\ modifier key on Linux and Windows, or \<^verbatim>\COMMAND\ on macOS. Hovering with the mouse while the modifier is pressed reveals a \<^emph>\tooltip\ (grey box over the text with a yellow popup) and/or a \<^emph>\hyperlink\ (black rectangle over the text with change of mouse pointer); see also \figref{fig:tooltip}. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{popup1} \end{center} \caption{Tooltip and hyperlink for some formal entity} \label{fig:tooltip} \end{figure} Tooltip popups use the same rendering technology as the main text area, and further tooltips and/or hyperlinks may be exposed recursively by the same mechanism; see \figref{fig:nested-tooltips}. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{popup2} \end{center} \caption{Nested tooltips over formal entities} \label{fig:nested-tooltips} \end{figure} The tooltip popup window provides some controls to \<^emph>\close\ or \<^emph>\detach\ the window, turning it into a separate \<^emph>\Info\ window managed by jEdit. The \<^verbatim>\ESCAPE\ key closes \<^emph>\all\ popups, which is particularly relevant when nested tooltips are stacking up. \<^medskip> A black rectangle in the text indicates a hyperlink that may be followed by a mouse click (while the \<^verbatim>\CONTROL\ or \<^verbatim>\COMMAND\ modifier key is still pressed). Such jumps to other text locations are recorded by the \<^emph>\Navigator\ plugin, which is bundled with Isabelle/jEdit and enabled by default. There are usually navigation arrows in the main jEdit toolbar. Note that the link target may be a file that is itself not subject to formal document processing of the editor session and thus prevents further exploration: the chain of hyperlinks may end in some source file of the underlying logic image, or within the ML bootstrap sources of Isabelle/Pure. \ section \Formal scopes and semantic selection\ text \ Formal entities are semantically annotated in the source text as explained in \secref{sec:tooltips-hyperlinks}. A \<^emph>\formal scope\ consists of the defining position with all its referencing positions. This correspondence is highlighted in the text according to the cursor position, see also \figref{fig:scope1}. Here the referencing positions are rendered with an additional border, in reminiscence to a hyperlink. A mouse click with \<^verbatim>\C\ modifier, or the action @{action_def "isabelle.goto-entity"} (shortcut \<^verbatim>\CS+d\) jumps to the original defining position. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{scope1} \end{center} \caption{Scope of formal entity: defining vs.\ referencing positions} \label{fig:scope1} \end{figure} The action @{action_def "isabelle.select-entity"} (shortcut \<^verbatim>\CS+ENTER\) supports semantic selection of all occurrences of the formal entity at the caret position, with a defining position in the current editor buffer. This facilitates systematic renaming, using regular jEdit editing of a multi-selection, see also \figref{fig:scope2}. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{scope2} \end{center} \caption{The result of semantic selection and systematic renaming} \label{fig:scope2} \end{figure} By default, the visual feedback on scopes is restricted to definitions within the visible text area. The keyboard modifier \<^verbatim>\CS\ overrides this: then all defining and referencing positions are shown. This modifier may be configured via option @{system_option jedit_focus_modifier}; the default coincides with the modifier for the above keyboard actions. The empty string means to disable this additional visual feedback. \ section \Completion \label{sec:completion}\ text \ Smart completion of partial input is the IDE functionality \<^emph>\par excellance\. Isabelle/jEdit combines several sources of information to achieve that. Despite its complexity, it should be possible to get some idea how completion works by experimentation, based on the overview of completion varieties in \secref{sec:completion-varieties}. The remaining subsections explain concepts around completion more systematically. \<^medskip> \<^emph>\Explicit completion\ is triggered by the action @{action_ref "isabelle.complete"}, which is bound to the keyboard shortcut \<^verbatim>\C+b\, and thus overrides the jEdit default for @{action_ref "complete-word"}. \<^emph>\Implicit completion\ hooks into the regular keyboard input stream of the editor, with some event filtering and optional delays. \<^medskip> Completion options may be configured in \<^emph>\Plugin Options~/ Isabelle~/ General~/ Completion\. These are explained in further detail below, whenever relevant. There is also a summary of options in \secref{sec:completion-options}. The asynchronous nature of PIDE interaction means that information from the prover is delayed --- at least by a full round-trip of the document update protocol. The default options already take this into account, with a sufficiently long completion delay to speculate on the availability of all relevant information from the editor and the prover, before completing text immediately or producing a popup. Although there is an inherent danger of non-deterministic behaviour due to such real-time parameters, the general completion policy aims at determined results as far as possible. \ subsection \Varieties of completion \label{sec:completion-varieties}\ subsubsection \Built-in templates\ text \ Isabelle is ultimately a framework of nested sub-languages of different kinds and purposes. The completion mechanism supports this by the following built-in templates: \<^descr> \<^verbatim>\`\ (single ASCII back-quote) or \<^verbatim>\"\ (double ASCII quote) support \<^emph>\quotations\ via text cartouches. There are three selections, which are always presented in the same order and do not depend on any context information. The default choice produces a template ``\\\\\'', where the box indicates the cursor position after insertion; the other choices help to repair the block structure of unbalanced text cartouches. \<^descr> \<^verbatim>\@{\ is completed to the template ``\@{\}\'', where the box indicates the cursor position after insertion. Here it is convenient to use the wildcard ``\<^verbatim>\__\'' or a more specific name prefix to let semantic completion of name-space entries propose antiquotation names. With some practice, input of quoted sub-languages and antiquotations of embedded languages should work smoothly. Note that national keyboard layouts might cause problems with back-quote as dead key, but double quote can be used instead. \ subsubsection \Syntax keywords\ text \ Syntax completion tables are determined statically from the keywords of the ``outer syntax'' of the underlying edit mode: for theory files this is the syntax of Isar commands according to the cumulative theory imports. Keywords are usually plain words, which means the completion mechanism only inserts them directly into the text for explicit completion (\secref{sec:completion-input}), but produces a popup (\secref{sec:completion-popup}) otherwise. At the point where outer syntax keywords are defined, it is possible to specify an alternative replacement string to be inserted instead of the keyword itself. An empty string means to suppress the keyword altogether, which is occasionally useful to avoid confusion, e.g.\ the rare keyword @{command simproc_setup} vs.\ the frequent name-space entry \simp\. \ subsubsection \Isabelle symbols\ text \ The completion tables for Isabelle symbols (\secref{sec:symbols}) are determined statically from \<^file>\$ISABELLE_HOME/etc/symbols\ and \<^path>\$ISABELLE_HOME_USER/etc/symbols\ for each symbol specification as follows: \<^medskip> \begin{tabular}{ll} \<^bold>\completion entry\ & \<^bold>\example\ \\\hline literal symbol & \<^verbatim>\\\ \\ symbol name with backslash & \<^verbatim>\\\\<^verbatim>\forall\ \\ symbol abbreviation & \<^verbatim>\ALL\ or \<^verbatim>\!\ \\ \end{tabular} \<^medskip> When inserted into the text, the above examples all produce the same Unicode rendering \\\ of the underlying symbol \<^verbatim>\\\. A symbol abbreviation that is a plain word, like \<^verbatim>\ALL\, is treated like a syntax keyword. Non-word abbreviations like \<^verbatim>\-->\ are inserted more aggressively, except for single-character abbreviations like \<^verbatim>\!\ above. Completion via abbreviations like \<^verbatim>\ALL\ or \<^verbatim>\-->\ depends on the semantic language context (\secref{sec:completion-context}). In contrast, backslash sequences like \<^verbatim>\\forall\ \<^verbatim>\\\ are always possible, but require additional interaction to confirm (via popup). This is important in ambiguous situations, e.g.\ for Isabelle document source, which may contain formal symbols or informal {\LaTeX} macros. Backslash sequences also help when input is broken, and thus escapes its normal semantic context: e.g.\ antiquotations or string literals in ML, which do not allow arbitrary backslash sequences. Special symbols like \<^verbatim>\\\ or control symbols like \<^verbatim>\\<^cancel>\, \<^verbatim>\\<^latex>\, \<^verbatim>\\<^binding>\ can have an argument: completing on a name prefix offers a template with an empty cartouche. Thus completion of \<^verbatim>\\co\ or \<^verbatim>\\ca\ allows to compose formal document comments quickly.\<^footnote>\It is customary to put a space between \<^verbatim>\\\ and its argument, while control symbols do \<^emph>\not\ allow extra space here.\ \ subsubsection \User-defined abbreviations\ text \ The theory header syntax supports abbreviations via the \<^theory_text>\abbrevs\ keyword - @{cite "isabelle-isar-ref"}. This is a slight generalization of built-in + \<^cite>\"isabelle-isar-ref"\. This is a slight generalization of built-in templates and abbreviations for Isabelle symbols, as explained above. Examples may be found in the Isabelle sources, by searching for ``\<^verbatim>\abbrevs\'' in \<^verbatim>\*.thy\ files. The \<^emph>\Symbols\ panel shows the abbreviations that are available in the current theory buffer (according to its \<^theory_text>\imports\) in the \<^verbatim>\Abbrevs\ tab. \ subsubsection \Name-space entries\ text \ This is genuine semantic completion, using information from the prover, so it requires some delay. A \<^emph>\failed name-space lookup\ produces an error message that is annotated with a list of alternative names that are legal. The list of results is truncated according to the system option @{system_option_ref completion_limit}. The completion mechanism takes this into account when collecting information on the prover side. Already recognized names are \<^emph>\not\ completed further, but completion may be extended by appending a suffix of underscores. This provokes a failed lookup, and another completion attempt (ignoring the underscores). For example, in a name space where \<^verbatim>\foo\ and \<^verbatim>\foobar\ are known, the input \<^verbatim>\foo\ remains unchanged, but \<^verbatim>\foo_\ may be completed to \<^verbatim>\foo\ or \<^verbatim>\foobar\. The special identifier ``\<^verbatim>\__\'' serves as a wild-card for arbitrary completion: it exposes the name-space content to the completion mechanism (truncated according to @{system_option completion_limit}). This is occasionally useful to explore an unknown name-space, e.g.\ in some template. \ subsubsection \File-system paths\ text \ Depending on prover markup about file-system paths in the source text, e.g.\ for the argument of a load command (\secref{sec:aux-files}), the completion mechanism explores the directory content and offers the result as completion popup. Relative path specifications are understood wrt.\ the \<^emph>\master directory\ of the document node (\secref{sec:buffer-node}) of the enclosing editor buffer; this requires a proper theory, not an auxiliary file. A suffix of slashes may be used to continue the exploration of an already recognized directory name. \ subsubsection \Spell-checking\ text \ The spell-checker combines semantic markup from the prover (regions of plain words) with static dictionaries (word lists) that are known to the editor. Unknown words are underlined in the text, using @{system_option_ref spell_checker_color} (blue by default). This is not an error, but a hint to the user that some action may be taken. The jEdit context menu provides various actions, as far as applicable: \<^medskip> \begin{tabular}{l} @{action_ref "isabelle.complete-word"} \\ @{action_ref "isabelle.exclude-word"} \\ @{action_ref "isabelle.exclude-word-permanently"} \\ @{action_ref "isabelle.include-word"} \\ @{action_ref "isabelle.include-word-permanently"} \\ \end{tabular} \<^medskip> Instead of the specific @{action_ref "isabelle.complete-word"}, it is also possible to use the generic @{action_ref "isabelle.complete"} with its default keyboard shortcut \<^verbatim>\C+b\. \<^medskip> Dictionary lookup uses some educated guesses about lower-case, upper-case, and capitalized words. This is oriented on common use in English, where this aspect is not decisive for proper spelling (in contrast to German, for example). \ subsection \Semantic completion context \label{sec:completion-context}\ text \ Completion depends on a semantic context that is provided by the prover, although with some delay, because at least a full PIDE protocol round-trip is required. Until that information becomes available in the PIDE document-model, the default context is given by the outer syntax of the editor mode (see also \secref{sec:buffer-node}). The semantic \<^emph>\language context\ provides information about nested sub-languages of Isabelle: keywords are only completed for outer syntax, and antiquotations for languages that support them. Symbol abbreviations only work for specific sub-languages: e.g.\ ``\<^verbatim>\=>\'' is \<^emph>\not\ completed in regular ML source, but is completed within ML strings, comments, antiquotations. Backslash representations of symbols like ``\<^verbatim>\\foobar\'' or ``\<^verbatim>\\\'' work in any context --- after additional confirmation. The prover may produce \<^emph>\no completion\ markup in exceptional situations, to tell that some language keywords should be excluded from further completion attempts. For example, ``\<^verbatim>\:\'' within accepted Isar syntax looses its meaning as abbreviation for symbol ``\\\''. \ subsection \Input events \label{sec:completion-input}\ text \ Completion is triggered by certain events produced by the user, with optional delay after keyboard input according to @{system_option jedit_completion_delay}. \<^descr>[Explicit completion] works via action @{action_ref "isabelle.complete"} with keyboard shortcut \<^verbatim>\C+b\. This overrides the shortcut for @{action_ref "complete-word"} in jEdit, but it is possible to restore the original jEdit keyboard mapping of @{action "complete-word"} via \<^emph>\Global Options~/ Shortcuts\ and invent a different one for @{action "isabelle.complete"}. \<^descr>[Explicit spell-checker completion] works via @{action_ref "isabelle.complete-word"}, which is exposed in the jEdit context menu, if the mouse points to a word that the spell-checker can complete. \<^descr>[Implicit completion] works via regular keyboard input of the editor. It depends on further side-conditions: \<^enum> The system option @{system_option_ref jedit_completion} needs to be enabled (default). \<^enum> Completion of syntax keywords requires at least 3 relevant characters in the text. \<^enum> The system option @{system_option_ref jedit_completion_delay} determines an additional delay (0.5 by default), before opening a completion popup. The delay gives the prover a chance to provide semantic completion information, notably the context (\secref{sec:completion-context}). \<^enum> The system option @{system_option_ref jedit_completion_immediate} (enabled by default) controls whether replacement text should be inserted immediately without popup, regardless of @{system_option jedit_completion_delay}. This aggressive mode of completion is restricted to symbol abbreviations that are not plain words (\secref{sec:symbols}). \<^enum> Completion of symbol abbreviations with only one relevant character in the text always enforces an explicit popup, regardless of @{system_option_ref jedit_completion_immediate}. \ subsection \Completion popup \label{sec:completion-popup}\ text \ A \<^emph>\completion popup\ is a minimally invasive GUI component over the text area that offers a selection of completion items to be inserted into the text, e.g.\ by mouse clicks. Items are sorted dynamically, according to the frequency of selection, with persistent history. The popup may interpret special keys \<^verbatim>\ENTER\, \<^verbatim>\TAB\, \<^verbatim>\ESCAPE\, \<^verbatim>\UP\, \<^verbatim>\DOWN\, \<^verbatim>\PAGE_UP\, \<^verbatim>\PAGE_DOWN\, but all other key events are passed to the underlying text area. This allows to ignore unwanted completions most of the time and continue typing quickly. Thus the popup serves as a mechanism of confirmation of proposed items, while the default is to continue without completion. The meaning of special keys is as follows: \<^medskip> \begin{tabular}{ll} \<^bold>\key\ & \<^bold>\action\ \\\hline \<^verbatim>\ENTER\ & select completion (if @{system_option jedit_completion_select_enter}) \\ \<^verbatim>\TAB\ & select completion (if @{system_option jedit_completion_select_tab}) \\ \<^verbatim>\ESCAPE\ & dismiss popup \\ \<^verbatim>\UP\ & move up one item \\ \<^verbatim>\DOWN\ & move down one item \\ \<^verbatim>\PAGE_UP\ & move up one page of items \\ \<^verbatim>\PAGE_DOWN\ & move down one page of items \\ \end{tabular} \<^medskip> Movement within the popup is only active for multiple items. Otherwise the corresponding key event retains its standard meaning within the underlying text area. \ subsection \Insertion \label{sec:completion-insert}\ text \ Completion may first propose replacements to be selected (via a popup), or replace text immediately in certain situations and depending on certain options like @{system_option jedit_completion_immediate}. In any case, insertion works uniformly, by imitating normal jEdit text insertion, depending on the state of the \<^emph>\text selection\. Isabelle/jEdit tries to accommodate the most common forms of advanced selections in jEdit, but not all combinations make sense. At least the following important cases are well-defined: \<^descr>[No selection.] The original is removed and the replacement inserted, depending on the caret position. \<^descr>[Rectangular selection of zero width.] This special case is treated by jEdit as ``tall caret'' and insertion of completion imitates its normal behaviour: separate copies of the replacement are inserted for each line of the selection. \<^descr>[Other rectangular selection or multiple selections.] Here the original is removed and the replacement is inserted for each line (or segment) of the selection. Support for multiple selections is particularly useful for \<^emph>\HyperSearch\: clicking on one of the items in the \<^emph>\HyperSearch Results\ window makes jEdit select all its occurrences in the corresponding line of text. Then explicit completion can be invoked via \<^verbatim>\C+b\, e.g.\ to replace occurrences of \<^verbatim>\-->\ by \\\. \<^medskip> Insertion works by removing and inserting pieces of text from the buffer. This counts as one atomic operation on the jEdit history. Thus unintended completions may be reverted by the regular @{action undo} action of jEdit. According to normal jEdit policies, the recovered text after @{action undo} is selected: \<^verbatim>\ESCAPE\ is required to reset the selection and to continue typing more text. \ subsection \Options \label{sec:completion-options}\ text \ This is a summary of Isabelle/Scala system options that are relevant for completion. They may be configured in \<^emph>\Plugin Options~/ Isabelle~/ General\ as usual. \<^item> @{system_option_def completion_limit} specifies the maximum number of items for various semantic completion operations (name-space entries etc.) \<^item> @{system_option_def jedit_completion} guards implicit completion via regular jEdit key events (\secref{sec:completion-input}): it allows to disable implicit completion altogether. \<^item> @{system_option_def jedit_completion_select_enter} and @{system_option_def jedit_completion_select_tab} enable keys to select a completion item from the popup (\secref{sec:completion-popup}). Note that a regular mouse click on the list of items is always possible. \<^item> @{system_option_def jedit_completion_context} specifies whether the language context provided by the prover should be used at all. Disabling that option makes completion less ``semantic''. Note that incomplete or severely broken input may cause some disagreement of the prover and the user about the intended language context. \<^item> @{system_option_def jedit_completion_delay} and @{system_option_def jedit_completion_immediate} determine the handling of keyboard events for implicit completion (\secref{sec:completion-input}). A @{system_option jedit_completion_delay}~\<^verbatim>\> 0\ postpones the processing of key events, until after the user has stopped typing for the given time span, but @{system_option jedit_completion_immediate}~\<^verbatim>\= true\ means that abbreviations of Isabelle symbols are handled nonetheless. \<^item> @{system_option_def completion_path_ignore} specifies ``glob'' patterns to ignore in file-system path completion (separated by colons), e.g.\ backup files ending with tilde. \<^item> @{system_option_def spell_checker} is a global guard for all spell-checker operations: it allows to disable that mechanism altogether. \<^item> @{system_option_def spell_checker_dictionary} determines the current dictionary, taken from the colon-separated list in the settings variable @{setting_def JORTHO_DICTIONARIES}. There are jEdit actions to specify local updates to a dictionary, by including or excluding words. The result of permanent dictionary updates is stored in the directory \<^path>\$ISABELLE_HOME_USER/dictionaries\, in a separate file for each dictionary. \<^item> @{system_option_def spell_checker_include} specifies a comma-separated list of markup elements that delimit words in the source that is subject to spell-checking, including various forms of comments. \<^item> @{system_option_def spell_checker_exclude} specifies a comma-separated list of markup elements that disable spell-checking (e.g.\ in nested antiquotations). \ section \Automatically tried tools \label{sec:auto-tools}\ text \ Continuous document processing works asynchronously in the background. Visible document source that has been evaluated may get augmented by additional results of \<^emph>\asynchronous print functions\. An example for that is proof state output, if that is enabled in the Output panel (\secref{sec:output}). More heavy-weight print functions may be applied as well, e.g.\ to prove or disprove parts of the formal text by other means. Isabelle/HOL provides various automatically tried tools that operate on outermost goal statements (e.g.\ @{command lemma}, @{command theorem}), independently of the state of the current proof attempt. They work implicitly without any arguments. Results are output as \<^emph>\information messages\, which are indicated in the text area by blue squiggles and a blue information sign in the gutter (see \figref{fig:auto-tools}). The message content may be shown as for other output (see also \secref{sec:output}). Some tools produce output with \<^emph>\sendback\ markup, which means that clicking on certain parts of the text inserts that into the source in the proper place. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{auto-tools} \end{center} \caption{Result of automatically tried tools} \label{fig:auto-tools} \end{figure} \<^medskip> The following Isabelle system options control the behavior of automatically tried tools (see also the jEdit dialog window \<^emph>\Plugin Options~/ Isabelle~/ General~/ Automatically tried tools\): \<^item> @{system_option_ref auto_methods} controls automatic use of a combination of standard proof methods (@{method auto}, @{method simp}, @{method blast}, - etc.). This corresponds to the Isar command @{command_ref "try0"} @{cite - "isabelle-isar-ref"}. + etc.). This corresponds to the Isar command @{command_ref "try0"} \<^cite>\"isabelle-isar-ref"\. The tool is disabled by default, since unparameterized invocation of standard proof methods often consumes substantial CPU resources without leading to success. \<^item> @{system_option_ref auto_nitpick} controls a slightly reduced version of @{command_ref nitpick}, which tests for counterexamples using first-order - relational logic. See also the Nitpick manual @{cite "isabelle-nitpick"}. + relational logic. See also the Nitpick manual \<^cite>\"isabelle-nitpick"\. This tool is disabled by default, due to the extra overhead of invoking an external Java process for each attempt to disprove a subgoal. \<^item> @{system_option_ref auto_quickcheck} controls automatic use of @{command_ref quickcheck}, which tests for counterexamples using a series of assignments for free variables of a subgoal. This tool is \<^emph>\enabled\ by default. It requires little overhead, but is a bit weaker than @{command nitpick}. \<^item> @{system_option_ref auto_sledgehammer} controls a significantly reduced version of @{command_ref sledgehammer}, which attempts to prove a subgoal - using external automatic provers. See also the Sledgehammer manual @{cite - "isabelle-sledgehammer"}. + using external automatic provers. See also the Sledgehammer manual \<^cite>\"isabelle-sledgehammer"\. This tool is disabled by default, due to the relatively heavy nature of Sledgehammer. \<^item> @{system_option_ref auto_solve_direct} controls automatic use of @{command_ref solve_direct}, which checks whether the current subgoals can be solved directly by an existing theorem. This also helps to detect duplicate lemmas. This tool is \<^emph>\enabled\ by default. Invocation of automatically tried tools is subject to some global policies of parallel execution, which may be configured as follows: \<^item> @{system_option_ref auto_time_limit} (default 2.0) determines the timeout (in seconds) for each tool execution. \<^item> @{system_option_ref auto_time_start} (default 1.0) determines the start delay (in seconds) for automatically tried tools, after the main command evaluation is finished. Each tool is submitted independently to the pool of parallel execution tasks in Isabelle/ML, using hardwired priorities according to its relative ``heaviness''. The main stages of evaluation and printing of proof states take precedence, but an already running tool is not canceled and may thus reduce reactivity of proof document processing. Users should experiment how the available CPU resources (number of cores) are best invested to get additional feedback from prover in the background, by using a selection of weaker or stronger tools. \ section \Sledgehammer \label{sec:sledgehammer}\ text \ The \<^emph>\Sledgehammer\ panel (\figref{fig:sledgehammer}) provides a view on some independent execution of the Isar command @{command_ref sledgehammer}, with process indicator (spinning wheel) and GUI elements for important Sledgehammer arguments and options. Any number of Sledgehammer panels may be active, according to the standard policies of Dockable Window Management in jEdit. Closing such windows also cancels the corresponding prover tasks. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{sledgehammer} \end{center} \caption{An instance of the Sledgehammer panel} \label{fig:sledgehammer} \end{figure} The \<^emph>\Apply\ button attaches a fresh invocation of @{command sledgehammer} to the command where the cursor is pointing in the text --- this should be some pending proof problem. Further buttons like \<^emph>\Cancel\ and \<^emph>\Locate\ help to manage the running process. Results appear incrementally in the output window of the panel. Proposed proof snippets are marked-up as \<^emph>\sendback\, which means a single mouse click inserts the text into a suitable place of the original source. Some manual editing may be required nonetheless, say to remove earlier proof attempts. \ chapter \Isabelle document preparation\ text \ The ultimate purpose of Isabelle is to produce nicely rendered documents with the Isabelle document preparation system, which is based on {\LaTeX}; - see also @{cite "isabelle-system" and "isabelle-isar-ref"}. Isabelle/jEdit + see also \<^cite>\"isabelle-system" and "isabelle-isar-ref"\. Isabelle/jEdit provides some additional support for document editing. \ section \Document outline\ text \ Theory sources may contain document markup commands, such as @{command_ref chapter}, @{command_ref section}, @{command subsection}. The Isabelle SideKick parser (\secref{sec:sidekick}) represents this document outline as structured tree view, with formal statements and proofs nested inside; see \figref{fig:sidekick-document}. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{sidekick-document} \end{center} \caption{Isabelle document outline via SideKick tree view} \label{fig:sidekick-document} \end{figure} It is also possible to use text folding according to this structure, by adjusting \<^emph>\Utilities / Buffer Options / Folding mode\ of jEdit. The default mode \<^verbatim>\isabelle\ uses the structure of formal definitions, statements, and proofs. The alternative mode \<^verbatim>\sidekick\ uses the document structure of the SideKick parser, as explained above. \ section \Markdown structure\ text \ Document text is internally structured in paragraphs and nested lists, using notation that is similar to Markdown\<^footnote>\\<^url>\https://commonmark.org\\. There are special control symbols for items of different kinds of lists, corresponding to \<^verbatim>\itemize\, \<^verbatim>\enumerate\, \<^verbatim>\description\ in {\LaTeX}. This is illustrated in for \<^verbatim>\itemize\ in \figref{fig:markdown-document}. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{markdown-document} \end{center} \caption{Markdown structure within document text} \label{fig:markdown-document} \end{figure} Items take colour according to the depth of nested lists. This helps to explore the implicit rules for list structure interactively. There is also markup for individual items and paragraphs in the text: it may be explored via mouse hovering with \<^verbatim>\CONTROL\ / \<^verbatim>\COMMAND\ as usual (\secref{sec:tooltips-hyperlinks}). \ section \Citations and Bib{\TeX} entries \label{sec:bibtex}\ text \ Citations are managed by {\LaTeX} and Bib{\TeX} in \<^verbatim>\.bib\ files. The - Isabelle session build process and the @{tool document} tool @{cite - "isabelle-system"} are smart enough to assemble the result, based on the + Isabelle session build process and the @{tool document} tool \<^cite>\"isabelle-system"\ are smart enough to assemble the result, based on the session directory layout. - The document antiquotation \@{cite}\ is described in @{cite - "isabelle-isar-ref"}. Within the Prover IDE it provides semantic markup for + The document antiquotation \@{cite}\ is described in \<^cite>\"isabelle-isar-ref"\. Within the Prover IDE it provides semantic markup for tooltips, hyperlinks, and completion for Bib{\TeX} database entries. Isabelle/jEdit does \<^emph>\not\ know about the actual Bib{\TeX} environment used in {\LaTeX} batch-mode, but it can take citations from those \<^verbatim>\.bib\ files that happen to be open in the editor; see \figref{fig:cite-completion}. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{cite-completion} \end{center} \caption{Semantic completion of citations from open Bib{\TeX} files} \label{fig:cite-completion} \end{figure} Isabelle/jEdit also provides IDE support for editing \<^verbatim>\.bib\ files themselves. There is syntax highlighting based on entry types (according to standard Bib{\TeX} styles), a context-menu to compose entries systematically, and a SideKick tree view of the overall content; see \figref{fig:bibtex-mode}. Semantic checking with errors and warnings is performed by the original \<^verbatim>\bibtex\ tool using style \<^verbatim>\plain\: different Bib{\TeX} styles may produce slightly different results. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{bibtex-mode} \end{center} \caption{Bib{\TeX} mode with context menu, SideKick tree view, and semantic output from the \<^verbatim>\bibtex\ tool} \label{fig:bibtex-mode} \end{figure} Regular document preview (\secref{sec:document-preview}) of \<^verbatim>\.bib\ files approximates the usual {\LaTeX} bibliography output in HTML (using style \<^verbatim>\unsort\). \ section \Document preview and printing \label{sec:document-preview}\ text \ The action @{action_def isabelle.preview} opens an HTML preview of the current document node in the default web browser. The content is derived from the semantic markup produced by the prover, and thus depends on the status of formal processing. Action @{action_def isabelle.draft} is similar to @{action isabelle.preview}, but shows a plain-text document draft. Both actions show document sources in a regular Web browser, which may be also used to print the result in a more portable manner than the Java printer dialog of the jEdit @{action_ref print} action. \ chapter \ML debugging within the Prover IDE\ text \ Isabelle/ML is based on Poly/ML\<^footnote>\\<^url>\https://www.polyml.org\\ and thus benefits from the source-level debugger of that implementation of Standard ML. The Prover IDE provides the \<^emph>\Debugger\ dockable to connect to running ML threads, inspect the stack frame with local ML bindings, and evaluate ML expressions in a particular run-time context. A typical debugger session is shown in \figref{fig:ml-debugger}. ML debugging depends on the following pre-requisites. \<^enum> ML source needs to be compiled with debugging enabled. This may be controlled for particular chunks of ML sources using any of the subsequent facilities. \<^enum> The system option @{system_option_ref ML_debugger} as implicit state of the Isabelle process. It may be changed in the menu \<^emph>\Plugins / Plugin Options / Isabelle / General\. ML modules need to be reloaded and recompiled to pick up that option as intended. \<^enum> The configuration option @{attribute_ref ML_debugger}, with an attribute of the same name, to update a global or local context (e.g.\ with the @{command declare} command). \<^enum> Commands that modify @{attribute ML_debugger} state for individual files: @{command_ref ML_file_debug}, @{command_ref ML_file_no_debug}, @{command_ref SML_file_debug}, @{command_ref SML_file_no_debug}. The instrumentation of ML code for debugging causes minor run-time overhead. ML modules that implement critical system infrastructure may lead to deadlocks or other undefined behaviour, when put under debugger control! \<^enum> The \<^emph>\Debugger\ panel needs to be active, otherwise the program ignores debugger instrumentation of the compiler and runs unmanaged. It is also possible to start debugging with the panel open, and later undock it, to let the program continue unhindered. \<^enum> The ML program needs to be stopped at a suitable breakpoint, which may be activated individually or globally as follows. For ML sources that have been compiled with debugger support, the IDE visualizes possible breakpoints in the text. A breakpoint may be toggled by pointing accurately with the mouse, with a right-click to activate jEdit's context menu and its \<^emph>\Toggle Breakpoint\ item. Alternatively, the \<^emph>\Break\ checkbox in the \<^emph>\Debugger\ panel may be enabled to stop ML threads always at the next possible breakpoint. Note that the state of individual breakpoints \<^emph>\gets lost\ when the coresponding ML source is re-compiled! This may happen unintentionally, e.g.\ when following hyperlinks into ML modules that have not been loaded into the IDE before. \begin{figure}[!htb] \begin{center} \includegraphics[scale=0.333]{ml-debugger} \end{center} \caption{ML debugger session} \label{fig:ml-debugger} \end{figure} The debugger panel (\figref{fig:ml-debugger}) shows a list of all threads that are presently stopped. Each thread shows a stack of all function invocations that lead to the current breakpoint at the top. It is possible to jump between stack positions freely, by clicking on this list. The current situation is displayed in the big output window, as a local ML environment with names and printed values. ML expressions may be evaluated in the current context by entering snippets of source into the text fields labeled \Context\ and \ML\, and pushing the \Eval\ button. By default, the source is interpreted as Isabelle/ML with the usual support for antiquotations (like @{command ML}, @{command ML_file}). Alternatively, strict Standard ML may be enforced via the \<^emph>\SML\ checkbox (like @{command SML_file}). The context for Isabelle/ML is optional, it may evaluate to a value of type \<^ML_type>\theory\, \<^ML_type>\Proof.context\, or \<^ML_type>\Context.generic\. Thus the given ML expression (with its antiquotations) may be subject to the intended dynamic run-time context, instead of the static compile-time context. \<^medskip> The buttons labeled \<^emph>\Continue\, \<^emph>\Step\, \<^emph>\Step over\, \<^emph>\Step out\ recommence execution of the program, with different policies concerning nested function invocations. The debugger always moves the cursor within the ML source to the next breakpoint position, and offers new stack frames as before. \ chapter \Miscellaneous tools\ section \Timing and monitoring\ text \ Managed evaluation of commands within PIDE documents includes timing information, which consists of elapsed (wall-clock) time, CPU time, and GC (garbage collection) time. Note that in a multithreaded system it is difficult to measure execution time precisely: elapsed time is closer to the real requirements of runtime resources than CPU or GC time, which are both subject to influences from the parallel environment that are outside the scope of the current command transaction. The \<^emph>\Timing\ panel provides an overview of cumulative command timings for each document node. Commands with elapsed time below the given threshold are ignored in the grand total. Nodes are sorted according to their overall timing. For the document node that corresponds to the current buffer, individual command timings are shown as well. A double-click on a theory node or command moves the editor focus to that particular source position. It is also possible to reveal individual timing information via some tooltip for the corresponding command keyword, using the technique of mouse hovering with \<^verbatim>\CONTROL\~/ \<^verbatim>\COMMAND\ modifier (\secref{sec:tooltips-hyperlinks}). Actual display of timing depends on the global option @{system_option_ref jedit_timing_threshold}, which can be configured in \<^emph>\Plugin Options~/ Isabelle~/ General\. \<^medskip> The jEdit status line includes a monitor widget for the current heap usage of the Isabelle/ML process; this includes information about ongoing garbage collection (shown as ``ML cleanup''). A double-click opens a new instance of the \<^emph>\Monitor\ panel, as explained below. There is a similar widget for the JVM: a double-click opens an external Java monitor process with detailed information and controls for the Java process underlying Isabelle/Scala/jEdit (this is based on \<^verbatim>\jconsole\). \<^medskip> The \<^emph>\Monitor\ panel visualizes various data collections about recent activity of the runtime system of Isabelle/ML and Java. There are buttons to request a full garbage collection and sharing of live data on the ML heap. The display is continuously updated according to @{system_option_ref editor_chart_delay}. Note that the painting of the chart takes considerable runtime itself --- on the Java Virtual Machine that runs Isabelle/Scala, not Isabelle/ML. \ section \Low-level output\ text \ Prover output is normally shown directly in the main text area or specific panels like \<^emph>\Output\ (\secref{sec:output}) or \<^emph>\State\ (\secref{sec:state-output}). Beyond this, it is occasionally useful to inspect low-level output channels via some of the following additional panels: \<^item> \<^emph>\Protocol\ shows internal messages between the Isabelle/Scala and Isabelle/ML side of the PIDE document editing protocol. Recording of messages starts with the first activation of the corresponding dockable window; earlier messages are lost. Display of protocol messages causes considerable slowdown, so it is important to undock all \<^emph>\Protocol\ panels for production work. \<^item> \<^emph>\Raw Output\ shows chunks of text from the \<^verbatim>\stdout\ and \<^verbatim>\stderr\ channels of the prover process. Recording of output starts with the first activation of the corresponding dockable window; earlier output is lost. The implicit stateful nature of physical I/O channels makes it difficult to relate raw output to the actual command from where it was originating. Parallel execution may add to the confusion. Peeking at physical process I/O is only the last resort to diagnose problems with tools that are not PIDE compliant. Under normal circumstances, prover output always works via managed message channels (corresponding to \<^ML>\writeln\, \<^ML>\warning\, \<^ML>\Output.error_message\ in Isabelle/ML), which are displayed by regular means within the document model (\secref{sec:output}). Unhandled Isabelle/ML exceptions are printed by the system via \<^ML>\Output.error_message\. \<^item> \<^emph>\Syslog\ shows system messages that might be relevant to diagnose problems with the startup or shutdown phase of the prover process; this also includes raw output on \<^verbatim>\stderr\. Isabelle/ML also provides an explicit \<^ML>\Output.system_message\ operation, which is occasionally useful for diagnostic purposes within the system infrastructure itself. A limited amount of syslog messages are buffered, independently of the docking state of the \<^emph>\Syslog\ panel. This allows to diagnose serious problems with Isabelle/PIDE process management, outside of the actual protocol layer. Under normal situations, such low-level system output can be ignored. \ chapter \Known problems and workarounds \label{sec:problems}\ text \ \<^item> \<^bold>\Problem:\ Keyboard shortcuts \<^verbatim>\C+PLUS\ and \<^verbatim>\C+MINUS\ for adjusting the editor font size depend on platform details and national keyboards. \<^bold>\Workaround:\ Rebind keys via \<^emph>\Global Options~/ Shortcuts\. \<^item> \<^bold>\Problem:\ The macOS key sequence \<^verbatim>\COMMAND+COMMA\ for application \<^emph>\Preferences\ is in conflict with the jEdit default keyboard shortcut for \<^emph>\Incremental Search Bar\ (action @{action_ref "quick-search"}). \<^bold>\Workaround:\ Rebind key via \<^emph>\Global Options~/ Shortcuts\ according to the national keyboard layout, e.g.\ \<^verbatim>\COMMAND+SLASH\ on English ones. \<^item> \<^bold>\Problem:\ On macOS with native Apple look-and-feel, some exotic national keyboards may cause a conflict of menu accelerator keys with regular jEdit key bindings. This leads to duplicate execution of the corresponding jEdit action. \<^bold>\Workaround:\ Disable the native Apple menu bar via Java runtime option \<^verbatim>\-Dapple.laf.useScreenMenuBar=false\. \<^item> \<^bold>\Problem:\ macOS system fonts sometimes lead to character drop-outs in the main text area. \<^bold>\Workaround:\ Use the default \<^verbatim>\Isabelle DejaVu\ fonts. \<^item> \<^bold>\Problem:\ On macOS the Java printer dialog sometimes does not work. \<^bold>\Workaround:\ Use action @{action isabelle.draft} and print via the Web browser. \<^item> \<^bold>\Problem:\ Antialiased text rendering may show bad performance or bad visual quality, notably on Linux/X11. \<^bold>\Workaround:\ The property \<^verbatim>\view.antiAlias\ (via menu item Utilities / Global Options / Text Area / Anti Aliased smooth text) has the main impact on text rendering, but some related properties may also change the behaviour. The default is \<^verbatim>\view.antiAlias=subpixel HRGB\: it can be much faster than \<^verbatim>\standard\, but occasionally causes problems with odd color shades. An alternative is to have \<^verbatim>\view.antiAlias=standard\ and set a Java system property like this:\<^footnote>\See also \<^url>\https://docs.oracle.com/javase/10/troubleshoot/java-2d-pipeline-rendering-and-properties.htm\.\ @{verbatim [display] \isabelle jedit -Dsun.java2d.opengl=true\} If this works reliably, it can be made persistent via @{setting JEDIT_JAVA_OPTIONS} within \<^path>\$ISABELLE_HOME_USER/etc/settings\. For the Isabelle desktop ``app'', there is a corresponding file with Java runtime options in the main directory (name depends on the OS platform). \<^item> \<^bold>\Problem:\ Some Linux/X11 input methods such as IBus tend to disrupt key event handling of Java/AWT/Swing. \<^bold>\Workaround:\ Do not use X11 input methods. Note that environment variable \<^verbatim>\XMODIFIERS\ is reset by default within Isabelle settings. \<^item> \<^bold>\Problem:\ Some Linux/X11 window managers that are not ``re-parenting'' cause problems with additional windows opened by Java. This affects either historic or neo-minimalistic window managers like \<^verbatim>\awesome\ or \<^verbatim>\xmonad\. \<^bold>\Workaround:\ Use a regular re-parenting X11 window manager. \<^item> \<^bold>\Problem:\ Various forks of Linux/X11 window managers and desktop environments (like Gnome) disrupt the handling of menu popups and mouse positions of Java/AWT/Swing. \<^bold>\Workaround:\ Use suitable version of Linux desktops. \<^item> \<^bold>\Problem:\ Full-screen mode via jEdit action @{action_ref "toggle-full-screen"} (default keyboard shortcut \<^verbatim>\F11\ or \<^verbatim>\S+F11\) works robustly on Windows, but not on macOS or various Linux/X11 window managers. For the latter platforms, it is approximated by educated guesses on the window size (excluding the macOS menu bar). \<^bold>\Workaround:\ Use native full-screen control of the macOS window manager. \<^item> \<^bold>\Problem:\ Heap space of the JVM may fill up and render the Prover IDE unresponsive, e.g.\ when editing big Isabelle sessions with many theories. \<^bold>\Workaround:\ Increase JVM heap parameters by editing platform-specific files (for ``properties'' or ``options'') that are associated with the main app bundle. \ end diff --git a/src/Doc/Locales/Examples.thy b/src/Doc/Locales/Examples.thy --- a/src/Doc/Locales/Examples.thy +++ b/src/Doc/Locales/Examples.thy @@ -1,763 +1,763 @@ theory Examples imports Main begin (* text {* The following presentation will use notation of Isabelle's meta logic, hence a few sentences to explain this. The logical primitives are universal quantification (@{text "\"}), entailment (@{text "\"}) and equality (@{text "\"}). Variables (not bound variables) are sometimes preceded by a question mark. The logic is typed. Type variables are denoted by~@{text "'a"},~@{text "'b"} etc., and~@{text "\"} is the function type. Double brackets~@{text "\"} and~@{text "\"} are used to abbreviate nested entailment. *} *) section \Introduction\ text \ Locales are based on contexts. A \emph{context} can be seen as a formula schema \[ \\x\<^sub>1\x\<^sub>n. \ A\<^sub>1; \ ;A\<^sub>m \ \ \\ \] where the variables~\x\<^sub>1\, \ldots,~\x\<^sub>n\ are called \emph{parameters} and the premises $\A\<^sub>1\, \ldots,~\A\<^sub>m\$ \emph{assumptions}. A formula~\C\ is a \emph{theorem} in the context if it is a conclusion \[ \\x\<^sub>1\x\<^sub>n. \ A\<^sub>1; \ ;A\<^sub>m \ \ C\. \] Isabelle/Isar's notion of context goes beyond this logical view. Its contexts record, in a consecutive order, proved conclusions along with \emph{attributes}, which can provide context specific configuration information for proof procedures and concrete syntax. From a logical perspective, locales are just contexts that have been made persistent. To the user, though, they provide powerful means for declaring and combining contexts, and for the reuse of theorems proved in these contexts. \ section \Simple Locales\ text \ In its simplest form, a \emph{locale declaration} consists of a sequence of context elements declaring parameters (keyword \isakeyword{fixes}) and assumptions (keyword \isakeyword{assumes}). The following is the specification of partial orders, as locale \partial_order\. \ locale partial_order = fixes le :: "'a \ 'a \ bool" (infixl "\" 50) assumes refl [intro, simp]: "x \ x" and anti_sym [intro]: "\ x \ y; y \ x \ \ x = y" and trans [trans]: "\ x \ y; y \ z \ \ x \ z" text (in partial_order) \The parameter of this locale is~\le\, which is a binary predicate with infix syntax~\\\. The parameter syntax is available in the subsequent assumptions, which are the familiar partial order axioms. Isabelle recognises unbound names as free variables. In locale assumptions, these are implicitly universally quantified. That is, \<^term>\\ x \ y; y \ z \ \ x \ z\ in fact means \<^term>\\x y z. \ x \ y; y \ z \ \ x \ z\. Two commands are provided to inspect locales: \isakeyword{print\_locales} lists the names of all locales of the current theory; \isakeyword{print\_locale}~$n$ prints the parameters and assumptions of locale $n$; the variation \isakeyword{print\_locale!}~$n$ additionally outputs the conclusions that are stored in the locale. We may inspect the new locale by issuing \isakeyword{print\_locale!} \<^term>\partial_order\. The output is the following list of context elements. \begin{small} \begin{alltt} \isakeyword{fixes} le :: "'a \(\Rightarrow\) 'a \(\Rightarrow\) bool" (\isakeyword{infixl} "\(\sqsubseteq\)" 50) \isakeyword{assumes} "partial_order (\(\sqsubseteq\))" \isakeyword{notes} assumption refl [intro, simp] = `?x \(\sqsubseteq\) ?x` \isakeyword{and} anti_sym [intro] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?x\(\isasymrbrakk\) \(\Longrightarrow\) ?x = ?y` \isakeyword{and} trans [trans] = `\(\isasymlbrakk\)?x \(\sqsubseteq\) ?y; ?y \(\sqsubseteq\) ?z\(\isasymrbrakk\) \(\Longrightarrow\) ?x \(\sqsubseteq\) ?z` \end{alltt} \end{small} This differs from the declaration. The assumptions have turned into conclusions, denoted by the keyword \isakeyword{notes}. Also, there is only one assumption, namely \<^term>\partial_order le\. The locale declaration has introduced the predicate \<^term>\partial_order\ to the theory. This predicate is the \emph{locale predicate}. Its definition may be inspected by issuing \isakeyword{thm} @{thm [source] partial_order_def}. @{thm [display, indent=2] partial_order_def} In our example, this is a unary predicate over the parameter of the locale. It is equivalent to the original assumptions, which have been turned into conclusions and are available as theorems in the context of the locale. The names and attributes from the locale declaration are associated to these theorems and are effective in the context of the locale. Each conclusion has a \emph{foundational theorem} as counterpart in the theory. Technically, this is simply the theorem composed of context and conclusion. For the transitivity theorem, this is @{thm [source] partial_order.trans}: @{thm [display, indent=2] partial_order.trans} \ subsection \Targets: Extending Locales\ text \ The specification of a locale is fixed, but its list of conclusions may be extended through Isar commands that take a \emph{target} argument. In the following, \isakeyword{definition} and \isakeyword{theorem} are illustrated. Table~\ref{tab:commands-with-target} lists Isar commands that accept a target. Isar provides various ways of specifying the target. A target for a single command may be indicated with keyword \isakeyword{in} in the following way: \begin{table} \hrule \vspace{2ex} \begin{center} \begin{tabular}{ll} \isakeyword{definition} & definition through an equation \\ \isakeyword{inductive} & inductive definition \\ \isakeyword{primrec} & primitive recursion \\ \isakeyword{fun}, \isakeyword{function} & general recursion \\ \isakeyword{abbreviation} & syntactic abbreviation \\ \isakeyword{theorem}, etc.\ & theorem statement with proof \\ \isakeyword{theorems}, etc.\ & redeclaration of theorems \\ \isakeyword{text}, etc.\ & document markup \end{tabular} \end{center} \hrule \caption{Isar commands that accept a target.} \label{tab:commands-with-target} \end{table} \ definition (in partial_order) less :: "'a \ 'a \ bool" (infixl "\" 50) where "(x \ y) = (x \ y \ x \ y)" text (in partial_order) \The strict order \less\ with infix syntax~\\\ is defined in terms of the locale parameter~\le\ and the general equality of the object logic we work in. The definition generates a \emph{foundational constant} \<^term>\partial_order.less\ with definition @{thm [source] partial_order.less_def}: @{thm [display, indent=2] partial_order.less_def} At the same time, the locale is extended by syntax transformations hiding this construction in the context of the locale. Here, the abbreviation \less\ is available for \partial_order.less le\, and it is printed and parsed as infix~\\\. Finally, the conclusion @{thm [source] less_def} is added to the locale: @{thm [display, indent=2] less_def} \ text \The treatment of theorem statements is more straightforward. As an example, here is the derivation of a transitivity law for the strict order relation.\ lemma (in partial_order) less_le_trans [trans]: "\ x \ y; y \ z \ \ x \ z" unfolding %visible less_def by %visible (blast intro: trans) text \In the context of the proof, conclusions of the locale may be used like theorems. Attributes are effective: \anti_sym\ was declared as introduction rule, hence it is in the context's set of rules used by the classical reasoner by default.\ subsection \Context Blocks\ text \When working with locales, sequences of commands with the same target are frequent. A block of commands, delimited by \isakeyword{begin} and \isakeyword{end}, makes a theory-like style of working possible. All commands inside the block refer to the same target. A block may immediately follow a locale declaration, which makes that locale the target. Alternatively the target for a block may be given with the \isakeyword{context} command. This style of working is illustrated in the block below, where notions of infimum and supremum for partial orders are introduced, together with theorems about their uniqueness.\ context partial_order begin definition is_inf where "is_inf x y i = (i \ x \ i \ y \ (\z. z \ x \ z \ y \ z \ i))" definition is_sup where "is_sup x y s = (x \ s \ y \ s \ (\z. x \ z \ y \ z \ s \ z))" lemma %invisible is_infI [intro?]: "i \ x \ i \ y \ (\z. z \ x \ z \ y \ z \ i) \ is_inf x y i" by (unfold is_inf_def) blast lemma %invisible is_inf_lower [elim?]: "is_inf x y i \ (i \ x \ i \ y \ C) \ C" by (unfold is_inf_def) blast lemma %invisible is_inf_greatest [elim?]: "is_inf x y i \ z \ x \ z \ y \ z \ i" by (unfold is_inf_def) blast theorem is_inf_uniq: "\is_inf x y i; is_inf x y i'\ \ i = i'" proof - assume inf: "is_inf x y i" assume inf': "is_inf x y i'" show ?thesis proof (rule anti_sym) from inf' show "i \ i'" proof (rule is_inf_greatest) from inf show "i \ x" .. from inf show "i \ y" .. qed from inf show "i' \ i" proof (rule is_inf_greatest) from inf' show "i' \ x" .. from inf' show "i' \ y" .. qed qed qed theorem %invisible is_inf_related [elim?]: "x \ y \ is_inf x y x" proof - assume "x \ y" show ?thesis proof show "x \ x" .. show "x \ y" by fact fix z assume "z \ x" and "z \ y" show "z \ x" by fact qed qed lemma %invisible is_supI [intro?]: "x \ s \ y \ s \ (\z. x \ z \ y \ z \ s \ z) \ is_sup x y s" by (unfold is_sup_def) blast lemma %invisible is_sup_least [elim?]: "is_sup x y s \ x \ z \ y \ z \ s \ z" by (unfold is_sup_def) blast lemma %invisible is_sup_upper [elim?]: "is_sup x y s \ (x \ s \ y \ s \ C) \ C" by (unfold is_sup_def) blast theorem is_sup_uniq: "\is_sup x y s; is_sup x y s'\ \ s = s'" proof - assume sup: "is_sup x y s" assume sup': "is_sup x y s'" show ?thesis proof (rule anti_sym) from sup show "s \ s'" proof (rule is_sup_least) from sup' show "x \ s'" .. from sup' show "y \ s'" .. qed from sup' show "s' \ s" proof (rule is_sup_least) from sup show "x \ s" .. from sup show "y \ s" .. qed qed qed theorem %invisible is_sup_related [elim?]: "x \ y \ is_sup x y y" proof - assume "x \ y" show ?thesis proof show "x \ y" by fact show "y \ y" .. fix z assume "x \ z" and "y \ z" show "y \ z" by fact qed qed end text \The syntax of the locale commands discussed in this tutorial is shown in Table~\ref{tab:commands}. The grammar is complete with the exception of the context elements \isakeyword{constrains} and \isakeyword{defines}, which are provided for backward compatibility. See the Isabelle/Isar Reference - Manual @{cite IsarRef} for full documentation.\ + Manual \<^cite>\IsarRef\ for full documentation.\ section \Import \label{sec:import}\ text \ Algebraic structures are commonly defined by adding operations and properties to existing structures. For example, partial orders are extended to lattices and total orders. Lattices are extended to distributive lattices.\ text \ With locales, this kind of inheritance is achieved through \emph{import} of locales. The import part of a locale declaration, if present, precedes the context elements. Here is an example, where partial orders are extended to lattices. \ locale lattice = partial_order + assumes ex_inf: "\inf. is_inf x y inf" and ex_sup: "\sup. is_sup x y sup" begin text \These assumptions refer to the predicates for infimum and supremum defined for \partial_order\ in the previous section. We now introduce the notions of meet and join, together with an example theorem.\ definition meet (infixl "\" 70) where "x \ y = (THE inf. is_inf x y inf)" definition join (infixl "\" 65) where "x \ y = (THE sup. is_sup x y sup)" lemma %invisible meet_equality [elim?]: "is_inf x y i \ x \ y = i" proof (unfold meet_def) assume "is_inf x y i" then show "(THE i. is_inf x y i) = i" by (rule the_equality) (rule is_inf_uniq [OF _ \is_inf x y i\]) qed lemma %invisible meetI [intro?]: "i \ x \ i \ y \ (\z. z \ x \ z \ y \ z \ i) \ x \ y = i" by (rule meet_equality, rule is_infI) blast+ lemma %invisible is_inf_meet [intro?]: "is_inf x y (x \ y)" proof (unfold meet_def) from ex_inf obtain i where "is_inf x y i" .. then show "is_inf x y (THE i. is_inf x y i)" by (rule theI) (rule is_inf_uniq [OF _ \is_inf x y i\]) qed lemma meet_left(*<*)[intro?](*>*): "x \ y \ x" by (rule is_inf_lower) (rule is_inf_meet) lemma %invisible meet_right [intro?]: "x \ y \ y" by (rule is_inf_lower) (rule is_inf_meet) lemma %invisible meet_le [intro?]: "\ z \ x; z \ y \ \ z \ x \ y" by (rule is_inf_greatest) (rule is_inf_meet) lemma %invisible join_equality [elim?]: "is_sup x y s \ x \ y = s" proof (unfold join_def) assume "is_sup x y s" then show "(THE s. is_sup x y s) = s" by (rule the_equality) (rule is_sup_uniq [OF _ \is_sup x y s\]) qed lemma %invisible joinI [intro?]: "x \ s \ y \ s \ (\z. x \ z \ y \ z \ s \ z) \ x \ y = s" by (rule join_equality, rule is_supI) blast+ lemma %invisible is_sup_join [intro?]: "is_sup x y (x \ y)" proof (unfold join_def) from ex_sup obtain s where "is_sup x y s" .. then show "is_sup x y (THE s. is_sup x y s)" by (rule theI) (rule is_sup_uniq [OF _ \is_sup x y s\]) qed lemma %invisible join_left [intro?]: "x \ x \ y" by (rule is_sup_upper) (rule is_sup_join) lemma %invisible join_right [intro?]: "y \ x \ y" by (rule is_sup_upper) (rule is_sup_join) lemma %invisible join_le [intro?]: "\ x \ z; y \ z \ \ x \ y \ z" by (rule is_sup_least) (rule is_sup_join) theorem %invisible meet_assoc: "(x \ y) \ z = x \ (y \ z)" proof (rule meetI) show "x \ (y \ z) \ x \ y" proof show "x \ (y \ z) \ x" .. show "x \ (y \ z) \ y" proof - have "x \ (y \ z) \ y \ z" .. also have "\ \ y" .. finally show ?thesis . qed qed show "x \ (y \ z) \ z" proof - have "x \ (y \ z) \ y \ z" .. also have "\ \ z" .. finally show ?thesis . qed fix w assume "w \ x \ y" and "w \ z" show "w \ x \ (y \ z)" proof show "w \ x" proof - have "w \ x \ y" by fact also have "\ \ x" .. finally show ?thesis . qed show "w \ y \ z" proof show "w \ y" proof - have "w \ x \ y" by fact also have "\ \ y" .. finally show ?thesis . qed show "w \ z" by fact qed qed qed theorem %invisible meet_commute: "x \ y = y \ x" proof (rule meetI) show "y \ x \ x" .. show "y \ x \ y" .. fix z assume "z \ y" and "z \ x" then show "z \ y \ x" .. qed theorem %invisible meet_join_absorb: "x \ (x \ y) = x" proof (rule meetI) show "x \ x" .. show "x \ x \ y" .. fix z assume "z \ x" and "z \ x \ y" show "z \ x" by fact qed theorem %invisible join_assoc: "(x \ y) \ z = x \ (y \ z)" proof (rule joinI) show "x \ y \ x \ (y \ z)" proof show "x \ x \ (y \ z)" .. show "y \ x \ (y \ z)" proof - have "y \ y \ z" .. also have "... \ x \ (y \ z)" .. finally show ?thesis . qed qed show "z \ x \ (y \ z)" proof - have "z \ y \ z" .. also have "... \ x \ (y \ z)" .. finally show ?thesis . qed fix w assume "x \ y \ w" and "z \ w" show "x \ (y \ z) \ w" proof show "x \ w" proof - have "x \ x \ y" .. also have "\ \ w" by fact finally show ?thesis . qed show "y \ z \ w" proof show "y \ w" proof - have "y \ x \ y" .. also have "... \ w" by fact finally show ?thesis . qed show "z \ w" by fact qed qed qed theorem %invisible join_commute: "x \ y = y \ x" proof (rule joinI) show "x \ y \ x" .. show "y \ y \ x" .. fix z assume "y \ z" and "x \ z" then show "y \ x \ z" .. qed theorem %invisible join_meet_absorb: "x \ (x \ y) = x" proof (rule joinI) show "x \ x" .. show "x \ y \ x" .. fix z assume "x \ z" and "x \ y \ z" show "x \ z" by fact qed theorem %invisible meet_idem: "x \ x = x" proof - have "x \ (x \ (x \ x)) = x" by (rule meet_join_absorb) also have "x \ (x \ x) = x" by (rule join_meet_absorb) finally show ?thesis . qed theorem %invisible meet_related [elim?]: "x \ y \ x \ y = x" proof (rule meetI) assume "x \ y" show "x \ x" .. show "x \ y" by fact fix z assume "z \ x" and "z \ y" show "z \ x" by fact qed theorem %invisible meet_related2 [elim?]: "y \ x \ x \ y = y" by (drule meet_related) (simp add: meet_commute) theorem %invisible join_related [elim?]: "x \ y \ x \ y = y" proof (rule joinI) assume "x \ y" show "y \ y" .. show "x \ y" by fact fix z assume "x \ z" and "y \ z" show "y \ z" by fact qed theorem %invisible join_related2 [elim?]: "y \ x \ x \ y = x" by (drule join_related) (simp add: join_commute) theorem %invisible meet_connection: "(x \ y) = (x \ y = x)" proof assume "x \ y" then have "is_inf x y x" .. then show "x \ y = x" .. next have "x \ y \ y" .. also assume "x \ y = x" finally show "x \ y" . qed theorem %invisible join_connection: "(x \ y) = (x \ y = y)" proof assume "x \ y" then have "is_sup x y y" .. then show "x \ y = y" .. next have "x \ x \ y" .. also assume "x \ y = y" finally show "x \ y" . qed theorem %invisible meet_connection2: "(x \ y) = (y \ x = x)" using meet_commute meet_connection by simp theorem %invisible join_connection2: "(x \ y) = (x \ y = y)" using join_commute join_connection by simp text %invisible \Naming according to Jacobson I, p.\ 459.\ lemmas %invisible L1 = join_commute meet_commute lemmas %invisible L2 = join_assoc meet_assoc (* lemmas L3 = join_idem meet_idem *) lemmas %invisible L4 = join_meet_absorb meet_join_absorb end text \Locales for total orders and distributive lattices follow to establish a sufficiently rich landscape of locales for further examples in this tutorial.\ locale total_order = partial_order + assumes total: "x \ y \ y \ x" lemma (in total_order) less_total: "x \ y \ x = y \ y \ x" using total by (unfold less_def) blast locale distrib_lattice = lattice + assumes meet_distr: "x \ (y \ z) = x \ y \ x \ z" lemma (in distrib_lattice) join_distr: "x \ (y \ z) = (x \ y) \ (x \ z)" (* txt {* Jacobson I, p.\ 462 *} *) proof - have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by (simp add: L4) also have "... = x \ ((x \ z) \ (y \ z))" by (simp add: L2) also have "... = x \ ((x \ y) \ z)" by (simp add: L1 meet_distr) also have "... = ((x \ y) \ x) \ ((x \ y) \ z)" by (simp add: L1 L4) also have "... = (x \ y) \ (x \ z)" by (simp add: meet_distr) finally show ?thesis . qed text \ The locale hierarchy obtained through these declarations is shown in Figure~\ref{fig:lattices}(a). \begin{figure} \hrule \vspace{2ex} \begin{center} \subfigure[Declared hierarchy]{ \begin{tikzpicture} \node (po) at (0,0) {\partial_order\}; \node (lat) at (-1.5,-1) {\lattice\}; \node (dlat) at (-1.5,-2) {\distrib_lattice\}; \node (to) at (1.5,-1) {\total_order\}; \draw (po) -- (lat); \draw (lat) -- (dlat); \draw (po) -- (to); % \draw[->, dashed] (lat) -- (to); \end{tikzpicture} } \\ \subfigure[Total orders are lattices]{ \begin{tikzpicture} \node (po) at (0,0) {\partial_order\}; \node (lat) at (0,-1) {\lattice\}; \node (dlat) at (-1.5,-2) {\distrib_lattice\}; \node (to) at (1.5,-2) {\total_order\}; \draw (po) -- (lat); \draw (lat) -- (dlat); \draw (lat) -- (to); % \draw[->, dashed] (dlat) -- (to); \end{tikzpicture} } \quad \subfigure[Total orders are distributive lattices]{ \begin{tikzpicture} \node (po) at (0,0) {\partial_order\}; \node (lat) at (0,-1) {\lattice\}; \node (dlat) at (0,-2) {\distrib_lattice\}; \node (to) at (0,-3) {\total_order\}; \draw (po) -- (lat); \draw (lat) -- (dlat); \draw (dlat) -- (to); \end{tikzpicture} } \end{center} \hrule \caption{Hierarchy of Lattice Locales.} \label{fig:lattices} \end{figure} \ section \Changing the Locale Hierarchy \label{sec:changing-the-hierarchy}\ text \ Locales enable to prove theorems abstractly, relative to sets of assumptions. These theorems can then be used in other contexts where the assumptions themselves, or instances of the assumptions, are theorems. This form of theorem reuse is called \emph{interpretation}. Locales generalise interpretation from theorems to conclusions, enabling the reuse of definitions and other constructs that are not part of the specifications of the locales. The first form of interpretation we will consider in this tutorial is provided by the \isakeyword{sublocale} command. It enables to modify the import hierarchy to reflect the \emph{logical} relation between locales. Consider the locale hierarchy from Figure~\ref{fig:lattices}(a). Total orders are lattices, although this is not reflected here, and definitions, theorems and other conclusions from \<^term>\lattice\ are not available in \<^term>\total_order\. To obtain the situation in Figure~\ref{fig:lattices}(b), it is sufficient to add the conclusions of the latter locale to the former. The \isakeyword{sublocale} command does exactly this. The declaration \isakeyword{sublocale} $l_1 \subseteq l_2$ causes locale $l_2$ to be \emph{interpreted} in the context of $l_1$. This means that all conclusions of $l_2$ are made available in $l_1$. Of course, the change of hierarchy must be supported by a theorem that reflects, in our example, that total orders are indeed lattices. Therefore the \isakeyword{sublocale} command generates a goal, which must be discharged by the user. This is illustrated in the following paragraphs. First the sublocale relation is stated. \ sublocale %visible total_order \ lattice txt \\normalsize This enters the context of locale \total_order\, in which the goal @{subgoals [display]} must be shown. Now the locale predicate needs to be unfolded --- for example, using its definition or by introduction rules provided by the locale package. For automation, the locale package provides the methods \intro_locales\ and \unfold_locales\. They are aware of the current context and dependencies between locales and automatically discharge goals implied by these. While \unfold_locales\ always unfolds locale predicates to assumptions, \intro_locales\ only unfolds definitions along the locale hierarchy, leaving a goal consisting of predicates defined by the locale package. Occasionally the latter is of advantage since the goal is smaller. For the current goal, we would like to get hold of the assumptions of \lattice\, which need to be shown, hence \unfold_locales\ is appropriate.\ proof unfold_locales txt \\normalsize Since the fact that both lattices and total orders are partial orders is already reflected in the locale hierarchy, the assumptions of \partial_order\ are discharged automatically, and only the assumptions introduced in \lattice\ remain as subgoals @{subgoals [display]} The proof for the first subgoal is obtained by constructing an infimum, whose existence is implied by totality.\ fix x y from total have "is_inf x y (if x \ y then x else y)" by (auto simp: is_inf_def) then show "\inf. is_inf x y inf" .. txt \\normalsize The proof for the second subgoal is analogous and not reproduced here.\ next %invisible fix x y from total have "is_sup x y (if x \ y then y else x)" by (auto simp: is_sup_def) then show "\sup. is_sup x y sup" .. qed %visible text \Similarly, we may establish that total orders are distributive lattices with a second \isakeyword{sublocale} statement.\ sublocale total_order \ distrib_lattice proof unfold_locales fix %"proof" x y z show "x \ (y \ z) = x \ y \ x \ z" (is "?l = ?r") txt \Jacobson I, p.\ 462\ proof - { assume c: "y \ x" "z \ x" from c have "?l = y \ z" by (metis c join_connection2 join_related2 meet_related2 total) also from c have "... = ?r" by (metis meet_related2) finally have "?l = ?r" . } moreover { assume c: "x \ y \ x \ z" from c have "?l = x" by (metis join_connection2 join_related2 meet_connection total trans) also from c have "... = ?r" by (metis join_commute join_related2 meet_connection meet_related2 total) finally have "?l = ?r" . } moreover note total ultimately show ?thesis by blast qed qed text \The locale hierarchy is now as shown in Figure~\ref{fig:lattices}(c).\ text \ Locale interpretation is \emph{dynamic}. The statement \isakeyword{sublocale} $l_1 \subseteq l_2$ will not just add the current conclusions of $l_2$ to $l_1$. Rather the dependency is stored, and conclusions that will be added to $l_2$ in future are automatically propagated to $l_1$. The sublocale relation is transitive --- that is, propagation takes effect along chains of sublocales. Even cycles in the sublocale relation are supported, as long as these cycles do not lead to infinite chains. - Details are discussed in the technical report @{cite Ballarin2006a}. + Details are discussed in the technical report \<^cite>\Ballarin2006a\. See also Section~\ref{sec:infinite-chains} of this tutorial.\ end diff --git a/src/Doc/Locales/Examples3.thy b/src/Doc/Locales/Examples3.thy --- a/src/Doc/Locales/Examples3.thy +++ b/src/Doc/Locales/Examples3.thy @@ -1,644 +1,644 @@ theory Examples3 imports Examples begin subsection \Third Version: Local Interpretation \label{sec:local-interpretation}\ text \In the above example, the fact that \<^term>\(\)\ is a partial order for the integers was used in the second goal to discharge the premise in the definition of \(\)\. In general, proofs of the equations not only may involve definitions from the interpreted locale but arbitrarily complex arguments in the context of the locale. Therefore it would be convenient to have the interpreted locale conclusions temporarily available in the proof. This can be achieved by a locale interpretation in the proof body. The command for local interpretations is \isakeyword{interpret}. We repeat the example from the previous section to illustrate this.\ interpretation %visible int: partial_order "(\) :: int \ int \ bool" rewrites "int.less x y = (x < y)" proof - show "partial_order ((\) :: int \ int \ bool)" by unfold_locales auto then interpret int: partial_order "(\) :: [int, int] \ bool" . show "int.less x y = (x < y)" unfolding int.less_def by auto qed text \The inner interpretation is immediate from the preceding fact and proved by assumption (Isar short hand ``.''). It enriches the local proof context by the theorems also obtained in the interpretation from Section~\ref{sec:po-first}, and \int.less_def\ may directly be used to unfold the definition. Theorems from the local interpretation disappear after leaving the proof context --- that is, after the succeeding \isakeyword{next} or \isakeyword{qed} statement.\ subsection \Further Interpretations\ text \Further interpretations are necessary for the other locales. In \lattice\ the operations~\\\ and~\\\ are substituted by \<^term>\min :: int \ int \ int\ and \<^term>\max :: int \ int \ int\. The entire proof for the interpretation is reproduced to give an example of a more elaborate interpretation proof. Note that the equations are named so they can be used in a later example.\ interpretation %visible int: lattice "(\) :: int \ int \ bool" rewrites int_min_eq: "int.meet x y = min x y" and int_max_eq: "int.join x y = max x y" proof - show "lattice ((\) :: int \ int \ bool)" txt \\normalsize We have already shown that this is a partial order,\ apply unfold_locales txt \\normalsize hence only the lattice axioms remain to be shown. @{subgoals [display]} By \is_inf\ and \is_sup\,\ apply (unfold int.is_inf_def int.is_sup_def) txt \\normalsize the goals are transformed to these statements: @{subgoals [display]} This is Presburger arithmetic, which can be solved by the method \arith\.\ by arith+ txt \\normalsize In order to show the equations, we put ourselves in a situation where the lattice theorems can be used in a convenient way.\ then interpret int: lattice "(\) :: int \ int \ bool" . show "int.meet x y = min x y" by (bestsimp simp: int.meet_def int.is_inf_def) show "int.join x y = max x y" by (bestsimp simp: int.join_def int.is_sup_def) qed text \Next follows that \(\)\ is a total order, again for the integers.\ interpretation %visible int: total_order "(\) :: int \ int \ bool" by unfold_locales arith text \Theorems that are available in the theory at this point are shown in Table~\ref{tab:int-lattice}. Two points are worth noting: \begin{table} \hrule \vspace{2ex} \begin{center} \begin{tabular}{l} @{thm [source] int.less_def} from locale \partial_order\: \\ \quad @{thm int.less_def} \\ @{thm [source] int.meet_left} from locale \lattice\: \\ \quad @{thm int.meet_left} \\ @{thm [source] int.join_distr} from locale \distrib_lattice\: \\ \quad @{thm int.join_distr} \\ @{thm [source] int.less_total} from locale \total_order\: \\ \quad @{thm int.less_total} \end{tabular} \end{center} \hrule \caption{Interpreted theorems for~\\\ on the integers.} \label{tab:int-lattice} \end{table} \begin{itemize} \item Locale \distrib_lattice\ was also interpreted. Since the locale hierarchy reflects that total orders are distributive lattices, the interpretation of the latter was inserted automatically with the interpretation of the former. In general, interpretation traverses the locale hierarchy upwards and interprets all encountered locales, regardless whether imported or proved via the \isakeyword{sublocale} command. Existing interpretations are skipped avoiding duplicate work. \item The predicate \<^term>\(<)\ appears in theorem @{thm [source] int.less_total} although an equation for the replacement of \(\)\ was only given in the interpretation of \partial_order\. The interpretation equations are pushed downwards the hierarchy for related interpretations --- that is, for interpretations that share the instances of parameters they have in common. \end{itemize} \ text \The interpretations for a locale $n$ within the current theory may be inspected with \isakeyword{print\_interps}~$n$. This prints the list of instances of $n$, for which interpretations exist. For example, \isakeyword{print\_interps} \<^term>\partial_order\ outputs the following: \begin{small} \begin{alltt} int : partial_order "(\(\le\))" \end{alltt} \end{small} Of course, there is only one interpretation. The interpretation qualifier on the left is mandatory. Qualifiers can either be \emph{mandatory} or \emph{optional}. Optional qualifiers are designated by ``?''. Mandatory qualifiers must occur in name references while optional ones need not. Mandatory qualifiers prevent accidental hiding of names, while optional qualifiers can be more convenient to use. The default is mandatory. \ section \Locale Expressions \label{sec:expressions}\ text \ A map~\<^term>\\\ between partial orders~\\\ and~\\\ is called order preserving if \x \ y\ implies \\ x \ \ y\. This situation is more complex than those encountered so far: it involves two partial orders, and it is desirable to use the existing locale for both. A locale for order preserving maps requires three parameters: \le\~(\isakeyword{infixl}~\\\) and \le'\~(\isakeyword{infixl}~\\\) for the orders and~\\\ for the map. In order to reuse the existing locale for partial orders, which has the single parameter~\le\, it must be imported twice, once mapping its parameter to~\le\ from the new locale and once to~\le'\. This can be achieved with a compound locale expression. In general, a locale expression is a sequence of \emph{locale instances} separated by~``$\textbf{+}$'' and followed by a \isakeyword{for} clause. An instance has the following format: \begin{quote} \textit{qualifier} \textbf{:} \textit{locale-name} \textit{parameter-instantiation} \end{quote} We have already seen locale instances as arguments to \isakeyword{interpretation} in Section~\ref{sec:interpretation}. As before, the qualifier serves to disambiguate names from different instances of the same locale, and unless designated with a ``?'' it must occur in name references. Since the parameters~\le\ and~\le'\ are to be partial orders, our locale for order preserving maps will import the these instances: \begin{small} \begin{alltt} le: partial_order le le': partial_order le' \end{alltt} \end{small} For matter of convenience we choose to name parameter names and qualifiers alike. This is an arbitrary decision. Technically, qualifiers and parameters are unrelated. Having determined the instances, let us turn to the \isakeyword{for} clause. It serves to declare locale parameters in the same way as the context element \isakeyword{fixes} does. Context elements can only occur after the import section, and therefore the parameters referred to in the instances must be declared in the \isakeyword{for} clause. The \isakeyword{for} clause is also where the syntax of these parameters is declared. Two context elements for the map parameter~\\\ and the assumptions that it is order preserving complete the locale declaration.\ locale order_preserving = le: partial_order le + le': partial_order le' for le (infixl "\" 50) and le' (infixl "\" 50) + fixes \ assumes hom_le: "x \ y \ \ x \ \ y" text (in order_preserving) \Here are examples of theorems that are available in the locale: \hspace*{1em}@{thm [source] hom_le}: @{thm hom_le} \hspace*{1em}@{thm [source] le.less_le_trans}: @{thm le.less_le_trans} \hspace*{1em}@{thm [source] le'.less_le_trans}: @{thm [display, indent=4] le'.less_le_trans} While there is infix syntax for the strict operation associated with \<^term>\(\)\, there is none for the strict version of \<^term>\(\)\. The syntax \\\ for \less\ is only available for the original instance it was declared for. We may introduce infix syntax for \le'.less\ with the following declaration:\ notation (in order_preserving) le'.less (infixl "\" 50) text (in order_preserving) \Now the theorem is displayed nicely as @{thm [source] le'.less_le_trans}: @{thm [display, indent=2] le'.less_le_trans}\ text \There are short notations for locale expressions. These are discussed in the following.\ subsection \Default Instantiations\ text \ It is possible to omit parameter instantiations. The instantiation then defaults to the name of the parameter itself. For example, the locale expression \partial_order\ is short for \partial_order le\, since the locale's single parameter is~\le\. We took advantage of this in the \isakeyword{sublocale} declarations of Section~\ref{sec:changing-the-hierarchy}.\ subsection \Implicit Parameters \label{sec:implicit-parameters}\ text \In a locale expression that occurs within a locale declaration, omitted parameters additionally extend the (possibly empty) \isakeyword{for} clause. The \isakeyword{for} clause is a general construct of Isabelle/Isar to mark names occurring in the preceding declaration as ``arbitrary but fixed''. This is necessary for example, if the name is already bound in a surrounding context. In a locale expression, names occurring in parameter instantiations should be bound by a \isakeyword{for} clause whenever these names are not introduced elsewhere in the context --- for example, on the left hand side of a \isakeyword{sublocale} declaration. There is an exception to this rule in locale declarations, where the \isakeyword{for} clause serves to declare locale parameters. Here, locale parameters for which no parameter instantiation is given are implicitly added, with their mixfix syntax, at the beginning of the \isakeyword{for} clause. For example, in a locale declaration, the expression \partial_order\ is short for \begin{small} \begin{alltt} partial_order le \isakeyword{for} le (\isakeyword{infixl} "\(\sqsubseteq\)" 50)\textrm{.} \end{alltt} \end{small} This short hand was used in the locale declarations throughout Section~\ref{sec:import}. \ text\ The following locale declarations provide more examples. A map~\\\ is a lattice homomorphism if it preserves meet and join.\ locale lattice_hom = le: lattice + le': lattice le' for le' (infixl "\" 50) + fixes \ assumes hom_meet: "\ (x \ y) = le'.meet (\ x) (\ y)" and hom_join: "\ (x \ y) = le'.join (\ x) (\ y)" text \The parameter instantiation in the first instance of \<^term>\lattice\ is omitted. This causes the parameter~\le\ to be added to the \isakeyword{for} clause, and the locale has parameters~\le\,~\le'\ and, of course,~\\\. Before turning to the second example, we complete the locale by providing infix syntax for the meet and join operations of the second lattice. \ context lattice_hom begin notation le'.meet (infixl "\''" 50) notation le'.join (infixl "\''" 50) end text \The next example makes radical use of the short hand facilities. A homomorphism is an endomorphism if both orders coincide.\ locale lattice_end = lattice_hom _ le text \The notation~\_\ enables to omit a parameter in a positional instantiation. The omitted parameter,~\le\ becomes the parameter of the declared locale and is, in the following position, used to instantiate the second parameter of \lattice_hom\. The effect is that of identifying the first in second parameter of the homomorphism locale.\ text \The inheritance diagram of the situation we have now is shown in Figure~\ref{fig:hom}, where the dashed line depicts an interpretation which is introduced below. Parameter instantiations are indicated by $\sqsubseteq \mapsto \preceq$ etc. By looking at the inheritance diagram it would seem that two identical copies of each of the locales \partial_order\ and \lattice\ are imported by \lattice_end\. This is not the case! Inheritance paths with identical morphisms are automatically detected and the conclusions of the respective locales appear only once. \begin{figure} \hrule \vspace{2ex} \begin{center} \begin{tikzpicture} \node (o) at (0,0) {\partial_order\}; \node (oh) at (1.5,-2) {\order_preserving\}; \node (oh1) at (1.5,-0.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; \node (oh2) at (0,-1.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; \node (l) at (-1.5,-2) {\lattice\}; \node (lh) at (0,-4) {\lattice_hom\}; \node (lh1) at (0,-2.7) {$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; \node (lh2) at (-1.5,-3.3) {$\scriptscriptstyle \sqsubseteq \mapsto \preceq$}; \node (le) at (0,-6) {\lattice_end\}; \node (le1) at (0,-4.8) [anchor=west]{$\scriptscriptstyle \sqsubseteq \mapsto \sqsubseteq$}; \node (le2) at (0,-5.2) [anchor=west]{$\scriptscriptstyle \preceq \mapsto \sqsubseteq$}; \draw (o) -- (l); \draw[dashed] (oh) -- (lh); \draw (lh) -- (le); \draw (o) .. controls (oh1.south west) .. (oh); \draw (o) .. controls (oh2.north east) .. (oh); \draw (l) .. controls (lh1.south west) .. (lh); \draw (l) .. controls (lh2.north east) .. (lh); \end{tikzpicture} \end{center} \hrule \caption{Hierarchy of Homomorphism Locales.} \label{fig:hom} \end{figure} \ text \It can be shown easily that a lattice homomorphism is order preserving. As the final example of this section, a locale interpretation is used to assert this:\ sublocale lattice_hom \ order_preserving proof unfold_locales fix x y assume "x \ y" then have "y = (x \ y)" by (simp add: le.join_connection) then have "\ y = (\ x \' \ y)" by (simp add: hom_join [symmetric]) then show "\ x \ \ y" by (simp add: le'.join_connection) qed text (in lattice_hom) \ Theorems and other declarations --- syntax, in particular --- from the locale \order_preserving\ are now active in \lattice_hom\, for example @{thm [source] hom_le}: @{thm [display, indent=2] hom_le} This theorem will be useful in the following section. \ section \Conditional Interpretation\ text \There are situations where an interpretation is not possible in the general case since the desired property is only valid if certain conditions are fulfilled. Take, for example, the function \\i. n * i\ that scales its argument by a constant factor. This function is order preserving (and even a lattice endomorphism) with respect to \<^term>\(\)\ provided \n \ 0\. It is not possible to express this using a global interpretation, because it is in general unspecified whether~\<^term>\n\ is non-negative, but one may make an interpretation in an inner context of a proof where full information is available. This is not fully satisfactory either, since potentially interpretations may be required to make interpretations in many contexts. What is required is an interpretation that depends on the condition --- and this can be done with the \isakeyword{sublocale} command. For this purpose, we introduce a locale for the condition.\ locale non_negative = fixes n :: int assumes non_neg: "0 \ n" text \It is again convenient to make the interpretation in an incremental fashion, first for order preserving maps, then for lattice endomorphisms.\ sublocale non_negative \ order_preserving "(\)" "(\)" "\i. n * i" using non_neg by unfold_locales (rule mult_left_mono) text \While the proof of the previous interpretation is straightforward from monotonicity lemmas for~\<^term>\(*)\, the second proof follows a useful pattern.\ sublocale %visible non_negative \ lattice_end "(\)" "\i. n * i" proof (unfold_locales, unfold int_min_eq int_max_eq) txt \\normalsize Unfolding the locale predicates \emph{and} the interpretation equations immediately yields two subgoals that reflect the core conjecture. @{subgoals [display]} It is now necessary to show, in the context of \<^term>\non_negative\, that multiplication by~\<^term>\n\ commutes with \<^term>\min\ and \<^term>\max\.\ qed (auto simp: hom_le) text (in order_preserving) \The lemma @{thm [source] hom_le} simplifies a proof that would have otherwise been lengthy and we may consider making it a default rule for the simplifier:\ lemmas (in order_preserving) hom_le [simp] subsection \Avoiding Infinite Chains of Interpretations \label{sec:infinite-chains}\ text \Similar situations arise frequently in formalisations of abstract algebra where it is desirable to express that certain constructions preserve certain properties. For example, polynomials over rings are rings, or --- an example from the domain where the illustrations of this tutorial are taken from --- a partial order may be obtained for a function space by point-wise lifting of the partial order of the co-domain. This corresponds to the following interpretation:\ sublocale %visible partial_order \ f: partial_order "\f g. \x. f x \ g x" oops text \Unfortunately this is a cyclic interpretation that leads to an infinite chain, namely @{text [display, indent=2] "partial_order \ partial_order (\f g. \x. f x \ g x) \ partial_order (\f g. \x y. f x y \ g x y) \ \"} and the interpretation is rejected. Instead it is necessary to declare a locale that is logically equivalent to \<^term>\partial_order\ but serves to collect facts about functions spaces where the co-domain is a partial order, and to make the interpretation in its context:\ locale fun_partial_order = partial_order sublocale fun_partial_order \ f: partial_order "\f g. \x. f x \ g x" by unfold_locales (fast,rule,fast,blast intro: trans) text \It is quite common in abstract algebra that such a construction maps a hierarchy of algebraic structures (or specifications) to a related hierarchy. By means of the same lifting, a function space is a lattice if its co-domain is a lattice:\ locale fun_lattice = fun_partial_order + lattice sublocale fun_lattice \ f: lattice "\f g. \x. f x \ g x" proof unfold_locales fix f g have "partial_order.is_inf (\f g. \x. f x \ g x) f g (\x. f x \ g x)" apply (rule f.is_infI) apply rule+ apply (drule spec, assumption)+ done then show "\inf. partial_order.is_inf (\f g. \x. f x \ g x) f g inf" by fast next fix f g have "partial_order.is_sup (\f g. \x. f x \ g x) f g (\x. f x \ g x)" apply (rule f.is_supI) apply rule+ apply (drule spec, assumption)+ done then show "\sup. partial_order.is_sup (\f g. \x. f x \ g x) f g sup" by fast qed section \Further Reading\ text \More information on locales and their interpretation is available. For the locale hierarchy of import and interpretation - dependencies see~@{cite Ballarin2006a}; interpretations in theories - and proofs are covered in~@{cite Ballarin2006b}. In the latter, I + dependencies see~\<^cite>\Ballarin2006a\; interpretations in theories + and proofs are covered in~\<^cite>\Ballarin2006b\. In the latter, I show how interpretation in proofs enables to reason about families of algebraic structures, which cannot be expressed with locales directly. - Haftmann and Wenzel~@{cite HaftmannWenzel2007} overcome a restriction + Haftmann and Wenzel~\<^cite>\HaftmannWenzel2007\ overcome a restriction of axiomatic type classes through a combination with locale interpretation. The result is a Haskell-style class system with a facility to generate ML and Haskell code. Classes are sufficient for simple specifications with a single type parameter. The locales for orders and lattices presented in this tutorial fall into this category. Order preserving maps, homomorphisms and vector spaces, on the other hand, do not. The locales reimplementation for Isabelle 2009 provides, among other improvements, a clean integration with Isabelle/Isar's local theory mechanisms, which are described in another paper by Haftmann and - Wenzel~@{cite HaftmannWenzel2009}. + Wenzel~\<^cite>\HaftmannWenzel2009\. - The original work of Kammüller on locales~@{cite KammullerEtAl1999} + The original work of Kammüller on locales~\<^cite>\KammullerEtAl1999\ may be of interest from a historical perspective. My previous - report on locales and locale expressions~@{cite Ballarin2004a} + report on locales and locale expressions~\<^cite>\Ballarin2004a\ describes a simpler form of expressions than available now and is outdated. The mathematical background on orders and lattices is - taken from Jacobson's textbook on algebra~@{cite \Chapter~8\ Jacobson1985}. + taken from Jacobson's textbook on algebra~\<^cite>\\Chapter~8\ in Jacobson1985\. The sources of this tutorial, which include all proofs, are available with the Isabelle distribution at \<^url>\https://isabelle.in.tum.de\. \ text \ \begin{table} \hrule \vspace{2ex} \begin{center} \begin{tabular}{l>$c<$l} \multicolumn{3}{l}{Miscellaneous} \\ \textit{attr-name} & ::= & \textit{name} $|$ \textit{attribute} $|$ \textit{name} \textit{attribute} \\ \textit{qualifier} & ::= & \textit{name} [``\textbf{?}''] \\[2ex] \multicolumn{3}{l}{Context Elements} \\ \textit{fixes} & ::= & \textit{name} [ ``\textbf{::}'' \textit{type} ] [ ``\textbf{(}'' \textbf{structure} ``\textbf{)}'' $|$ \textit{mixfix} ] \\ \begin{comment} \textit{constrains} & ::= & \textit{name} ``\textbf{::}'' \textit{type} \\ \end{comment} \textit{assumes} & ::= & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ \begin{comment} \textit{defines} & ::= & [ \textit{attr-name} ``\textbf{:}'' ] \textit{proposition} \\ \textit{notes} & ::= & [ \textit{attr-name} ``\textbf{=}'' ] ( \textit{qualified-name} [ \textit{attribute} ] )$^+$ \\ \end{comment} \textit{element} & ::= & \textbf{fixes} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ \\ \begin{comment} & | & \textbf{constrains} \textit{constrains} ( \textbf{and} \textit{constrains} )$^*$ \\ \end{comment} & | & \textbf{assumes} \textit{assumes} ( \textbf{and} \textit{assumes} )$^*$ \\[2ex] %\begin{comment} % & | % & \textbf{defines} \textit{defines} ( \textbf{and} \textit{defines} )$^*$ \\ % & | % & \textbf{notes} \textit{notes} ( \textbf{and} \textit{notes} )$^*$ \\ %\end{comment} \multicolumn{3}{l}{Locale Expressions} \\ \textit{pos-insts} & ::= & ( \textit{term} $|$ ``\textbf{\_}'' )$^*$ \\ \textit{named-insts} & ::= & \textbf{where} \textit{name} ``\textbf{=}'' \textit{term} ( \textbf{and} \textit{name} ``\textbf{=}'' \textit{term} )$^*$ \\ \textit{instance} & ::= & [ \textit{qualifier} ``\textbf{:}'' ] \textit{name} ( \textit{pos-insts} $|$ \textit{named-inst} ) \\ \textit{expression} & ::= & \textit{instance} ( ``\textbf{+}'' \textit{instance} )$^*$ [ \textbf{for} \textit{fixes} ( \textbf{and} \textit{fixes} )$^*$ ] \\[2ex] \multicolumn{3}{l}{Declaration of Locales} \\ \textit{locale} & ::= & \textit{element}$^+$ \\ & | & \textit{expression} [ ``\textbf{+}'' \textit{element}$^+$ ] \\ \textit{toplevel} & ::= & \textbf{locale} \textit{name} [ ``\textbf{=}'' \textit{locale} ] \\[2ex] \multicolumn{3}{l}{Interpretation} \\ \textit{equation} & ::= & [ \textit{attr-name} ``\textbf{:}'' ] \textit{prop} \\ \textit{equations} & ::= & \textbf{rewrites} \textit{equation} ( \textbf{and} \textit{equation} )$^*$ \\ \textit{toplevel} & ::= & \textbf{sublocale} \textit{name} ( ``$<$'' $|$ ``$\subseteq$'' ) \textit{expression} \textit{proof} \\ & | & \textbf{interpretation} \textit{expression} [ \textit{equations} ] \textit{proof} \\ & | & \textbf{interpret} \textit{expression} \textit{proof} \\[2ex] \multicolumn{3}{l}{Diagnostics} \\ \textit{toplevel} & ::= & \textbf{print\_locales} \\ & | & \textbf{print\_locale} [ ``\textbf{!}'' ] \textit{name} \\ & | & \textbf{print\_interps} \textit{name} \end{tabular} \end{center} \hrule \caption{Syntax of Locale Commands.} \label{tab:commands} \end{table} \ text \\textbf{Revision History.} The present fourth revision of the tutorial accommodates minor updates to the syntax of the locale commands and the handling of notation under morphisms introduced with Isabelle 2016. For the third revision of the tutorial, which corresponds to the published version, much of the explanatory text was rewritten. Inheritance of interpretation equations was made available with Isabelle 2009-1. The second revision accommodates changes introduced by the locales reimplementation for Isabelle 2009. Most notably locale expressions were generalised from renaming to instantiation.\ text \\textbf{Acknowledgements.} Alexander Krauss, Tobias Nipkow, Randy Pollack, Andreas Schropp, Christian Sternagel and Makarius Wenzel have made useful comments on earlier versions of this document. The section on conditional interpretation was inspired by a number of e-mail enquiries the author received from locale users, and which suggested that this use case is important enough to deserve explicit explanation. The term \emph{conditional interpretation} is due to Larry Paulson.\ end diff --git a/src/Doc/Logics_ZF/ZF_Isar.thy b/src/Doc/Logics_ZF/ZF_Isar.thy --- a/src/Doc/Logics_ZF/ZF_Isar.thy +++ b/src/Doc/Logics_ZF/ZF_Isar.thy @@ -1,137 +1,137 @@ theory ZF_Isar imports ZF begin (*<*) ML_file \../antiquote_setup.ML\ (*>*) chapter \Some Isar language elements\ section \Type checking\ text \ The ZF logic is essentially untyped, so the concept of ``type checking'' is performed as logical reasoning about set-membership statements. A special method assists users in this task; a version of this is already declared as a ``solver'' in the standard Simplifier setup. \begin{matharray}{rcl} @{command_def (ZF) "print_tcset"}\\<^sup>*\ & : & \context \\ \\ @{method_def (ZF) typecheck} & : & \method\ \\ @{attribute_def (ZF) TC} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute (ZF) TC} (() | 'add' | 'del') \ \begin{description} \item @{command (ZF) "print_tcset"} prints the collection of typechecking rules of the current context. \item @{method (ZF) typecheck} attempts to solve any pending type-checking problems in subgoals. \item @{attribute (ZF) TC} adds or deletes type-checking rules from the context. \end{description} \ section \(Co)Inductive sets and datatypes\ subsection \Set definitions\ text \ In ZF everything is a set. The generic inductive package also provides a specific view for ``datatype'' specifications. Coinductive definitions are available in both cases, too. \begin{matharray}{rcl} @{command_def (ZF) "inductive"} & : & \theory \ theory\ \\ @{command_def (ZF) "coinductive"} & : & \theory \ theory\ \\ @{command_def (ZF) "datatype"} & : & \theory \ theory\ \\ @{command_def (ZF) "codatatype"} & : & \theory \ theory\ \\ \end{matharray} \<^rail>\ (@@{command (ZF) inductive} | @@{command (ZF) coinductive}) domains intros hints ; domains: @'domains' (@{syntax term} + '+') ('<=' | '\') @{syntax term} ; intros: @'intros' (@{syntax thmdecl}? @{syntax prop} +) ; hints: @{syntax (ZF) "monos"}? condefs? \ @{syntax (ZF) typeintros}? @{syntax (ZF) typeelims}? ; @{syntax_def (ZF) "monos"}: @'monos' @{syntax thms} ; condefs: @'con_defs' @{syntax thms} ; @{syntax_def (ZF) typeintros}: @'type_intros' @{syntax thms} ; @{syntax_def (ZF) typeelims}: @'type_elims' @{syntax thms} \ In the following syntax specification \monos\, \typeintros\, and \typeelims\ are the same as above. \<^rail>\ (@@{command (ZF) datatype} | @@{command (ZF) codatatype}) domain? (dtspec + @'and') hints ; domain: ('<=' | '\') @{syntax term} ; dtspec: @{syntax term} '=' (con + '|') ; con: @{syntax name} ('(' (@{syntax term} ',' +) ')')? ; hints: @{syntax (ZF) "monos"}? @{syntax (ZF) typeintros}? @{syntax (ZF) typeelims}? \ - See @{cite "isabelle-ZF"} for further information on inductive + See \<^cite>\"isabelle-ZF"\ for further information on inductive definitions in ZF, but note that this covers the old-style theory format. \ subsection \Primitive recursive functions\ text \ \begin{matharray}{rcl} @{command_def (ZF) "primrec"} & : & \theory \ theory\ \\ \end{matharray} \<^rail>\ @@{command (ZF) primrec} (@{syntax thmdecl}? @{syntax prop} +) \ \ subsection \Cases and induction: emulating tactic scripts\ text \ The following important tactical tools of Isabelle/ZF have been ported to Isar. These should not be used in proper proof texts. \begin{matharray}{rcl} @{method_def (ZF) case_tac}\\<^sup>*\ & : & \method\ \\ @{method_def (ZF) induct_tac}\\<^sup>*\ & : & \method\ \\ @{method_def (ZF) ind_cases}\\<^sup>*\ & : & \method\ \\ @{command_def (ZF) "inductive_cases"} & : & \theory \ theory\ \\ \end{matharray} \<^rail>\ (@@{method (ZF) case_tac} | @@{method (ZF) induct_tac}) @{syntax goal_spec}? @{syntax name} ; @@{method (ZF) ind_cases} (@{syntax prop} +) ; @@{command (ZF) inductive_cases} (@{syntax thmdecl}? (@{syntax prop} +) + @'and') \ \ end diff --git a/src/Doc/Prog_Prove/Isar.thy b/src/Doc/Prog_Prove/Isar.thy --- a/src/Doc/Prog_Prove/Isar.thy +++ b/src/Doc/Prog_Prove/Isar.thy @@ -1,1242 +1,1242 @@ (*<*) theory Isar imports LaTeXsugar begin declare [[quick_and_dirty]] (*>*) text\ Apply-scripts are unreadable and hard to maintain. The language of choice for larger proofs is \concept{Isar}. The two key features of Isar are: \begin{itemize} \item It is structured, not linear. \item It is readable without its being run because you need to state what you are proving at any given point. \end{itemize} Whereas apply-scripts are like assembly language programs, Isar proofs are like structured programs with comments. A typical Isar proof looks like this: \text\ \begin{tabular}{@ {}l} \isacom{proof}\\ \quad\isacom{assume} \"\$\mathit{formula}_0$\"\\\ \quad\isacom{have} \"\$\mathit{formula}_1$\"\ \quad\isacom{by} \simp\\\ \quad\vdots\\ \quad\isacom{have} \"\$\mathit{formula}_n$\"\ \quad\isacom{by} \blast\\\ \quad\isacom{show} \"\$\mathit{formula}_{n+1}$\"\ \quad\isacom{by} \\\\\ \isacom{qed} \end{tabular} \text\ It proves $\mathit{formula}_0 \Longrightarrow \mathit{formula}_{n+1}$ (provided each proof step succeeds). The intermediate \isacom{have} statements are merely stepping stones on the way towards the \isacom{show} statement that proves the actual goal. In more detail, this is the Isar core syntax: \medskip \begin{tabular}{@ {}lcl@ {}} \textit{proof} &=& \indexed{\isacom{by}}{by} \textit{method}\\ &$\mid$& \indexed{\isacom{proof}}{proof} [\textit{method}] \ \textit{step}$^*$ \ \indexed{\isacom{qed}}{qed} \end{tabular} \medskip \begin{tabular}{@ {}lcl@ {}} \textit{step} &=& \indexed{\isacom{fix}}{fix} \textit{variables} \\ &$\mid$& \indexed{\isacom{assume}}{assume} \textit{proposition} \\ &$\mid$& [\indexed{\isacom{from}}{from} \textit{fact}$^+$] (\indexed{\isacom{have}}{have} $\mid$ \indexed{\isacom{show}}{show}) \ \textit{proposition} \ \textit{proof} \end{tabular} \medskip \begin{tabular}{@ {}lcl@ {}} \textit{proposition} &=& [\textit{name}:] \"\\textit{formula}\"\ \end{tabular} \medskip \begin{tabular}{@ {}lcl@ {}} \textit{fact} &=& \textit{name} \ $\mid$ \ \dots \end{tabular} \medskip \noindent A proof can either be an atomic \isacom{by} with a single proof method which must finish off the statement being proved, for example \auto\, or it can be a \isacom{proof}--\isacom{qed} block of multiple steps. Such a block can optionally begin with a proof method that indicates how to start off the proof, e.g., \mbox{\(induction xs)\}. A step either assumes a proposition or states a proposition together with its proof. The optional \isacom{from} clause indicates which facts are to be used in the proof. Intermediate propositions are stated with \isacom{have}, the overall goal is stated with \isacom{show}. A step can also introduce new local variables with \isacom{fix}. Logically, \isacom{fix} introduces \\\-quantified variables, \isacom{assume} introduces the assumption of an implication (\\\) and \isacom{have}/\isacom{show} introduce the conclusion. Propositions are optionally named formulas. These names can be referred to in later \isacom{from} clauses. In the simplest case, a fact is such a name. But facts can also be composed with \OF\ and \of\ as shown in \autoref{sec:forward-proof} --- hence the \dots\ in the above grammar. Note that assumptions, intermediate \isacom{have} statements and global lemmas all have the same status and are thus collectively referred to as \conceptidx{facts}{fact}. Fact names can stand for whole lists of facts. For example, if \f\ is defined by command \isacom{fun}, \f.simps\ refers to the whole list of recursion equations defining \f\. Individual facts can be selected by writing \f.simps(2)\, whole sublists by writing \f.simps(2-4)\. \section{Isar by Example} We show a number of proofs of Cantor's theorem that a function from a set to its powerset cannot be surjective, illustrating various features of Isar. The constant \<^const>\surj\ is predefined. \ lemma "\ surj(f :: 'a \ 'a set)" proof assume 0: "surj f" from 0 have 1: "\A. \a. A = f a" by(simp add: surj_def) from 1 have 2: "\a. {x. x \ f x} = f a" by blast from 2 show "False" by blast qed text\ The \isacom{proof} command lacks an explicit method by which to perform the proof. In such cases Isabelle tries to use some standard introduction rule, in the above case for \\\: \[ \inferrule{ \mbox{@{thm (prem 1) notI}}} {\mbox{@{thm (concl) notI}}} \] In order to prove \<^prop>\~ P\, assume \P\ and show \False\. Thus we may assume \mbox{\noquotes{@{prop [source] "surj f"}}}. The proof shows that names of propositions may be (single!) digits --- meaningful names are hard to invent and are often not necessary. Both \isacom{have} steps are obvious. The second one introduces the diagonal set \<^term>\{x. x \ f x}\, the key idea in the proof. If you wonder why \2\ directly implies \False\: from \2\ it follows that \<^prop>\a \ f a \ a \ f a\. \subsection{\indexed{\this\}{this}, \indexed{\isacom{then}}{then}, \indexed{\isacom{hence}}{hence} and \indexed{\isacom{thus}}{thus}} Labels should be avoided. They interrupt the flow of the reader who has to scan the context for the point where the label was introduced. Ideally, the proof is a linear flow, where the output of one step becomes the input of the next step, piping the previously proved fact into the next proof, like in a UNIX pipe. In such cases the predefined name \this\ can be used to refer to the proposition proved in the previous step. This allows us to eliminate all labels from our proof (we suppress the \isacom{lemma} statement): \ (*<*) lemma "\ surj(f :: 'a \ 'a set)" (*>*) proof assume "surj f" from this have "\a. {x. x \ f x} = f a" by(auto simp: surj_def) from this show "False" by blast qed text\We have also taken the opportunity to compress the two \isacom{have} steps into one. To compact the text further, Isar has a few convenient abbreviations: \medskip \begin{tabular}{r@ {\quad=\quad}l} \isacom{then} & \isacom{from} \this\\\ \isacom{thus} & \isacom{then} \isacom{show}\\ \isacom{hence} & \isacom{then} \isacom{have} \end{tabular} \medskip \noindent With the help of these abbreviations the proof becomes \ (*<*) lemma "\ surj(f :: 'a \ 'a set)" (*>*) proof assume "surj f" hence "\a. {x. x \ f x} = f a" by(auto simp: surj_def) thus "False" by blast qed text\ There are two further linguistic variations: \medskip \begin{tabular}{r@ {\quad=\quad}l} (\isacom{have}$\mid$\isacom{show}) \ \textit{prop} \ \indexed{\isacom{using}}{using} \ \textit{facts} & \isacom{from} \ \textit{facts} \ (\isacom{have}$\mid$\isacom{show}) \ \textit{prop}\\ \indexed{\isacom{with}}{with} \ \textit{facts} & \isacom{from} \ \textit{facts} \isa{this} \end{tabular} \medskip \noindent The \isacom{using} idiom de-emphasizes the used facts by moving them behind the proposition. \subsection{Structured Lemma Statements: \indexed{\isacom{fixes}}{fixes}, \indexed{\isacom{assumes}}{assumes}, \indexed{\isacom{shows}}{shows}} \index{lemma@\isacom{lemma}} Lemmas can also be stated in a more structured fashion. To demonstrate this feature with Cantor's theorem, we rephrase \noquotes{@{prop[source]"\ surj f"}} a little: \ lemma fixes f :: "'a \ 'a set" assumes s: "surj f" shows "False" txt\The optional \isacom{fixes} part allows you to state the types of variables up front rather than by decorating one of their occurrences in the formula with a type constraint. The key advantage of the structured format is the \isacom{assumes} part that allows you to name each assumption; multiple assumptions can be separated by \isacom{and}. The \isacom{shows} part gives the goal. The actual theorem that will come out of the proof is \noquotes{@{prop[source]"surj f \ False"}}, but during the proof the assumption \noquotes{@{prop[source]"surj f"}} is available under the name \s\ like any other fact. \ proof - have "\ a. {x. x \ f x} = f a" using s by(auto simp: surj_def) thus "False" by blast qed text\ \begin{warn} Note the hyphen after the \isacom{proof} command. It is the null method that does nothing to the goal. Leaving it out would be asking Isabelle to try some suitable introduction rule on the goal \<^const>\False\ --- but there is no such rule and \isacom{proof} would fail. \end{warn} In the \isacom{have} step the assumption \noquotes{@{prop[source]"surj f"}} is now referenced by its name \s\. The duplication of \noquotes{@{prop[source]"surj f"}} in the above proofs (once in the statement of the lemma, once in its proof) has been eliminated. Stating a lemma with \isacom{assumes}-\isacom{shows} implicitly introduces the name \indexed{\assms\}{assms} that stands for the list of all assumptions. You can refer to individual assumptions by \assms(1)\, \assms(2)\, etc., thus obviating the need to name them individually. \section{Proof Patterns} We show a number of important basic proof patterns. Many of them arise from the rules of natural deduction that are applied by \isacom{proof} by default. The patterns are phrased in terms of \isacom{show} but work for \isacom{have} and \isacom{lemma}, too. \ifsem\else \subsection{Logic} \fi We start with two forms of \concept{case analysis}: starting from a formula \P\ we have the two cases \P\ and \<^prop>\~P\, and starting from a fact \<^prop>\P \ Q\ we have the two cases \P\ and \Q\: \text_raw\ \begin{tabular}{@ {}ll@ {}} \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "R" proof-(*>*) show "R" proof cases assume "P" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "R" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ next assume "\ P" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "R" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed(*<*)oops(*>*) text_raw \} \end{minipage}\index{cases@\cases\} & \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "R" proof-(*>*) have "P \ Q" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ then show "R" proof assume "P" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "R" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ next assume "Q" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "R" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed(*<*)oops(*>*) text_raw \} \end{minipage} \end{tabular} \medskip \begin{isamarkuptext}% How to prove a logical equivalence: \end{isamarkuptext}% \isa{% \ (*<*)lemma "P\Q" proof-(*>*) show "P \ Q" proof assume "P" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "Q" (*<*)sorry(*>*) text_raw\\ \isasymproof\\\ next assume "Q" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "P" (*<*)sorry(*>*) text_raw\\ \isasymproof\\\ qed(*<*)qed(*>*) text_raw \} \medskip \begin{isamarkuptext}% Proofs by contradiction (@{thm[source] ccontr} stands for ``classical contradiction''): \end{isamarkuptext}% \begin{tabular}{@ {}ll@ {}} \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "\ P" proof-(*>*) show "\ P" proof assume "P" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "False" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed(*<*)oops(*>*) text_raw \} \end{minipage} & \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "P" proof-(*>*) show "P" proof (rule ccontr) assume "\P" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "False" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed(*<*)oops(*>*) text_raw \} \end{minipage} \end{tabular} \medskip \begin{isamarkuptext}% How to prove quantified formulas: \end{isamarkuptext}% \begin{tabular}{@ {}ll@ {}} \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "\x. P x" proof-(*>*) show "\x. P(x)" proof fix x text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "P(x)" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed(*<*)oops(*>*) text_raw \} \end{minipage} & \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "\x. P(x)" proof-(*>*) show "\x. P(x)" proof text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "P(witness)" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed (*<*)oops(*>*) text_raw \} \end{minipage} \end{tabular} \medskip \begin{isamarkuptext}% In the proof of \noquotes{@{prop[source]"\x. P(x)"}}, the step \indexed{\isacom{fix}}{fix}~\x\ introduces a locally fixed variable \x\ into the subproof, the proverbial ``arbitrary but fixed value''. Instead of \x\ we could have chosen any name in the subproof. In the proof of \noquotes{@{prop[source]"\x. P(x)"}}, \witness\ is some arbitrary term for which we can prove that it satisfies \P\. How to reason forward from \noquotes{@{prop[source] "\x. P(x)"}}: \end{isamarkuptext}% \ (*<*)lemma True proof- assume 1: "\x. P x"(*>*) have "\x. P(x)" (*<*)by(rule 1)(*>*)text_raw\\ \isasymproof\\\ then obtain x where p: "P(x)" by blast (*<*)oops(*>*) text\ After the \indexed{\isacom{obtain}}{obtain} step, \x\ (we could have chosen any name) is a fixed local variable, and \p\ is the name of the fact \noquotes{@{prop[source] "P(x)"}}. This pattern works for one or more \x\. As an example of the \isacom{obtain} command, here is the proof of Cantor's theorem in more detail: \ lemma "\ surj(f :: 'a \ 'a set)" proof assume "surj f" hence "\a. {x. x \ f x} = f a" by(auto simp: surj_def) then obtain a where "{x. x \ f x} = f a" by blast hence "a \ f a \ a \ f a" by blast thus "False" by blast qed text_raw\ \begin{isamarkuptext}% Finally, how to prove set equality and subset relationship: \end{isamarkuptext}% \begin{tabular}{@ {}ll@ {}} \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "A = (B::'a set)" proof-(*>*) show "A = B" proof show "A \ B" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ next show "B \ A" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed(*<*)qed(*>*) text_raw \} \end{minipage} & \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "A <= (B::'a set)" proof-(*>*) show "A \ B" proof fix x assume "x \ A" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "x \ B" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed(*<*)qed(*>*) text_raw \} \end{minipage} \end{tabular} \begin{isamarkuptext}% \ifsem\else \subsection{Chains of (In)Equations} In textbooks, chains of equations (and inequations) are often displayed like this: \begin{quote} \begin{tabular}{@ {}l@ {\qquad}l@ {}} $t_1 = t_2$ & \isamath{\,\langle\mathit{justification}\rangle}\\ $\phantom{t_1} = t_3$ & \isamath{\,\langle\mathit{justification}\rangle}\\ \quad $\vdots$\\ $\phantom{t_1} = t_n$ & \isamath{\,\langle\mathit{justification}\rangle} \end{tabular} \end{quote} The Isar equivalent is this: \begin{samepage} \begin{quote} \isacom{have} \"t\<^sub>1 = t\<^sub>2"\ \isasymproof\\ \isacom{also have} \"... = t\<^sub>3"\ \isasymproof\\ \quad $\vdots$\\ \isacom{also have} \"... = t\<^sub>n"\ \isasymproof \\ \isacom{finally show} \"t\<^sub>1 = t\<^sub>n"\\ \texttt{.} \end{quote} \end{samepage} \noindent The ``\...\'' and ``\.\'' deserve some explanation: \begin{description} \item[``\...\''] is literally three dots. It is the name of an unknown that Isar automatically instantiates with the right-hand side of the previous equation. In general, if \this\ is the theorem \<^term>\p t\<^sub>1 t\<^sub>2\ then ``\...\'' stands for \t\<^sub>2\. \item[``\.\''] (a single dot) is a proof method that solves a goal by one of the assumptions. This works here because the result of \isacom{finally} is the theorem \mbox{\t\<^sub>1 = t\<^sub>n\}, \isacom{show} \"t\<^sub>1 = t\<^sub>n"\ states the theorem explicitly, and ``\.\'' proves the theorem with the result of \isacom{finally}. \end{description} The above proof template also works for arbitrary mixtures of \=\, \\\ and \<\, for example: \begin{quote} \isacom{have} \"t\<^sub>1 < t\<^sub>2"\ \isasymproof\\ \isacom{also have} \"... = t\<^sub>3"\ \isasymproof\\ \quad $\vdots$\\ \isacom{also have} \"... \ t\<^sub>n"\ \isasymproof \\ \isacom{finally show} \"t\<^sub>1 < t\<^sub>n"\\ \texttt{.} \end{quote} The relation symbol in the \isacom{finally} step needs to be the most precise one possible. In the example above, you must not write \t\<^sub>1 \ t\<^sub>n\ instead of \mbox{\t\<^sub>1 < t\<^sub>n\}. \begin{warn} Isabelle only supports \=\, \\\ and \<\ but not \\\ and \>\ in (in)equation chains (by default). \end{warn} If you want to go beyond merely using the above proof patterns and want to understand what \isacom{also} and \isacom{finally} mean, read on. There is an Isar theorem variable called \calculation\, similar to \this\. When the first \isacom{also} in a chain is encountered, Isabelle sets \calculation := this\. In each subsequent \isacom{also} step, Isabelle composes the theorems \calculation\ and \this\ (i.e.\ the two previous (in)equalities) using some predefined set of rules including transitivity of \=\, \\\ and \<\ but also mixed rules like \<^prop>\\ x \ y; y < z \ \ x < z\. The result of this composition is assigned to \calculation\. Consider \begin{quote} \isacom{have} \"t\<^sub>1 \ t\<^sub>2"\ \isasymproof\\ \isacom{also} \isacom{have} \"... < t\<^sub>3"\ \isasymproof\\ \isacom{also} \isacom{have} \"... = t\<^sub>4"\ \isasymproof\\ \isacom{finally show} \"t\<^sub>1 < t\<^sub>4"\\ \texttt{.} \end{quote} After the first \isacom{also}, \calculation\ is \"t\<^sub>1 \ t\<^sub>2"\, and after the second \isacom{also}, \calculation\ is \"t\<^sub>1 < t\<^sub>3"\. The command \isacom{finally} is short for \isacom{also from} \calculation\. Therefore the \isacom{also} hidden in \isacom{finally} sets \calculation\ to \t\<^sub>1 < t\<^sub>4\ and the final ``\texttt{.}'' succeeds. -For more information on this style of proof see @{cite "BauerW-TPHOLs01"}. +For more information on this style of proof see \<^cite>\"BauerW-TPHOLs01"\. \fi \section{Streamlining Proofs} \subsection{Pattern Matching and Quotations} In the proof patterns shown above, formulas are often duplicated. This can make the text harder to read, write and maintain. Pattern matching is an abbreviation mechanism to avoid such duplication. Writing \begin{quote} \isacom{show} \ \textit{formula} \(\\indexed{\isacom{is}}{is} \textit{pattern}\)\ \end{quote} matches the pattern against the formula, thus instantiating the unknowns in the pattern for later use. As an example, consider the proof pattern for \\\: \end{isamarkuptext}% \ (*<*)lemma "formula\<^sub>1 \ formula\<^sub>2" proof-(*>*) show "formula\<^sub>1 \ formula\<^sub>2" (is "?L \ ?R") proof assume "?L" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "?R" (*<*)sorry(*>*) text_raw\\ \isasymproof\\\ next assume "?R" text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show "?L" (*<*)sorry(*>*) text_raw\\ \isasymproof\\\ qed(*<*)qed(*>*) text\Instead of duplicating \formula\<^sub>i\ in the text, we introduce the two abbreviations \?L\ and \?R\ by pattern matching. Pattern matching works wherever a formula is stated, in particular with \isacom{have} and \isacom{lemma}. The unknown \indexed{\?thesis\}{thesis} is implicitly matched against any goal stated by \isacom{lemma} or \isacom{show}. Here is a typical example:\ lemma "formula" proof - text_raw\\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\ show ?thesis (*<*)sorry(*>*) text_raw\\ \isasymproof\\\ qed text\ Unknowns can also be instantiated with \indexed{\isacom{let}}{let} commands \begin{quote} \isacom{let} \?t\ = \"\\textit{some-big-term}\"\ \end{quote} Later proof steps can refer to \?t\: \begin{quote} \isacom{have} \"\\dots \?t\ \dots\"\ \end{quote} \begin{warn} Names of facts are introduced with \name:\ and refer to proved theorems. Unknowns \?X\ refer to terms or formulas. \end{warn} Although abbreviations shorten the text, the reader needs to remember what they stand for. Similarly for names of facts. Names like \1\, \2\ and \3\ are not helpful and should only be used in short proofs. For longer proofs, descriptive names are better. But look at this example: \begin{quote} \isacom{have} \ \x_gr_0: "x > 0"\\\ $\vdots$\\ \isacom{from} \x_gr_0\ \dots \end{quote} The name is longer than the fact it stands for! Short facts do not need names; one can refer to them easily by quoting them: \begin{quote} \isacom{have} \ \"x > 0"\\\ $\vdots$\\ \isacom{from} \`x>0`\ \dots\index{$IMP053@\`...`\} \end{quote} Note that the quotes around \x>0\ are \conceptnoidx{back quotes}. They refer to the fact not by name but by value. \subsection{\indexed{\isacom{moreover}}{moreover}} \index{ultimately@\isacom{ultimately}} Sometimes one needs a number of facts to enable some deduction. Of course one can name these facts individually, as shown on the right, but one can also combine them with \isacom{moreover}, as shown on the left: \text_raw\ \begin{tabular}{@ {}ll@ {}} \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "P" proof-(*>*) have "P\<^sub>1" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ moreover have "P\<^sub>2" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ moreover text_raw\\\$\vdots$\\\hspace{-1.4ex}\(*<*)have "True" ..(*>*) moreover have "P\<^sub>n" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ ultimately have "P" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ (*<*)oops(*>*) text_raw \} \end{minipage} & \qquad \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "P" proof-(*>*) have lab\<^sub>1: "P\<^sub>1" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ have lab\<^sub>2: "P\<^sub>2" (*<*)sorry(*>*)text_raw\\ \isasymproof\ text_raw\\\$\vdots$\\\hspace{-1.4ex}\ have lab\<^sub>n: "P\<^sub>n" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ from lab\<^sub>1 lab\<^sub>2 text_raw\\ $\dots$\\\ have "P" (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ (*<*)oops(*>*) text_raw \} \end{minipage} \end{tabular} \begin{isamarkuptext}% The \isacom{moreover} version is no shorter but expresses the structure a bit more clearly and avoids new names. \subsection{Local Lemmas} Sometimes one would like to prove some lemma locally within a proof, a lemma that shares the current context of assumptions but that has its own assumptions and is generalized over its locally fixed variables at the end. This is simply an extension of the basic \indexed{\isacom{have}}{have} construct: \begin{quote} \indexed{\isacom{have}}{have} \B\\ \indexed{\isacom{if}}{if} \name:\ \A\<^sub>1 \ A\<^sub>m\\ \indexed{\isacom{for}}{for} \x\<^sub>1 \ x\<^sub>n\\\ \isasymproof \end{quote} proves \\ A\<^sub>1; \ ; A\<^sub>m \ \ B\ where all \x\<^sub>i\ have been replaced by unknowns \?x\<^sub>i\. As an example we prove a simple fact about divisibility on integers. The definition of \dvd\ is @{thm dvd_def}. \end{isamarkuptext}% \ lemma fixes a b :: int assumes "b dvd (a+b)" shows "b dvd a" proof - have "\k'. a = b*k'" if asm: "a+b = b*k" for k proof show "a = b*(k - 1)" using asm by(simp add: algebra_simps) qed then show ?thesis using assms by(auto simp add: dvd_def) qed text\ \subsection*{Exercises} \exercise Give a readable, structured proof of the following lemma: \ lemma assumes T: "\x y. T x y \ T y x" and A: "\x y. A x y \ A y x \ x = y" and TA: "\x y. T x y \ A x y" and "A x y" shows "T x y" (*<*)oops(*>*) text\ \endexercise \exercise Give a readable, structured proof of the following lemma: \ lemma "\ys zs. xs = ys @ zs \ (length ys = length zs \ length ys = length zs + 1)" (*<*)oops(*>*) text\ Hint: There are predefined functions @{const_typ take} and @{const_typ drop} such that \take k [x\<^sub>1,\] = [x\<^sub>1,\,x\<^sub>k]\ and \drop k [x\<^sub>1,\] = [x\<^bsub>k+1\<^esub>,\]\. Let sledgehammer find and apply the relevant \<^const>\take\ and \<^const>\drop\ lemmas for you. \endexercise \section{Case Analysis and Induction} \subsection{Datatype Case Analysis} \index{case analysis|(} We have seen case analysis on formulas. Now we want to distinguish which form some term takes: is it \0\ or of the form \<^term>\Suc n\, is it \<^term>\[]\ or of the form \<^term>\x#xs\, etc. Here is a typical example proof by case analysis on the form of \xs\: \ lemma "length(tl xs) = length xs - 1" proof (cases xs) assume "xs = []" thus ?thesis by simp next fix y ys assume "xs = y#ys" thus ?thesis by simp qed text\\index{cases@\cases\|(}Function \tl\ (''tail'') is defined by @{thm list.sel(2)} and @{thm list.sel(3)}. Note that the result type of \<^const>\length\ is \<^typ>\nat\ and \<^prop>\0 - 1 = (0::nat)\. This proof pattern works for any term \t\ whose type is a datatype. The goal has to be proved for each constructor \C\: \begin{quote} \isacom{fix} \ \x\<^sub>1 \ x\<^sub>n\ \isacom{assume} \"t = C x\<^sub>1 \ x\<^sub>n"\ \end{quote}\index{case@\isacom{case}|(} Each case can be written in a more compact form by means of the \isacom{case} command: \begin{quote} \isacom{case} \(C x\<^sub>1 \ x\<^sub>n)\ \end{quote} This is equivalent to the explicit \isacom{fix}-\isacom{assume} line but also gives the assumption \"t = C x\<^sub>1 \ x\<^sub>n"\ a name: \C\, like the constructor. Here is the \isacom{case} version of the proof above: \ (*<*)lemma "length(tl xs) = length xs - 1"(*>*) proof (cases xs) case Nil thus ?thesis by simp next case (Cons y ys) thus ?thesis by simp qed text\Remember that \Nil\ and \Cons\ are the alphanumeric names for \[]\ and \#\. The names of the assumptions are not used because they are directly piped (via \isacom{thus}) into the proof of the claim. \index{case analysis|)} \subsection{Structural Induction} \index{induction|(} \index{structural induction|(} We illustrate structural induction with an example based on natural numbers: the sum (\\\) of the first \n\ natural numbers (\{0..n::nat}\) is equal to \mbox{\<^term>\n*(n+1) div 2::nat\}. Never mind the details, just focus on the pattern: \ lemma "\{0..n::nat} = n*(n+1) div 2" proof (induction n) show "\{0..0::nat} = 0*(0+1) div 2" by simp next fix n assume "\{0..n::nat} = n*(n+1) div 2" thus "\{0..Suc n} = Suc n*(Suc n+1) div 2" by simp qed text\Except for the rewrite steps, everything is explicitly given. This makes the proof easily readable, but the duplication means it is tedious to write and maintain. Here is how pattern matching can completely avoid any duplication:\ lemma "\{0..n::nat} = n*(n+1) div 2" (is "?P n") proof (induction n) show "?P 0" by simp next fix n assume "?P n" thus "?P(Suc n)" by simp qed text\The first line introduces an abbreviation \?P n\ for the goal. Pattern matching \?P n\ with the goal instantiates \?P\ to the function \<^term>\\n. \{0..n::nat} = n*(n+1) div 2\. Now the proposition to be proved in the base case can be written as \?P 0\, the induction hypothesis as \?P n\, and the conclusion of the induction step as \?P(Suc n)\. Induction also provides the \isacom{case} idiom that abbreviates the \isacom{fix}-\isacom{assume} step. The above proof becomes \ (*<*)lemma "\{0..n::nat} = n*(n+1) div 2"(*>*) proof (induction n) case 0 show ?case by simp next case (Suc n) thus ?case by simp qed text\ The unknown \?case\\index{case?@\?case\|(} is set in each case to the required claim, i.e., \?P 0\ and \mbox{\?P(Suc n)\} in the above proof, without requiring the user to define a \?P\. The general pattern for induction over \<^typ>\nat\ is shown on the left-hand side: \text_raw\ \begin{tabular}{@ {}ll@ {}} \begin{minipage}[t]{.4\textwidth} \isa{% \ (*<*)lemma "P(n::nat)" proof -(*>*) show "P(n)" proof (induction n) case 0 text_raw\\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}\ show ?case (*<*)sorry(*>*) text_raw\\ \isasymproof\\\ next case (Suc n) text_raw\\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}\ show ?case (*<*)sorry(*>*) text_raw\\ \isasymproof\\\ qed(*<*)qed(*>*) text_raw \} \end{minipage} & \begin{minipage}[t]{.4\textwidth} ~\\ ~\\ \isacom{let} \?case = "P(0)"\\\ ~\\ ~\\ ~\\[1ex] \isacom{fix} \n\ \isacom{assume} \Suc: "P(n)"\\\ \isacom{let} \?case = "P(Suc n)"\\\ \end{minipage} \end{tabular} \medskip \ text\ On the right side you can see what the \isacom{case} command on the left stands for. In case the goal is an implication, induction does one more thing: the proposition to be proved in each case is not the whole implication but only its conclusion; the premises of the implication are immediately made assumptions of that case. That is, if in the above proof we replace \isacom{show}~\"P(n)"\ by \mbox{\isacom{show}~\"A(n) \ P(n)"\} then \isacom{case}~\0\ stands for \begin{quote} \isacom{assume} \ \0: "A(0)"\\\ \isacom{let} \?case = "P(0)"\ \end{quote} and \isacom{case}~\(Suc n)\ stands for \begin{quote} \isacom{fix} \n\\\ \isacom{assume} \Suc:\ \begin{tabular}[t]{l}\"A(n) \ P(n)"\\\\"A(Suc n)"\\end{tabular}\\ \isacom{let} \?case = "P(Suc n)"\ \end{quote} The list of assumptions \Suc\ is actually subdivided into \Suc.IH\, the induction hypotheses (here \A(n) \ P(n)\), and \Suc.prems\, the premises of the goal being proved (here \A(Suc n)\). Induction works for any datatype. Proving a goal \\ A\<^sub>1(x); \; A\<^sub>k(x) \ \ P(x)\ by induction on \x\ generates a proof obligation for each constructor \C\ of the datatype. The command \isacom{case}~\(C x\<^sub>1 \ x\<^sub>n)\ performs the following steps: \begin{enumerate} \item \isacom{fix} \x\<^sub>1 \ x\<^sub>n\ \item \isacom{assume} the induction hypotheses (calling them \C.IH\\index{IH@\.IH\}) and the premises \mbox{\A\<^sub>i(C x\<^sub>1 \ x\<^sub>n)\} (calling them \C.prems\\index{prems@\.prems\}) and calling the whole list \C\ \item \isacom{let} \?case = "P(C x\<^sub>1 \ x\<^sub>n)"\ \end{enumerate} \index{structural induction|)} \ifsem\else \subsection{Computation Induction} \index{rule induction} In \autoref{sec:recursive-funs} we introduced computation induction and its realization in Isabelle: the definition of a recursive function \f\ via \isacom{fun} proves the corresponding computation induction rule called \f.induct\. Induction with this rule looks like in \autoref{sec:recursive-funs}, but now with \isacom{proof} instead of \isacom{apply}: \begin{quote} \isacom{proof} (\induction x\<^sub>1 \ x\<^sub>k rule: f.induct\) \end{quote} Just as for structural induction, this creates several cases, one for each defining equation for \f\. By default (if the equations have not been named by the user), the cases are numbered. That is, they are started by \begin{quote} \isacom{case} (\i x y ...\) \end{quote} where \i = 1,...,n\, \n\ is the number of equations defining \f\, and \x y ...\ are the variables in equation \i\. Note the following: \begin{itemize} \item Although \i\ is an Isar name, \i.IH\ (or similar) is not. You need double quotes: "\i.IH\". When indexing the name, write "\i.IH\"(1), not "\i.IH\(1)". \item If defining equations for \f\ overlap, \isacom{fun} instantiates them to make them nonoverlapping. This means that one user-provided equation may lead to several equations and thus to several cases in the induction rule. These have names of the form "\i_j\", where \i\ is the number of the original equation and the system-generated \j\ indicates the subcase. \end{itemize} In Isabelle/jEdit, the \induction\ proof method displays a proof skeleton with all \isacom{case}s. This is particularly useful for computation induction and the following rule induction. \fi \subsection{Rule Induction} \index{rule induction|(} Recall the inductive and recursive definitions of even numbers in \autoref{sec:inductive-defs}: \ inductive ev :: "nat \ bool" where ev0: "ev 0" | evSS: "ev n \ ev(Suc(Suc n))" fun evn :: "nat \ bool" where "evn 0 = True" | "evn (Suc 0) = False" | "evn (Suc(Suc n)) = evn n" text\We recast the proof of \<^prop>\ev n \ evn n\ in Isar. The left column shows the actual proof text, the right column shows the implicit effect of the two \isacom{case} commands:\text_raw\ \begin{tabular}{@ {}l@ {\qquad}l@ {}} \begin{minipage}[t]{.5\textwidth} \isa{% \ lemma "ev n \ evn n" proof(induction rule: ev.induct) case ev0 show ?case by simp next case evSS thus ?case by simp qed text_raw \} \end{minipage} & \begin{minipage}[t]{.5\textwidth} ~\\ ~\\ \isacom{let} \?case = "evn 0"\\\ ~\\ ~\\ \isacom{fix} \n\\\ \isacom{assume} \evSS:\ \begin{tabular}[t]{l} \"ev n"\\\\"evn n"\\end{tabular}\\ \isacom{let} \?case = "evn(Suc(Suc n))"\\\ \end{minipage} \end{tabular} \medskip \ text\ The proof resembles structural induction, but the induction rule is given explicitly and the names of the cases are the names of the rules in the inductive definition. Let us examine the two assumptions named @{thm[source]evSS}: \<^prop>\ev n\ is the premise of rule @{thm[source]evSS}, which we may assume because we are in the case where that rule was used; \<^prop>\evn n\ is the induction hypothesis. \begin{warn} Because each \isacom{case} command introduces a list of assumptions named like the case name, which is the name of a rule of the inductive definition, those rules now need to be accessed with a qualified name, here @{thm[source] ev.ev0} and @{thm[source] ev.evSS}. \end{warn} In the case @{thm[source]evSS} of the proof above we have pretended that the system fixes a variable \n\. But unless the user provides the name \n\, the system will just invent its own name that cannot be referred to. In the above proof, we do not need to refer to it, hence we do not give it a specific name. In case one needs to refer to it one writes \begin{quote} \isacom{case} \(evSS m)\ \end{quote} like \isacom{case}~\(Suc n)\ in earlier structural inductions. The name \m\ is an arbitrary choice. As a result, case @{thm[source] evSS} is derived from a renamed version of rule @{thm[source] evSS}: \ev m \ ev(Suc(Suc m))\. Here is an example with a (contrived) intermediate step that refers to \m\: \ lemma "ev n \ evn n" proof(induction rule: ev.induct) case ev0 show ?case by simp next case (evSS m) have "evn(Suc(Suc m)) = evn m" by simp thus ?case using `evn m` by blast qed text\ \indent In general, let \I\ be a (for simplicity unary) inductively defined predicate and let the rules in the definition of \I\ be called \rule\<^sub>1\, \dots, \rule\<^sub>n\. A proof by rule induction follows this pattern:\index{inductionrule@\induction ... rule:\} \ (*<*) inductive I where rule\<^sub>1: "I()" | rule\<^sub>2: "I()" | rule\<^sub>n: "I()" lemma "I x \ P x" proof-(*>*) show "I x \ P x" proof(induction rule: I.induct) case rule\<^sub>1 text_raw\\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\ show ?case (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ next text_raw\\\[-.4ex]$\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\ (*<*) case rule\<^sub>2 show ?case sorry (*>*) next case rule\<^sub>n text_raw\\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\ show ?case (*<*)sorry(*>*)text_raw\\ \isasymproof\\\ qed(*<*)qed(*>*) text\ One can provide explicit variable names by writing \isacom{case}~\(rule\<^sub>i x\<^sub>1 \ x\<^sub>k)\, thus renaming the first \k\ free variables in rule \i\ to \x\<^sub>1 \ x\<^sub>k\, going through rule \i\ from left to right. \subsection{Assumption Naming} \label{sec:assm-naming} In any induction, \isacom{case}~\name\ sets up a list of assumptions also called \name\, which is subdivided into three parts: \begin{description} \item[\name.IH\]\index{IH@\.IH\} contains the induction hypotheses. \item[\name.hyps\]\index{hyps@\.hyps\} contains all the other hypotheses of this case in the induction rule. For rule inductions these are the hypotheses of rule \name\, for structural inductions these are empty. \item[\name.prems\]\index{prems@\.prems\} contains the (suitably instantiated) premises of the statement being proved, i.e., the \A\<^sub>i\ when proving \\ A\<^sub>1; \; A\<^sub>n \ \ A\. \end{description} \begin{warn} Proof method \induct\ differs from \induction\ only in this naming policy: \induct\ does not distinguish \IH\ from \hyps\ but subsumes \IH\ under \hyps\. \end{warn} More complicated inductive proofs than the ones we have seen so far often need to refer to specific assumptions --- just \name\ or even \name.prems\ and \name.IH\ can be too unspecific. This is where the indexing of fact lists comes in handy, e.g., \name.IH(2)\ or \name.prems(1-2)\. \subsection{Rule Inversion} \label{sec:rule-inversion} \index{rule inversion|(} Rule inversion is case analysis of which rule could have been used to derive some fact. The name \conceptnoidx{rule inversion} emphasizes that we are reasoning backwards: by which rules could some given fact have been proved? For the inductive definition of \<^const>\ev\, rule inversion can be summarized like this: @{prop[display]"ev n \ n = 0 \ (\k. n = Suc(Suc k) \ ev k)"} The realisation in Isabelle is a case analysis. A simple example is the proof that \<^prop>\ev n \ ev (n - 2)\. We already went through the details informally in \autoref{sec:Logic:even}. This is the Isar proof: \ (*<*) notepad begin fix n (*>*) assume "ev n" from this have "ev(n - 2)" proof cases case ev0 thus "ev(n - 2)" by (simp add: ev.ev0) next case (evSS k) thus "ev(n - 2)" by (simp add: ev.evSS) qed (*<*) end (*>*) text\The key point here is that a case analysis over some inductively defined predicate is triggered by piping the given fact (here: \isacom{from}~\this\) into a proof by \cases\. Let us examine the assumptions available in each case. In case \ev0\ we have \n = 0\ and in case \evSS\ we have \<^prop>\n = Suc(Suc k)\ and \<^prop>\ev k\. In each case the assumptions are available under the name of the case; there is no fine-grained naming schema like there is for induction. Sometimes some rules could not have been used to derive the given fact because constructors clash. As an extreme example consider rule inversion applied to \<^prop>\ev(Suc 0)\: neither rule \ev0\ nor rule \evSS\ can yield \<^prop>\ev(Suc 0)\ because \Suc 0\ unifies neither with \0\ nor with \<^term>\Suc(Suc n)\. Impossible cases do not have to be proved. Hence we can prove anything from \<^prop>\ev(Suc 0)\: \ (*<*) notepad begin fix P (*>*) assume "ev(Suc 0)" then have P by cases (*<*) end (*>*) text\That is, \<^prop>\ev(Suc 0)\ is simply not provable:\ lemma "\ ev(Suc 0)" proof assume "ev(Suc 0)" then show False by cases qed text\Normally not all cases will be impossible. As a simple exercise, prove that \mbox{\<^prop>\\ ev(Suc(Suc(Suc 0)))\.} \subsection{Advanced Rule Induction} \label{sec:advanced-rule-induction} So far, rule induction was always applied to goals of the form \I x y z \ \\ where \I\ is some inductively defined predicate and \x\, \y\, \z\ are variables. In some rare situations one needs to deal with an assumption where not all arguments \r\, \s\, \t\ are variables: \begin{isabelle} \isacom{lemma} \"I r s t \ \"\ \end{isabelle} Applying the standard form of rule induction in such a situation will lead to strange and typically unprovable goals. We can easily reduce this situation to the standard one by introducing new variables \x\, \y\, \z\ and reformulating the goal like this: \begin{isabelle} \isacom{lemma} \"I x y z \ x = r \ y = s \ z = t \ \"\ \end{isabelle} Standard rule induction will work fine now, provided the free variables in \r\, \s\, \t\ are generalized via \arbitrary\. However, induction can do the above transformation for us, behind the curtains, so we never need to see the expanded version of the lemma. This is what we need to write: \begin{isabelle} \isacom{lemma} \"I r s t \ \"\\isanewline \isacom{proof}\(induction "r" "s" "t" arbitrary: \ rule: I.induct)\\index{inductionrule@\induction ... rule:\}\index{arbitrary@\arbitrary:\} \end{isabelle} Like for rule inversion, cases that are impossible because of constructor clashes will not show up at all. Here is a concrete example:\ lemma "ev (Suc m) \ \ ev m" proof(induction "Suc m" arbitrary: m rule: ev.induct) fix n assume IH: "\m. n = Suc m \ \ ev m" show "\ ev (Suc n)" proof \ \contradiction\ assume "ev(Suc n)" thus False proof cases \ \rule inversion\ fix k assume "n = Suc k" "ev k" thus False using IH by auto qed qed qed text\ Remarks: \begin{itemize} \item Instead of the \isacom{case} and \?case\ magic we have spelled all formulas out. This is merely for greater clarity. \item We only need to deal with one case because the @{thm[source] ev0} case is impossible. \item The form of the \IH\ shows us that internally the lemma was expanded as explained above: \noquotes{@{prop[source]"ev x \ x = Suc m \ \ ev m"}}. \item The goal \<^prop>\\ ev (Suc n)\ may surprise. The expanded version of the lemma would suggest that we have a \isacom{fix} \m\ \isacom{assume} \<^prop>\Suc(Suc n) = Suc m\ and need to show \<^prop>\\ ev m\. What happened is that Isabelle immediately simplified \<^prop>\Suc(Suc n) = Suc m\ to \<^prop>\Suc n = m\ and could then eliminate \m\. Beware of such nice surprises with this advanced form of induction. \end{itemize} \begin{warn} This advanced form of induction does not support the \IH\ naming schema explained in \autoref{sec:assm-naming}: the induction hypotheses are instead found under the name \hyps\, as they are for the simpler \induct\ method. \end{warn} \index{induction|)} \index{cases@\cases\|)} \index{case@\isacom{case}|)} \index{case?@\?case\|)} \index{rule induction|)} \index{rule inversion|)} \subsection*{Exercises} \exercise Give a structured proof by rule inversion: \ lemma assumes a: "ev(Suc(Suc n))" shows "ev n" (*<*)oops(*>*) text\ \endexercise \begin{exercise} Give a structured proof of \<^prop>\\ ev(Suc(Suc(Suc 0)))\ by rule inversions. If there are no cases to be proved you can close a proof immediately with \isacom{qed}. \end{exercise} \begin{exercise} Recall predicate \star\ from \autoref{sec:star} and \iter\ from Exercise~\ref{exe:iter}. Prove \<^prop>\iter r n x y \ star r x y\ in a structured style; do not just sledgehammer each case of the required induction. \end{exercise} \begin{exercise} Define a recursive function \elems ::\ \<^typ>\'a list \ 'a set\ and prove \<^prop>\x \ elems xs \ \ys zs. xs = ys @ x # zs \ x \ elems ys\. \end{exercise} \begin{exercise} Extend Exercise~\ref{exe:cfg} with a function that checks if some \mbox{\alpha list\} is a balanced string of parentheses. More precisely, define a \mbox{recursive} function \balanced :: nat \ alpha list \ bool\ such that \<^term>\balanced n w\ is true iff (informally) \S (a\<^sup>n @ w)\. Formally, prove that \<^prop>\balanced n w \ S (replicate n a @ w)\ where \<^const>\replicate\ \::\ \<^typ>\nat \ 'a \ 'a list\ is predefined and \<^term>\replicate n x\ yields the list \[x, \, x]\ of length \n\. \end{exercise} \ (*<*) end (*>*) diff --git a/src/Doc/Prog_Prove/Logic.thy b/src/Doc/Prog_Prove/Logic.thy --- a/src/Doc/Prog_Prove/Logic.thy +++ b/src/Doc/Prog_Prove/Logic.thy @@ -1,880 +1,880 @@ (*<*) theory Logic imports LaTeXsugar begin (*>*) text\ \vspace{-5ex} \section{Formulas} The core syntax of formulas (\textit{form} below) provides the standard logical constructs, in decreasing order of precedence: \[ \begin{array}{rcl} \mathit{form} & ::= & \(form)\ ~\mid~ \<^const>\True\ ~\mid~ \<^const>\False\ ~\mid~ \<^prop>\term = term\\\ &\mid& \<^prop>\\ form\\index{$HOL4@\isasymnot} ~\mid~ \<^prop>\form \ form\\index{$HOL0@\isasymand} ~\mid~ \<^prop>\form \ form\\index{$HOL1@\isasymor} ~\mid~ \<^prop>\form \ form\\index{$HOL2@\isasymlongrightarrow}\\ &\mid& \<^prop>\\x. form\\index{$HOL6@\isasymforall} ~\mid~ \<^prop>\\x. form\\index{$HOL7@\isasymexists} \end{array} \] Terms are the ones we have seen all along, built from constants, variables, function application and \\\-abstraction, including all the syntactic sugar like infix symbols, \if\, \case\, etc. \begin{warn} Remember that formulas are simply terms of type \bool\. Hence \=\ also works for formulas. Beware that \=\ has a higher precedence than the other logical operators. Hence \<^prop>\s = t \ A\ means \(s = t) \ A\, and \<^prop>\A\B = B\A\ means \A \ (B = B) \ A\. Logical equivalence can also be written with \\\ instead of \=\, where \\\ has the same low precedence as \\\. Hence \A \ B \ B \ A\ really means \(A \ B) \ (B \ A)\. \end{warn} \begin{warn} Quantifiers need to be enclosed in parentheses if they are nested within other constructs (just like \if\, \case\ and \let\). \end{warn} The most frequent logical symbols and their ASCII representations are shown in Fig.~\ref{fig:log-symbols}. \begin{figure} \begin{center} \begin{tabular}{l@ {\qquad}l@ {\qquad}l} \\\ & \xsymbol{forall} & \texttt{ALL}\\ \\\ & \xsymbol{exists} & \texttt{EX}\\ \\\ & \xsymbol{lambda} & \texttt{\%}\\ \\\ & \texttt{-{\kern0pt}->}\\ \\\ & \texttt{<->}\\ \\\ & \texttt{/\char`\\} & \texttt{\&}\\ \\\ & \texttt{\char`\\/} & \texttt{|}\\ \\\ & \xsymbol{not} & \texttt{\char`~}\\ \\\ & \xsymbol{noteq} & \texttt{\char`~=} \end{tabular} \end{center} \caption{Logical symbols and their ASCII forms} \label{fig:log-symbols} \end{figure} The first column shows the symbols, the other columns ASCII representations. The \texttt{\char`\\}\texttt{<...>} form is always converted into the symbolic form by the Isabelle interfaces, the treatment of the other ASCII forms depends on the interface. The ASCII forms \texttt{/\char`\\} and \texttt{\char`\\/} are special in that they are merely keyboard shortcuts for the interface and not logical symbols by themselves. \begin{warn} The implication \\\ is part of the Isabelle framework. It structures theorems and proof states, separating assumptions from conclusions. The implication \\\ is part of the logic HOL and can occur inside the formulas that make up the assumptions and conclusion. Theorems should be of the form \\ A\<^sub>1; \; A\<^sub>n \ \ A\, not \A\<^sub>1 \ \ \ A\<^sub>n \ A\. Both are logically equivalent but the first one works better when using the theorem in further proofs. The ASCII representation of \\\ and \\\ is \texttt{[|} and \texttt{|]}. \end{warn} \section{Sets} \label{sec:Sets} Sets of elements of type \<^typ>\'a\ have type \<^typ>\'a set\\index{set@\set\}. They can be finite or infinite. Sets come with the usual notation: \begin{itemize} \item \indexed{\<^term>\{}\}{$IMP042},\quad \{e\<^sub>1,\,e\<^sub>n}\ \item \<^prop>\e \ A\\index{$HOLSet0@\isasymin},\quad \<^prop>\A \ B\\index{$HOLSet2@\isasymsubseteq} \item \<^term>\A \ B\\index{$HOLSet4@\isasymunion},\quad \<^term>\A \ B\\index{$HOLSet5@\isasyminter},\quad \<^term>\A - B\,\quad \<^term>\-A\ \end{itemize} (where \<^term>\A-B\ and \-A\ are set difference and complement) and much more. \<^const>\UNIV\ is the set of all elements of some type. Set comprehension\index{set comprehension} is written \<^term>\{x. P}\\index{$IMP042@\<^term>\{x. P}\} rather than \{x | P}\. \begin{warn} In \<^term>\{x. P}\ the \x\ must be a variable. Set comprehension involving a proper term \t\ must be written \noquotes{@{term[source] "{t | x y. P}"}}\index{$IMP042@\{t |x. P}\}, where \x y\ are those free variables in \t\ that occur in \P\. This is just a shorthand for \<^term>\{v. \x y. v = t \ P}\, where \v\ is a new variable. For example, \<^term>\{x+y|x. x \ A}\ is short for \noquotes{@{term[source]"{v. \x. v = x+y \ x \ A}"}}. \end{warn} Here are the ASCII representations of the mathematical symbols: \begin{center} \begin{tabular}{l@ {\quad}l@ {\quad}l} \\\ & \texttt{\char`\\\char`\} & \texttt{:}\\ \\\ & \texttt{\char`\\\char`\} & \texttt{<=}\\ \\\ & \texttt{\char`\\\char`\} & \texttt{Un}\\ \\\ & \texttt{\char`\\\char`\} & \texttt{Int} \end{tabular} \end{center} Sets also allow bounded quantifications \<^prop>\\x \ A. P\ and \<^prop>\\x \ A. P\. For the more ambitious, there are also \\\\index{$HOLSet6@\isasymUnion} and \\\\index{$HOLSet7@\isasymInter}: \begin{center} @{thm Union_eq} \qquad @{thm Inter_eq} \end{center} The ASCII forms of \\\ are \texttt{\char`\\\char`\} and \texttt{Union}, those of \\\ are \texttt{\char`\\\char`\} and \texttt{Inter}. There are also indexed unions and intersections: \begin{center} @{thm[eta_contract=false] UNION_eq} \\ @{thm[eta_contract=false] INTER_eq} \end{center} The ASCII forms are \ \texttt{UN x:A.~B} \ and \ \texttt{INT x:A. B} \ where \texttt{x} may occur in \texttt{B}. If \texttt{A} is \texttt{UNIV} you can write \ \texttt{UN x.~B} \ and \ \texttt{INT x. B}. Some other frequently useful functions on sets are the following: \begin{center} \begin{tabular}{l@ {\quad}l} @{const_typ set}\index{set@\<^const>\set\} & converts a list to the set of its elements\\ @{const_typ finite}\index{finite@\<^const>\finite\} & is true iff its argument is finite\\ \noquotes{@{term[source] "card :: 'a set \ nat"}}\index{card@\<^const>\card\} & is the cardinality of a finite set\\ & and is \0\ for all infinite sets\\ @{thm image_def}\index{$IMP042@\<^term>\f ` A\} & is the image of a function over a set \end{tabular} \end{center} -See @{cite "Nipkow-Main"} for the wealth of further predefined functions in theory +See \<^cite>\"Nipkow-Main"\ for the wealth of further predefined functions in theory \<^theory>\Main\. \subsection*{Exercises} \exercise Start from the data type of binary trees defined earlier: \ datatype 'a tree = Tip | Node "'a tree" 'a "'a tree" text\ Define a function \set ::\ \<^typ>\'a tree \ 'a set\ that returns the elements in a tree and a function \ord ::\ \<^typ>\int tree \ bool\ that tests if an \<^typ>\int tree\ is ordered. Define a function \ins\ that inserts an element into an ordered \<^typ>\int tree\ while maintaining the order of the tree. If the element is already in the tree, the same tree should be returned. Prove correctness of \ins\: \<^prop>\set(ins x t) = {x} \ set t\ and \<^prop>\ord t \ ord(ins i t)\. \endexercise \section{Proof Automation} So far we have only seen \simp\ and \indexed{\auto\}{auto}: Both perform rewriting, both can also prove linear arithmetic facts (no multiplication), and \auto\ is also able to prove simple logical or set-theoretic goals: \ lemma "\x. \y. x = y" by auto lemma "A \ B \ C \ A \ B \ C" by auto text\where \begin{quote} \isacom{by} \textit{proof-method} \end{quote} is short for \begin{quote} \isacom{apply} \textit{proof-method}\\ \isacom{done} \end{quote} The key characteristics of both \simp\ and \auto\ are \begin{itemize} \item They show you where they got stuck, giving you an idea how to continue. \item They perform the obvious steps but are highly incomplete. \end{itemize} A proof method is \conceptnoidx{complete} if it can prove all true formulas. There is no complete proof method for HOL, not even in theory. Hence all our proof methods only differ in how incomplete they are. A proof method that is still incomplete but tries harder than \auto\ is \indexed{\fastforce\}{fastforce}. It either succeeds or fails, it acts on the first subgoal only, and it can be modified like \auto\, e.g., with \simp add\. Here is a typical example of what \fastforce\ can do: \ lemma "\ \xs \ A. \ys. xs = ys @ ys; us \ A \ \ \n. length us = n+n" by fastforce text\This lemma is out of reach for \auto\ because of the quantifiers. Even \fastforce\ fails when the quantifier structure becomes more complicated. In a few cases, its slow version \force\ succeeds where \fastforce\ fails. The method of choice for complex logical goals is \indexed{\blast\}{blast}. In the following example, \T\ and \A\ are two binary predicates. It is shown that if \T\ is total, \A\ is antisymmetric and \T\ is a subset of \A\, then \A\ is a subset of \T\: \ lemma "\ \x y. T x y \ T y x; \x y. A x y \ A y x \ x = y; \x y. T x y \ A x y \ \ \x y. A x y \ T x y" by blast text\ We leave it to the reader to figure out why this lemma is true. Method \blast\ \begin{itemize} \item is (in principle) a complete proof procedure for first-order formulas, a fragment of HOL. In practice there is a search bound. \item does no rewriting and knows very little about equality. \item covers logic, sets and relations. \item either succeeds or fails. \end{itemize} Because of its strength in logic and sets and its weakness in equality reasoning, it complements the earlier proof methods. \subsection{\concept{Sledgehammer}} Command \isacom{sledgehammer} calls a number of external automatic theorem provers (ATPs) that run for up to 30 seconds searching for a proof. Some of these ATPs are part of the Isabelle installation, others are queried over the internet. If successful, a proof command is generated and can be inserted into your proof. The biggest win of \isacom{sledgehammer} is that it will take into account the whole lemma library and you do not need to feed in any lemma explicitly. For example,\ lemma "\ xs @ ys = ys @ xs; length xs = length ys \ \ xs = ys" txt\cannot be solved by any of the standard proof methods, but \isacom{sledgehammer} finds the following proof:\ by (metis append_eq_conv_conj) text\We do not explain how the proof was found but what this command means. For a start, Isabelle does not trust external tools (and in particular not the translations from Isabelle's logic to those tools!) and insists on a proof that it can check. This is what \indexed{\metis\}{metis} does. It is given a list of lemmas and tries to find a proof using just those lemmas (and pure logic). In contrast to using \simp\ and friends who know a lot of lemmas already, using \metis\ manually is tedious because one has to find all the relevant lemmas first. But that is precisely what \isacom{sledgehammer} does for us. In this case lemma @{thm[source]append_eq_conv_conj} alone suffices: @{thm[display] append_eq_conv_conj} We leave it to the reader to figure out why this lemma suffices to prove the above lemma, even without any knowledge of what the functions \<^const>\take\ and \<^const>\drop\ do. Keep in mind that the variables in the two lemmas are independent of each other, despite the same names, and that you can substitute arbitrary values for the free variables in a lemma. Just as for the other proof methods we have seen, there is no guarantee that \isacom{sledgehammer} will find a proof if it exists. Nor is \isacom{sledgehammer} superior to the other proof methods. They are incomparable. Therefore it is recommended to apply \simp\ or \auto\ before invoking \isacom{sledgehammer} on what is left. \subsection{Arithmetic} By arithmetic formulas we mean formulas involving variables, numbers, \+\, \-\, \=\, \<\, \\\ and the usual logical connectives \\\, \\\, \\\, \\\, \\\. Strictly speaking, this is known as \concept{linear arithmetic} because it does not involve multiplication, although multiplication with numbers, e.g., \2*n\, is allowed. Such formulas can be proved by \indexed{\arith\}{arith}: \ lemma "\ (a::nat) \ x + b; 2*x < c \ \ 2*a + 1 \ 2*b + c" by arith text\In fact, \auto\ and \simp\ can prove many linear arithmetic formulas already, like the one above, by calling a weak but fast version of \arith\. Hence it is usually not necessary to invoke \arith\ explicitly. The above example involves natural numbers, but integers (type \<^typ>\int\) and real numbers (type \real\) are supported as well. As are a number of further operators like \<^const>\min\ and \<^const>\max\. On \<^typ>\nat\ and \<^typ>\int\, \arith\ can even prove theorems with quantifiers in them, but we will not enlarge on that here. \subsection{Trying Them All} If you want to try all of the above automatic proof methods you simply type \begin{isabelle} \isacom{try} \end{isabelle} There is also a lightweight variant \isacom{try0} that does not call sledgehammer. If desired, specific simplification and introduction rules can be added: \begin{isabelle} \isacom{try0} \simp: \ intro: \\ \end{isabelle} \section{Single Step Proofs} Although automation is nice, it often fails, at least initially, and you need to find out why. When \fastforce\ or \blast\ simply fail, you have no clue why. At this point, the stepwise application of proof rules may be necessary. For example, if \blast\ fails on \<^prop>\A \ B\, you want to attack the two conjuncts \A\ and \B\ separately. This can be achieved by applying \emph{conjunction introduction} \[ @{thm[mode=Rule,show_question_marks]conjI}\ \conjI\ \] to the proof state. We will now examine the details of this process. \subsection{Instantiating Unknowns} We had briefly mentioned earlier that after proving some theorem, Isabelle replaces all free variables \x\ by so called \conceptidx{unknowns}{unknown} \?x\. We can see this clearly in rule @{thm[source] conjI}. These unknowns can later be instantiated explicitly or implicitly: \begin{itemize} \item By hand, using \indexed{\of\}{of}. The expression \conjI[of "a=b" "False"]\ instantiates the unknowns in @{thm[source] conjI} from left to right with the two formulas \a=b\ and \False\, yielding the rule @{thm[display,mode=Rule,margin=100]conjI[of "a=b" False]} In general, \th[of string\<^sub>1 \ string\<^sub>n]\ instantiates the unknowns in the theorem \th\ from left to right with the terms \string\<^sub>1\ to \string\<^sub>n\. \item By unification. \conceptidx{Unification}{unification} is the process of making two terms syntactically equal by suitable instantiations of unknowns. For example, unifying \?P \ ?Q\ with \mbox{\<^prop>\a=b \ False\} instantiates \?P\ with \<^prop>\a=b\ and \?Q\ with \<^prop>\False\. \end{itemize} We need not instantiate all unknowns. If we want to skip a particular one we can write \_\ instead, for example \conjI[of _ "False"]\. Unknowns can also be instantiated by name using \indexed{\where\}{where}, for example \conjI[where ?P = "a=b"\ \isacom{and} \?Q = "False"]\. \subsection{Rule Application} \conceptidx{Rule application}{rule application} means applying a rule backwards to a proof state. For example, applying rule @{thm[source]conjI} to a proof state \begin{quote} \1. \ \ A \ B\ \end{quote} results in two subgoals, one for each premise of @{thm[source]conjI}: \begin{quote} \1. \ \ A\\\ \2. \ \ B\ \end{quote} In general, the application of a rule \\ A\<^sub>1; \; A\<^sub>n \ \ A\ to a subgoal \mbox{\\ \ C\} proceeds in two steps: \begin{enumerate} \item Unify \A\ and \C\, thus instantiating the unknowns in the rule. \item Replace the subgoal \C\ with \n\ new subgoals \A\<^sub>1\ to \A\<^sub>n\. \end{enumerate} This is the command to apply rule \xyz\: \begin{quote} \isacom{apply}\(rule xyz)\\index{rule@\rule\} \end{quote} This is also called \concept{backchaining} with rule \xyz\. \subsection{Introduction Rules} Conjunction introduction (@{thm[source] conjI}) is one example of a whole class of rules known as \conceptidx{introduction rules}{introduction rule}. They explain under which premises some logical construct can be introduced. Here are some further useful introduction rules: \[ \inferrule*[right=\mbox{\impI\}]{\mbox{\?P \ ?Q\}}{\mbox{\?P \ ?Q\}} \qquad \inferrule*[right=\mbox{\allI\}]{\mbox{\\x. ?P x\}}{\mbox{\\x. ?P x\}} \] \[ \inferrule*[right=\mbox{\iffI\}]{\mbox{\?P \ ?Q\} \\ \mbox{\?Q \ ?P\}} {\mbox{\?P = ?Q\}} \] These rules are part of the logical system of \concept{natural deduction} -(e.g., @{cite HuthRyan}). Although we intentionally de-emphasize the basic rules +(e.g., \<^cite>\HuthRyan\). Although we intentionally de-emphasize the basic rules of logic in favour of automatic proof methods that allow you to take bigger steps, these rules are helpful in locating where and why automation fails. When applied backwards, these rules decompose the goal: \begin{itemize} \item @{thm[source] conjI} and @{thm[source]iffI} split the goal into two subgoals, \item @{thm[source] impI} moves the left-hand side of a HOL implication into the list of assumptions, \item and @{thm[source] allI} removes a \\\ by turning the quantified variable into a fixed local variable of the subgoal. \end{itemize} Isabelle knows about these and a number of other introduction rules. The command \begin{quote} \isacom{apply} \rule\\index{rule@\rule\} \end{quote} automatically selects the appropriate rule for the current subgoal. You can also turn your own theorems into introduction rules by giving them the \indexed{\intro\}{intro} attribute, analogous to the \simp\ attribute. In that case \blast\, \fastforce\ and (to a limited extent) \auto\ will automatically backchain with those theorems. The \intro\ attribute should be used with care because it increases the search space and can lead to nontermination. Sometimes it is better to use it only in specific calls of \blast\ and friends. For example, @{thm[source] le_trans}, transitivity of \\\ on type \<^typ>\nat\, is not an introduction rule by default because of the disastrous effect on the search space, but can be useful in specific situations: \ lemma "\ (a::nat) \ b; b \ c; c \ d; d \ e \ \ a \ e" by(blast intro: le_trans) text\ Of course this is just an example and could be proved by \arith\, too. \subsection{Forward Proof} \label{sec:forward-proof} Forward proof means deriving new theorems from old theorems. We have already seen a very simple form of forward proof: the \of\ operator for instantiating unknowns in a theorem. The big brother of \of\ is \indexed{\OF\}{OF} for applying one theorem to others. Given a theorem \<^prop>\A \ B\ called \r\ and a theorem \A'\ called \r'\, the theorem \r[OF r']\ is the result of applying \r\ to \r'\, where \r\ should be viewed as a function taking a theorem \A\ and returning \B\. More precisely, \A\ and \A'\ are unified, thus instantiating the unknowns in \B\, and the result is the instantiated \B\. Of course, unification may also fail. \begin{warn} Application of rules to other rules operates in the forward direction: from the premises to the conclusion of the rule; application of rules to proof states operates in the backward direction, from the conclusion to the premises. \end{warn} In general \r\ can be of the form \\ A\<^sub>1; \; A\<^sub>n \ \ A\ and there can be multiple argument theorems \r\<^sub>1\ to \r\<^sub>m\ (with \m \ n\), in which case \r[OF r\<^sub>1 \ r\<^sub>m]\ is obtained by unifying and thus proving \A\<^sub>i\ with \r\<^sub>i\, \i = 1\m\. Here is an example, where @{thm[source]refl} is the theorem @{thm[show_question_marks] refl}: \ thm conjI[OF refl[of "a"] refl[of "b"]] text\yields the theorem @{thm conjI[OF refl[of "a"] refl[of "b"]]}. The command \isacom{thm} merely displays the result. Forward reasoning also makes sense in connection with proof states. Therefore \blast\, \fastforce\ and \auto\ support a modifier \dest\ which instructs the proof method to use certain rules in a forward fashion. If \r\ is of the form \mbox{\A \ B\}, the modifier \mbox{\dest: r\}\index{dest@\dest:\} allows proof search to reason forward with \r\, i.e., to replace an assumption \A'\, where \A'\ unifies with \A\, with the correspondingly instantiated \B\. For example, @{thm[source,show_question_marks] Suc_leD} is the theorem \mbox{@{thm Suc_leD}}, which works well for forward reasoning: \ lemma "Suc(Suc(Suc a)) \ b \ a \ b" by(blast dest: Suc_leD) text\In this particular example we could have backchained with @{thm[source] Suc_leD}, too, but because the premise is more complicated than the conclusion this can easily lead to nontermination. %\subsection{Finding Theorems} % %Command \isacom{find{\isacharunderscorekeyword}theorems} searches for specific theorems in the current %theory. Search criteria include pattern matching on terms and on names. -%For details see the Isabelle/Isar Reference Manual~@{cite IsarRef}. +%For details see the Isabelle/Isar Reference Manual~\<^cite>\IsarRef\. %\bigskip \begin{warn} To ease readability we will drop the question marks in front of unknowns from now on. \end{warn} \section{Inductive Definitions} \label{sec:inductive-defs}\index{inductive definition|(} Inductive definitions are the third important definition facility, after datatypes and recursive function. \ifsem In fact, they are the key construct in the definition of operational semantics in the second part of the book. \fi \subsection{An Example: Even Numbers} \label{sec:Logic:even} Here is a simple example of an inductively defined predicate: \begin{itemize} \item 0 is even \item If $n$ is even, so is $n+2$. \end{itemize} The operative word ``inductive'' means that these are the only even numbers. In Isabelle we give the two rules the names \ev0\ and \evSS\ and write \ inductive ev :: "nat \ bool" where ev0: "ev 0" | evSS: (*<*)"ev n \ ev (Suc(Suc n))"(*>*) text_raw\@{prop[source]"ev n \ ev (n + 2)"}\ text\To get used to inductive definitions, we will first prove a few properties of \<^const>\ev\ informally before we descend to the Isabelle level. How do we prove that some number is even, e.g., \<^prop>\ev 4\? Simply by combining the defining rules for \<^const>\ev\: \begin{quote} \ev 0 \ ev (0 + 2) \ ev((0 + 2) + 2) = ev 4\ \end{quote} \subsubsection{Rule Induction}\index{rule induction|(} Showing that all even numbers have some property is more complicated. For example, let us prove that the inductive definition of even numbers agrees with the following recursive one:\ fun evn :: "nat \ bool" where "evn 0 = True" | "evn (Suc 0) = False" | "evn (Suc(Suc n)) = evn n" text\We prove \<^prop>\ev m \ evn m\. That is, we assume \<^prop>\ev m\ and by induction on the form of its derivation prove \<^prop>\evn m\. There are two cases corresponding to the two rules for \<^const>\ev\: \begin{description} \item[Case @{thm[source]ev0}:] \<^prop>\ev m\ was derived by rule \<^prop>\ev 0\: \\ \\\ \<^prop>\m=(0::nat)\ \\\ \evn m = evn 0 = True\ \item[Case @{thm[source]evSS}:] \<^prop>\ev m\ was derived by rule \<^prop>\ev n \ ev(n+2)\: \\ \\\ \<^prop>\m=n+(2::nat)\ and by induction hypothesis \<^prop>\evn n\\\ \\\ \evn m = evn(n + 2) = evn n = True\ \end{description} What we have just seen is a special case of \concept{rule induction}. Rule induction applies to propositions of this form \begin{quote} \<^prop>\ev n \ P n\ \end{quote} That is, we want to prove a property \<^prop>\P n\ for all even \n\. But if we assume \<^prop>\ev n\, then there must be some derivation of this assumption using the two defining rules for \<^const>\ev\. That is, we must prove \begin{description} \item[Case @{thm[source]ev0}:] \<^prop>\P(0::nat)\ \item[Case @{thm[source]evSS}:] \<^prop>\\ ev n; P n \ \ P(n + 2::nat)\ \end{description} The corresponding rule is called @{thm[source] ev.induct} and looks like this: \[ \inferrule{ \mbox{@{thm (prem 1) ev.induct[of "n"]}}\\ \mbox{@{thm (prem 2) ev.induct}}\\ \mbox{\<^prop>\!!n. \ ev n; P n \ \ P(n+2)\}} {\mbox{@{thm (concl) ev.induct[of "n"]}}} \] The first premise \<^prop>\ev n\ enforces that this rule can only be applied in situations where we know that \n\ is even. Note that in the induction step we may not just assume \<^prop>\P n\ but also \mbox{\<^prop>\ev n\}, which is simply the premise of rule @{thm[source] evSS}. Here is an example where the local assumption \<^prop>\ev n\ comes in handy: we prove \<^prop>\ev m \ ev(m - 2)\ by induction on \<^prop>\ev m\. Case @{thm[source]ev0} requires us to prove \<^prop>\ev(0 - 2)\, which follows from \<^prop>\ev 0\ because \<^prop>\0 - 2 = (0::nat)\ on type \<^typ>\nat\. In case @{thm[source]evSS} we have \mbox{\<^prop>\m = n+(2::nat)\} and may assume \<^prop>\ev n\, which implies \<^prop>\ev (m - 2)\ because \m - 2 = (n + 2) - 2 = n\. We did not need the induction hypothesis at all for this proof (it is just a case analysis of which rule was used) but having \<^prop>\ev n\ at our disposal in case @{thm[source]evSS} was essential. This case analysis of rules is also called ``rule inversion'' and is discussed in more detail in \autoref{ch:Isar}. \subsubsection{In Isabelle} Let us now recast the above informal proofs in Isabelle. For a start, we use \<^const>\Suc\ terms instead of numerals in rule @{thm[source]evSS}: @{thm[display] evSS} This avoids the difficulty of unifying \n+2\ with some numeral, which is not automatic. The simplest way to prove \<^prop>\ev(Suc(Suc(Suc(Suc 0))))\ is in a forward direction: \evSS[OF evSS[OF ev0]]\ yields the theorem @{thm evSS[OF evSS[OF ev0]]}. Alternatively, you can also prove it as a lemma in backwards fashion. Although this is more verbose, it allows us to demonstrate how each rule application changes the proof state:\ lemma "ev(Suc(Suc(Suc(Suc 0))))" txt\ @{subgoals[display,indent=0,goals_limit=1]} \ apply(rule evSS) txt\ @{subgoals[display,indent=0,goals_limit=1]} \ apply(rule evSS) txt\ @{subgoals[display,indent=0,goals_limit=1]} \ apply(rule ev0) done text\\indent Rule induction is applied by giving the induction rule explicitly via the \rule:\ modifier:\index{inductionrule@\induction ... rule:\}\ lemma "ev m \ evn m" apply(induction rule: ev.induct) by(simp_all) text\Both cases are automatic. Note that if there are multiple assumptions of the form \<^prop>\ev t\, method \induction\ will induct on the leftmost one. As a bonus, we also prove the remaining direction of the equivalence of \<^const>\ev\ and \<^const>\evn\: \ lemma "evn n \ ev n" apply(induction n rule: evn.induct) txt\This is a proof by computation induction on \n\ (see \autoref{sec:recursive-funs}) that sets up three subgoals corresponding to the three equations for \<^const>\evn\: @{subgoals[display,indent=0]} The first and third subgoals follow with @{thm[source]ev0} and @{thm[source]evSS}, and the second subgoal is trivially true because \<^prop>\evn(Suc 0)\ is \<^const>\False\: \ by (simp_all add: ev0 evSS) text\The rules for \<^const>\ev\ make perfect simplification and introduction rules because their premises are always smaller than the conclusion. It makes sense to turn them into simplification and introduction rules permanently, to enhance proof automation. They are named @{thm[source] ev.intros} \index{intros@\.intros\} by Isabelle:\ declare ev.intros[simp,intro] text\The rules of an inductive definition are not simplification rules by default because, in contrast to recursive functions, there is no termination requirement for inductive definitions. \subsubsection{Inductive Versus Recursive} We have seen two definitions of the notion of evenness, an inductive and a recursive one. Which one is better? Much of the time, the recursive one is more convenient: it allows us to do rewriting in the middle of terms, and it expresses both the positive information (which numbers are even) and the negative information (which numbers are not even) directly. An inductive definition only expresses the positive information directly. The negative information, for example, that \1\ is not even, has to be proved from it (by induction or rule inversion). On the other hand, rule induction is tailor-made for proving \mbox{\<^prop>\ev n \ P n\} because it only asks you to prove the positive cases. In the proof of \<^prop>\evn n \ P n\ by computation induction via @{thm[source]evn.induct}, we are also presented with the trivial negative cases. If you want the convenience of both rewriting and rule induction, you can make two definitions and show their equivalence (as above) or make one definition and prove additional properties from it, for example rule induction from computation induction. But many concepts do not admit a recursive definition at all because there is no datatype for the recursion (for example, the transitive closure of a relation), or the recursion would not terminate (for example, an interpreter for a programming language). Even if there is a recursive definition, if we are only interested in the positive information, the inductive definition may be much simpler. \subsection{The Reflexive Transitive Closure} \label{sec:star} Evenness is really more conveniently expressed recursively than inductively. As a second and very typical example of an inductive definition we define the reflexive transitive closure. \ifsem It will also be an important building block for some of the semantics considered in the second part of the book. \fi The reflexive transitive closure, called \star\ below, is a function that maps a binary predicate to another binary predicate: if \r\ is of type \\ \ \ \ bool\ then \<^term>\star r\ is again of type \\ \ \ \ bool\, and \<^prop>\star r x y\ means that \x\ and \y\ are in the relation \<^term>\star r\. Think \<^term>\r\<^sup>*\ when you see \<^term>\star r\, because \star r\ is meant to be the reflexive transitive closure. That is, \<^prop>\star r x y\ is meant to be true if from \x\ we can reach \y\ in finitely many \r\ steps. This concept is naturally defined inductively:\ inductive star :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" for r where refl: "star r x x" | step: "r x y \ star r y z \ star r x z" text\The base case @{thm[source] refl} is reflexivity: \<^term>\x=y\. The step case @{thm[source]step} combines an \r\ step (from \x\ to \y\) and a \<^term>\star r\ step (from \y\ to \z\) into a \<^term>\star r\ step (from \x\ to \z\). The ``\isacom{for}~\r\'' in the header is merely a hint to Isabelle that \r\ is a fixed parameter of \<^const>\star\, in contrast to the further parameters of \<^const>\star\, which change. As a result, Isabelle generates a simpler induction rule. By definition \<^term>\star r\ is reflexive. It is also transitive, but we need rule induction to prove that:\ lemma star_trans: "star r x y \ star r y z \ star r x z" apply(induction rule: star.induct) (*<*) defer apply(rename_tac u x y) defer (*>*) txt\The induction is over \<^prop>\star r x y\ (the first matching assumption) and we try to prove \mbox{\<^prop>\star r y z \ star r x z\}, which we abbreviate by \<^prop>\P x y\. These are our two subgoals: @{subgoals[display,indent=0]} The first one is \<^prop>\P x x\, the result of case @{thm[source]refl}, and it is trivial:\index{assumption@\assumption\} \ apply(assumption) txt\Let us examine subgoal \2\, case @{thm[source] step}. Assumptions \<^prop>\r u x\ and \mbox{\<^prop>\star r x y\} are the premises of rule @{thm[source]step}. Assumption \<^prop>\star r y z \ star r x z\ is \mbox{\<^prop>\P x y\}, the IH coming from \<^prop>\star r x y\. We have to prove \<^prop>\P u y\, which we do by assuming \<^prop>\star r y z\ and proving \<^prop>\star r u z\. The proof itself is straightforward: from \mbox{\<^prop>\star r y z\} the IH leads to \<^prop>\star r x z\ which, together with \<^prop>\r u x\, leads to \mbox{\<^prop>\star r u z\} via rule @{thm[source]step}: \ apply(metis step) done text\\index{rule induction|)} \subsection{The General Case} Inductive definitions have approximately the following general form: \begin{quote} \isacom{inductive} \I :: "\ \ bool"\ \isacom{where} \end{quote} followed by a sequence of (possibly named) rules of the form \begin{quote} \\ I a\<^sub>1; \; I a\<^sub>n \ \ I a\ \end{quote} separated by \|\. As usual, \n\ can be 0. The corresponding rule induction principle \I.induct\ applies to propositions of the form \begin{quote} \<^prop>\I x \ P x\ \end{quote} where \P\ may itself be a chain of implications. \begin{warn} Rule induction is always on the leftmost premise of the goal. Hence \I x\ must be the first premise. \end{warn} Proving \<^prop>\I x \ P x\ by rule induction means proving for every rule of \I\ that \P\ is invariant: \begin{quote} \\ I a\<^sub>1; P a\<^sub>1; \; I a\<^sub>n; P a\<^sub>n \ \ P a\ \end{quote} The above format for inductive definitions is simplified in a number of respects. \I\ can have any number of arguments and each rule can have additional premises not involving \I\, so-called \conceptidx{side conditions}{side condition}. In rule inductions, these side conditions appear as additional assumptions. The \isacom{for} clause seen in the definition of the reflexive transitive closure simplifies the induction rule. \index{inductive definition|)} \subsection*{Exercises} \begin{exercise} Formalize the following definition of palindromes \begin{itemize} \item The empty list and a singleton list are palindromes. \item If \xs\ is a palindrome, so is \<^term>\a # xs @ [a]\. \end{itemize} as an inductive predicate \palindrome ::\ \<^typ>\'a list \ bool\ and prove that \<^prop>\rev xs = xs\ if \xs\ is a palindrome. \end{exercise} \exercise We could also have defined \<^const>\star\ as follows: \ inductive star' :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" for r where refl': "star' r x x" | step': "star' r x y \ r y z \ star' r x z" text\ The single \r\ step is performed after rather than before the \star'\ steps. Prove \<^prop>\star' r x y \ star r x y\ and \<^prop>\star r x y \ star' r x y\. You may need lemmas. Note that rule induction fails if the assumption about the inductive predicate is not the first assumption. \endexercise \begin{exercise}\label{exe:iter} Analogous to \<^const>\star\, give an inductive definition of the \n\-fold iteration of a relation \r\: \<^term>\iter r n x y\ should hold if there are \x\<^sub>0\, \dots, \x\<^sub>n\ such that \<^prop>\x = x\<^sub>0\, \<^prop>\x\<^sub>n = y\ and \r x\<^bsub>i\<^esub> x\<^bsub>i+1\<^esub>\ for all \<^prop>\i < n\. Correct and prove the following claim: \<^prop>\star r x y \ iter r n x y\. \end{exercise} \begin{exercise}\label{exe:cfg} A context-free grammar can be seen as an inductive definition where each nonterminal $A$ is an inductively defined predicate on lists of terminal symbols: $A(w)$ means that $w$ is in the language generated by $A$. For example, the production $S \to a S b$ can be viewed as the implication \<^prop>\S w \ S (a # w @ [b])\ where \a\ and \b\ are terminal symbols, i.e., elements of some alphabet. The alphabet can be defined like this: \isacom{datatype} \alpha = a | b | \\ Define the two grammars (where $\varepsilon$ is the empty word) \[ \begin{array}{r@ {\quad}c@ {\quad}l} S &\to& \varepsilon \quad\mid\quad aSb \quad\mid\quad SS \\ T &\to& \varepsilon \quad\mid\quad TaTb \end{array} \] as two inductive predicates. If you think of \a\ and \b\ as ``\(\'' and ``\)\'', the grammar defines strings of balanced parentheses. Prove \<^prop>\T w \ S w\ and \mbox{\<^prop>\S w \ T w\} separately and conclude \<^prop>\S w = T w\. \end{exercise} \ifsem \begin{exercise} In \autoref{sec:AExp} we defined a recursive evaluation function \aval :: aexp \ state \ val\. Define an inductive evaluation predicate \aval_rel :: aexp \ state \ val \ bool\ and prove that it agrees with the recursive function: \<^prop>\aval_rel a s v \ aval a s = v\, \<^prop>\aval a s = v \ aval_rel a s v\ and thus \noquotes{@{prop [source] "aval_rel a s v \ aval a s = v"}}. \end{exercise} \begin{exercise} Consider the stack machine from Chapter~3 and recall the concept of \concept{stack underflow} from Exercise~\ref{exe:stack-underflow}. Define an inductive predicate \ok :: nat \ instr list \ nat \ bool\ such that \ok n is n'\ means that with any initial stack of length \n\ the instructions \is\ can be executed without stack underflow and that the final stack has length \n'\. Prove that \ok\ correctly computes the final stack size @{prop[display] "\ok n is n'; length stk = n\ \ length (exec is s stk) = n'"} and that instruction sequences generated by \comp\ cannot cause stack underflow: \ \ok n (comp a) ?\ \ for some suitable value of \?\. \end{exercise} \fi \ (*<*) end (*>*) diff --git a/src/Doc/Prog_Prove/Types_and_funs.thy b/src/Doc/Prog_Prove/Types_and_funs.thy --- a/src/Doc/Prog_Prove/Types_and_funs.thy +++ b/src/Doc/Prog_Prove/Types_and_funs.thy @@ -1,605 +1,605 @@ (*<*) theory Types_and_funs imports Main begin (*>*) text\ \vspace{-5ex} \section{Type and Function Definitions} Type synonyms are abbreviations for existing types, for example \index{string@\string\}\ type_synonym string = "char list" text\ Type synonyms are expanded after parsing and are not present in internal representation and output. They are mere conveniences for the reader. \subsection{Datatypes} \label{sec:datatypes} The general form of a datatype definition looks like this: \begin{quote} \begin{tabular}{@ {}rclcll} \indexed{\isacom{datatype}}{datatype} \('a\<^sub>1,\,'a\<^sub>n)t\ & = & $C_1 \ \"\\tau_{1,1}\"\ \dots \"\\tau_{1,n_1}\"\$ \\ & $|$ & \dots \\ & $|$ & $C_k \ \"\\tau_{k,1}\"\ \dots \"\\tau_{k,n_k}\"\$ \end{tabular} \end{quote} It introduces the constructors \ $C_i :: \tau_{i,1}\Rightarrow \cdots \Rightarrow \tau_{i,n_i} \Rightarrow$~\('a\<^sub>1,\,'a\<^sub>n)t\ \ and expresses that any value of this type is built from these constructors in a unique manner. Uniqueness is implied by the following properties of the constructors: \begin{itemize} \item \emph{Distinctness:} $C_i\ \ldots \neq C_j\ \dots$ \quad if $i \neq j$ \item \emph{Injectivity:} \begin{tabular}[t]{l} $(C_i \ x_1 \dots x_{n_i} = C_i \ y_1 \dots y_{n_i}) =$\\ $(x_1 = y_1 \land \dots \land x_{n_i} = y_{n_i})$ \end{tabular} \end{itemize} The fact that any value of the datatype is built from the constructors implies the \concept{structural induction}\index{induction} rule: to show $P~x$ for all $x$ of type \('a\<^sub>1,\,'a\<^sub>n)t\, one needs to show $P(C_i\ x_1 \dots x_{n_i})$ (for each $i$) assuming $P(x_j)$ for all $j$ where $\tau_{i,j} =$~\('a\<^sub>1,\,'a\<^sub>n)t\. Distinctness and injectivity are applied automatically by \auto\ and other proof methods. Induction must be applied explicitly. Like in functional programming languages, datatype values can be taken apart with case expressions\index{case expression}\index{case expression@\case ... of\}, for example \begin{quote} \noquotes{@{term[source] "(case xs of [] \ 0 | x # _ \ Suc x)"}} \end{quote} Case expressions must be enclosed in parentheses. As an example of a datatype beyond \<^typ>\nat\ and \list\, consider binary trees: \ datatype 'a tree = Tip | Node "'a tree" 'a "'a tree" text\with a mirror function:\ fun mirror :: "'a tree \ 'a tree" where "mirror Tip = Tip" | "mirror (Node l a r) = Node (mirror r) a (mirror l)" text\The following lemma illustrates induction:\ lemma "mirror(mirror t) = t" apply(induction t) txt\yields @{subgoals[display]} The induction step contains two induction hypotheses, one for each subtree. An application of \auto\ finishes the proof. A very simple but also very useful datatype is the predefined @{datatype[display] option}\index{option@\option\}\index{None@\<^const>\None\}\index{Some@\<^const>\Some\} Its sole purpose is to add a new element \<^const>\None\ to an existing type \<^typ>\'a\. To make sure that \<^const>\None\ is distinct from all the elements of \<^typ>\'a\, you wrap them up in \<^const>\Some\ and call the new type \<^typ>\'a option\. A typical application is a lookup function on a list of key-value pairs, often called an association list: \ (*<*) apply auto done (*>*) fun lookup :: "('a * 'b) list \ 'a \ 'b option" where "lookup [] x = None" | "lookup ((a,b) # ps) x = (if a = x then Some b else lookup ps x)" text\ Note that \\\<^sub>1 * \\<^sub>2\ is the type of pairs, also written \\\<^sub>1 \ \\<^sub>2\. Pairs can be taken apart either by pattern matching (as above) or with the projection functions \<^const>\fst\ and \<^const>\snd\: @{thm fst_conv[of x y]} and @{thm snd_conv[of x y]}. Tuples are simulated by pairs nested to the right: \<^term>\(a,b,c)\ is short for \(a, (b, c))\ and \\\<^sub>1 \ \\<^sub>2 \ \\<^sub>3\ is short for \\\<^sub>1 \ (\\<^sub>2 \ \\<^sub>3)\. \subsection{Definitions} Non-recursive functions can be defined as in the following example: \index{definition@\isacom{definition}}\ definition sq :: "nat \ nat" where "sq n = n * n" text\Such definitions do not allow pattern matching but only \f x\<^sub>1 \ x\<^sub>n = t\, where \f\ does not occur in \t\. \subsection{Abbreviations} Abbreviations are similar to definitions: \index{abbreviation@\isacom{abbreviation}}\ abbreviation sq' :: "nat \ nat" where "sq' n \ n * n" text\The key difference is that \<^const>\sq'\ is only syntactic sugar: after parsing, \<^term>\sq' t\ is replaced by \mbox{\<^term>\t*t\}; before printing, every occurrence of \<^term>\u*u\ is replaced by \mbox{\<^term>\sq' u\}. Internally, \<^const>\sq'\ does not exist. This is the advantage of abbreviations over definitions: definitions need to be expanded explicitly (\autoref{sec:rewr-defs}) whereas abbreviations are already expanded upon parsing. However, abbreviations should be introduced sparingly: if abused, they can lead to a confusing discrepancy between the internal and external view of a term. The ASCII representation of \\\ is \texttt{==} or \xsymbol{equiv}. \subsection{Recursive Functions} \label{sec:recursive-funs} Recursive functions are defined with \indexed{\isacom{fun}}{fun} by pattern matching over datatype constructors. The order of equations matters, as in functional programming languages. However, all HOL functions must be total. This simplifies the logic --- terms are always defined --- but means that recursive functions must terminate. Otherwise one could define a function \<^prop>\f n = f n + (1::nat)\ and conclude \mbox{\<^prop>\(0::nat) = 1\} by subtracting \<^term>\f n\ on both sides. Isabelle's automatic termination checker requires that the arguments of recursive calls on the right-hand side must be strictly smaller than the arguments on the left-hand side. In the simplest case, this means that one fixed argument position decreases in size with each recursive call. The size is measured as the number of constructors (excluding 0-ary ones, e.g., \Nil\). Lexicographic combinations are also recognized. In more complicated situations, the user may have to prove termination by hand. For details -see~@{cite Krauss}. +see~\<^cite>\Krauss\. Functions defined with \isacom{fun} come with their own induction schema that mirrors the recursion schema and is derived from the termination order. For example, \ fun div2 :: "nat \ nat" where "div2 0 = 0" | "div2 (Suc 0) = 0" | "div2 (Suc(Suc n)) = Suc(div2 n)" text\does not just define \<^const>\div2\ but also proves a customized induction rule: \[ \inferrule{ \mbox{@{thm (prem 1) div2.induct}}\\ \mbox{@{thm (prem 2) div2.induct}}\\ \mbox{@{thm (prem 3) div2.induct}}} {\mbox{@{thm (concl) div2.induct[of _ "m"]}}} \] This customized induction rule can simplify inductive proofs. For example, \ lemma "div2 n = n div 2" apply(induction n rule: div2.induct) txt\(where the infix \div\ is the predefined division operation) yields the subgoals @{subgoals[display,margin=65]} An application of \auto\ finishes the proof. Had we used ordinary structural induction on \n\, the proof would have needed an additional case analysis in the induction step. This example leads to the following induction heuristic: \begin{quote} \emph{Let \f\ be a recursive function. If the definition of \f\ is more complicated than having one equation for each constructor of some datatype, then properties of \f\ are best proved via \f.induct\.\index{inductionrule@\.induct\}} \end{quote} The general case is often called \concept{computation induction}, because the induction follows the (terminating!) computation. For every defining equation \begin{quote} \f(e) = \ f(r\<^sub>1) \ f(r\<^sub>k) \\ \end{quote} where \f(r\<^sub>i)\, \i=1\k\, are all the recursive calls, the induction rule \f.induct\ contains one premise of the form \begin{quote} \P(r\<^sub>1) \ \ \ P(r\<^sub>k) \ P(e)\ \end{quote} If \f :: \\<^sub>1 \ \ \ \\<^sub>n \ \\ then \f.induct\ is applied like this: \begin{quote} \isacom{apply}\(induction x\<^sub>1 \ x\<^sub>n rule: f.induct)\\index{inductionrule@\induction ... rule:\} \end{quote} where typically there is a call \f x\<^sub>1 \ x\<^sub>n\ in the goal. But note that the induction rule does not mention \f\ at all, except in its name, and is applicable independently of \f\. \subsection*{Exercises} \begin{exercise} Starting from the type \'a tree\ defined in the text, define a function \contents ::\ \<^typ>\'a tree \ 'a list\ that collects all values in a tree in a list, in any order, without removing duplicates. Then define a function \sum_tree ::\ \<^typ>\nat tree \ nat\ that sums up all values in a tree of natural numbers and prove \<^prop>\sum_tree t = sum_list(contents t)\ where \<^const>\sum_list\ is predefined by the equations @{thm sum_list.Nil[where 'a=nat]} and @{thm sum_list.Cons}. \end{exercise} \begin{exercise} Define the two functions \pre_order\ and \post_order\ of type @{typ "'a tree \ 'a list"} that traverse a tree and collect all stored values in the respective order in a list. Prove \<^prop>\pre_order (mirror t) = rev (post_order t)\. \end{exercise} \begin{exercise} Define a function \intersperse ::\ \<^typ>\'a \ 'a list \ 'a list\ such that \intersperse a [x\<^sub>1, ..., x\<^sub>n] = [x\<^sub>1, a, x\<^sub>2, a, ..., a, x\<^sub>n]\. Now prove that \<^prop>\map f (intersperse a xs) = intersperse (f a) (map f xs)\. \end{exercise} \section{Induction Heuristics}\index{induction heuristics} We have already noted that theorems about recursive functions are proved by induction. In case the function has more than one argument, we have followed the following heuristic in the proofs about the append function: \begin{quote} \emph{Perform induction on argument number $i$\\ if the function is defined by recursion on argument number $i$.} \end{quote} The key heuristic, and the main point of this section, is to \emph{generalize the goal before induction}. The reason is simple: if the goal is too specific, the induction hypothesis is too weak to allow the induction step to go through. Let us illustrate the idea with an example. Function \<^const>\rev\ has quadratic worst-case running time because it calls append for each element of the list and append is linear in its first argument. A linear time version of \<^const>\rev\ requires an extra argument where the result is accumulated gradually, using only~\#\: \ (*<*) apply auto done (*>*) fun itrev :: "'a list \ 'a list \ 'a list" where "itrev [] ys = ys" | "itrev (x#xs) ys = itrev xs (x#ys)" text\The behaviour of \<^const>\itrev\ is simple: it reverses its first argument by stacking its elements onto the second argument, and it returns that second argument when the first one becomes empty. Note that \<^const>\itrev\ is tail-recursive: it can be compiled into a loop; no stack is necessary for executing it. Naturally, we would like to show that \<^const>\itrev\ reverses its first argument: \ lemma "itrev xs [] = rev xs" txt\There is no choice as to the induction variable: \ apply(induction xs) apply(auto) txt\ Unfortunately, this attempt does not prove the induction step: @{subgoals[display,margin=70]} The induction hypothesis is too weak. The fixed argument,~\<^term>\[]\, prevents it from rewriting the conclusion. This example suggests a heuristic: \begin{quote} \emph{Generalize goals for induction by replacing constants by variables.} \end{quote} Of course one cannot do this naively: \<^prop>\itrev xs ys = rev xs\ is just not true. The correct generalization is \ (*<*)oops(*>*) lemma "itrev xs ys = rev xs @ ys" (*<*)apply(induction xs, auto)(*>*) txt\ If \ys\ is replaced by \<^term>\[]\, the right-hand side simplifies to \<^term>\rev xs\, as required. In this instance it was easy to guess the right generalization. Other situations can require a good deal of creativity. Although we now have two variables, only \xs\ is suitable for induction, and we repeat our proof attempt. Unfortunately, we are still not there: @{subgoals[display,margin=65]} The induction hypothesis is still too weak, but this time it takes no intuition to generalize: the problem is that the \ys\ in the induction hypothesis is fixed, but the induction hypothesis needs to be applied with \<^term>\a # ys\ instead of \ys\. Hence we prove the theorem for all \ys\ instead of a fixed one. We can instruct induction to perform this generalization for us by adding \arbitrary: ys\\index{arbitrary@\arbitrary:\}. \ (*<*)oops lemma "itrev xs ys = rev xs @ ys" (*>*) apply(induction xs arbitrary: ys) txt\The induction hypothesis is now universally quantified over \ys\: @{subgoals[display,margin=65]} Thus the proof succeeds: \ apply(auto) done text\ This leads to another heuristic for generalization: \begin{quote} \emph{Generalize induction by generalizing all free variables\\ {\em(except the induction variable itself)}.} \end{quote} Generalization is best performed with \arbitrary: y\<^sub>1 \ y\<^sub>k\. This heuristic prevents trivial failures like the one above. However, it should not be applied blindly. It is not always required, and the additional quantifiers can complicate matters in some cases. The variables that need to be quantified are typically those that change in recursive calls. \subsection*{Exercises} \begin{exercise} Write a tail-recursive variant of the \add\ function on \<^typ>\nat\: \<^term>\itadd :: nat \ nat \ nat\. Tail-recursive means that in the recursive case, \itadd\ needs to call itself directly: \mbox{\<^term>\itadd (Suc m) n\} \= itadd \\. Prove \<^prop>\itadd m n = add m n\. \end{exercise} \section{Simplification} So far we have talked a lot about simplifying terms without explaining the concept. \conceptidx{Simplification}{simplification} means \begin{itemize} \item using equations $l = r$ from left to right (only), \item as long as possible. \end{itemize} To emphasize the directionality, equations that have been given the \simp\ attribute are called \conceptidx{simplification rules}{simplification rule}. Logically, they are still symmetric, but proofs by simplification use them only in the left-to-right direction. The proof tool that performs simplifications is called the \concept{simplifier}. It is the basis of \auto\ and other related proof methods. The idea of simplification is best explained by an example. Given the simplification rules \[ \begin{array}{rcl@ {\quad}l} \<^term>\0 + n::nat\ &\=\& \n\ & (1) \\ \<^term>\(Suc m) + n\ &\=\& \<^term>\Suc (m + n)\ & (2) \\ \(Suc m \ Suc n)\ &\=\& \(m \ n)\ & (3)\\ \(0 \ m)\ &\=\& \<^const>\True\ & (4) \end{array} \] the formula \<^prop>\0 + Suc 0 \ Suc 0 + x\ is simplified to \<^const>\True\ as follows: \[ \begin{array}{r@ {}c@ {}l@ {\quad}l} \(0 + Suc 0\ & \leq & \Suc 0 + x)\ & \stackrel{(1)}{=} \\ \(Suc 0\ & \leq & \Suc 0 + x)\ & \stackrel{(2)}{=} \\ \(Suc 0\ & \leq & \Suc (0 + x))\ & \stackrel{(3)}{=} \\ \(0\ & \leq & \0 + x)\ & \stackrel{(4)}{=} \\[1ex] & \<^const>\True\ \end{array} \] Simplification is often also called \concept{rewriting} and simplification rules \conceptidx{rewrite rules}{rewrite rule}. \subsection{Simplification Rules} The attribute \simp\ declares theorems to be simplification rules, which the simplifier will use automatically. In addition, \isacom{datatype} and \isacom{fun} commands implicitly declare some simplification rules: \isacom{datatype} the distinctness and injectivity rules, \isacom{fun} the defining equations. Definitions are not declared as simplification rules automatically! Nearly any theorem can become a simplification rule. The simplifier will try to transform it into an equation. For example, the theorem \<^prop>\\ P\ is turned into \<^prop>\P = False\. Only equations that really simplify, like \<^prop>\rev (rev xs) = xs\ and \<^prop>\xs @ [] = xs\, should be declared as simplification rules. Equations that may be counterproductive as simplification rules should only be used in specific proof steps (see \autoref{sec:simp} below). Distributivity laws, for example, alter the structure of terms and can produce an exponential blow-up. \subsection{Conditional Simplification Rules} Simplification rules can be conditional. Before applying such a rule, the simplifier will first try to prove the preconditions, again by simplification. For example, given the simplification rules \begin{quote} \<^prop>\p(0::nat) = True\\\ \<^prop>\p(x) \ f(x) = g(x)\, \end{quote} the term \<^term>\f(0::nat)\ simplifies to \<^term>\g(0::nat)\ but \<^term>\f(1::nat)\ does not simplify because \<^prop>\p(1::nat)\ is not provable. \subsection{Termination} Simplification can run forever, for example if both \<^prop>\f x = g x\ and \<^prop>\g(x) = f(x)\ are simplification rules. It is the user's responsibility not to include simplification rules that can lead to nontermination, either on their own or in combination with other simplification rules. The right-hand side of a simplification rule should always be ``simpler'' than the left-hand side --- in some sense. But since termination is undecidable, such a check cannot be automated completely and Isabelle makes little attempt to detect nontermination. When conditional simplification rules are applied, their preconditions are proved first. Hence all preconditions need to be simpler than the left-hand side of the conclusion. For example \begin{quote} \<^prop>\n < m \ (n < Suc m) = True\ \end{quote} is suitable as a simplification rule: both \<^prop>\n and \<^const>\True\ are simpler than \mbox{\<^prop>\n < Suc m\}. But \begin{quote} \<^prop>\Suc n < m \ (n < m) = True\ \end{quote} leads to nontermination: when trying to rewrite \<^prop>\n to \<^const>\True\ one first has to prove \mbox{\<^prop>\Suc n < m\}, which can be rewritten to \<^const>\True\ provided \<^prop>\Suc(Suc n) < m\, \emph{ad infinitum}. \subsection{The \indexed{\simp\}{simp} Proof Method} \label{sec:simp} So far we have only used the proof method \auto\. Method \simp\ is the key component of \auto\, but \auto\ can do much more. In some cases, \auto\ is overeager and modifies the proof state too much. In such cases the more predictable \simp\ method should be used. Given a goal \begin{quote} \1. \ P\<^sub>1; \; P\<^sub>m \ \ C\ \end{quote} the command \begin{quote} \isacom{apply}\(simp add: th\<^sub>1 \ th\<^sub>n)\ \end{quote} simplifies the assumptions \P\<^sub>i\ and the conclusion \C\ using \begin{itemize} \item all simplification rules, including the ones coming from \isacom{datatype} and \isacom{fun}, \item the additional lemmas \th\<^sub>1 \ th\<^sub>n\, and \item the assumptions. \end{itemize} In addition to or instead of \add\ there is also \del\ for removing simplification rules temporarily. Both are optional. Method \auto\ can be modified similarly: \begin{quote} \isacom{apply}\(auto simp add: \ simp del: \)\ \end{quote} Here the modifiers are \simp add\ and \simp del\ instead of just \add\ and \del\ because \auto\ does not just perform simplification. Note that \simp\ acts only on subgoal 1, \auto\ acts on all subgoals. There is also \simp_all\, which applies \simp\ to all subgoals. \subsection{Rewriting with Definitions} \label{sec:rewr-defs} Definitions introduced by the command \isacom{definition} can also be used as simplification rules, but by default they are not: the simplifier does not expand them automatically. Definitions are intended for introducing abstract concepts and not merely as abbreviations. Of course, we need to expand the definition initially, but once we have proved enough abstract properties of the new constant, we can forget its original definition. This style makes proofs more robust: if the definition has to be changed, only the proofs of the abstract properties will be affected. The definition of a function \f\ is a theorem named \f_def\ and can be added to a call of \simp\ like any other theorem: \begin{quote} \isacom{apply}\(simp add: f_def)\ \end{quote} In particular, let-expressions can be unfolded by making @{thm[source] Let_def} a simplification rule. \subsection{Case Splitting With \simp\} Goals containing if-expressions are automatically split into two cases by \simp\ using the rule \begin{quote} \<^prop>\P(if A then s else t) = ((A \ P(s)) \ (~A \ P(t)))\ \end{quote} For example, \simp\ can prove \begin{quote} \<^prop>\(A \ B) = (if A then B else False)\ \end{quote} because both \<^prop>\A \ (A & B) = B\ and \<^prop>\~A \ (A & B) = False\ simplify to \<^const>\True\. We can split case-expressions similarly. For \nat\ the rule looks like this: @{prop[display,margin=65,indent=4]"P(case e of 0 \ a | Suc n \ b n) = ((e = 0 \ P a) \ (\n. e = Suc n \ P(b n)))"} Case expressions are not split automatically by \simp\, but \simp\ can be instructed to do so: \begin{quote} \isacom{apply}\(simp split: nat.split)\ \end{quote} splits all case-expressions over natural numbers. For an arbitrary datatype \t\ it is \t.split\\index{split@\.split\} instead of @{thm[source] nat.split}. Method \auto\ can be modified in exactly the same way. The modifier \indexed{\split:\}{split} can be followed by multiple names. Splitting if or case-expressions in the assumptions requires \split: if_splits\ or \split: t.splits\. \ifsem\else \subsection{Rewriting \let\ and Numerals} Let-expressions (\<^term>\let x = s in t\) can be expanded explicitly with the simplification rule @{thm[source] Let_def}. The simplifier will not expand \let\s automatically in many cases. Numerals of type \<^typ>\nat\ can be converted to \<^const>\Suc\ terms with the simplification rule @{thm[source] numeral_eq_Suc}. This is required, for example, when a function that is defined by pattern matching with \<^const>\Suc\ is applied to a numeral: if \f\ is defined by \f 0 = ...\ and \f (Suc n) = ...\, the simplifier cannot simplify \f 2\ unless \2\ is converted to \<^term>\Suc(Suc 0)\ (via @{thm[source] numeral_eq_Suc}). \fi \subsection*{Exercises} \exercise\label{exe:tree0} Define a datatype \tree0\ of binary tree skeletons which do not store any information, neither in the inner nodes nor in the leaves. Define a function \nodes :: tree0 \ nat\ that counts the number of all nodes (inner nodes and leaves) in such a tree. Consider the following recursive function: \ (*<*) datatype tree0 = Tip | Node tree0 tree0 (*>*) fun explode :: "nat \ tree0 \ tree0" where "explode 0 t = t" | "explode (Suc n) t = explode n (Node t t)" text \ Find an equation expressing the size of a tree after exploding it (\noquotes{@{term [source] "nodes (explode n t)"}}) as a function of \<^term>\nodes t\ and \n\. Prove your equation. You may use the usual arithmetic operators, including the exponentiation operator ``\^\''. For example, \noquotes{@{prop [source] "2 ^ 2 = 4"}}. Hint: simplifying with the list of theorems @{thm[source] algebra_simps} takes care of common algebraic properties of the arithmetic operators. \endexercise \exercise Define arithmetic expressions in one variable over integers (type \<^typ>\int\) as a data type: \ datatype exp = Var | Const int | Add exp exp | Mult exp exp text\ Define a function \noquotes{@{term [source]"eval :: exp \ int \ int"}} such that \<^term>\eval e x\ evaluates \e\ at the value \x\. A polynomial can be represented as a list of coefficients, starting with the constant. For example, \<^term>\[4, 2, -1, 3::int]\ represents the polynomial $4 + 2x - x^2 + 3x^3$. Define a function \noquotes{@{term [source] "evalp :: int list \ int \ int"}} that evaluates a polynomial at the given value. Define a function \noquotes{@{term[source] "coeffs :: exp \ int list"}} that transforms an expression into a polynomial. This may require auxiliary functions. Prove that \coeffs\ preserves the value of the expression: \mbox{\<^prop>\evalp (coeffs e) x = eval e x\.} Hint: consider the hint in Exercise~\ref{exe:tree0}. \endexercise \ (*<*) end (*>*) diff --git a/src/Doc/Sugar/Sugar.thy b/src/Doc/Sugar/Sugar.thy --- a/src/Doc/Sugar/Sugar.thy +++ b/src/Doc/Sugar/Sugar.thy @@ -1,583 +1,583 @@ (*<*) theory Sugar imports "HOL-Library.LaTeXsugar" "HOL-Library.OptionalSugar" begin no_translations ("prop") "P \ Q \ R" <= ("prop") "P \ Q \ R" (*>*) text\ \section{Introduction} This document is for those Isabelle users who have mastered the art of mixing \LaTeX\ text and Isabelle theories and never want to typeset a theorem by hand anymore because they have experienced the bliss of writing \verb!@!\verb!{thm[display,mode=latex_sum] sum_Suc_diff [no_vars]}! and seeing Isabelle typeset it for them: @{thm[display,mode=latex_sum] sum_Suc_diff[no_vars]} No typos, no omissions, no sweat. If you have not experienced that joy, read Chapter 4, \emph{Presenting -Theories}, @{cite LNCS2283} first. +Theories}, \<^cite>\LNCS2283\ first. If you have mastered the art of Isabelle's \emph{antiquotations}, i.e.\ things like the above \verb!@!\verb!{thm...}!, beware: in your vanity you may be tempted to think that all readers of the stunning documents you can now produce at the drop of a hat will be struck with awe at the beauty unfolding in front of their eyes. Until one day you come across that very critical of readers known as the ``common referee''. He has the nasty habit of refusing to understand unfamiliar notation like Isabelle's infamous \\ \ \\ no matter how many times you explain it in your paper. Even worse, he thinks that using \\ \\ for anything other than denotational semantics is a cardinal sin that must be punished by instant rejection. This document shows you how to make Isabelle and \LaTeX\ cooperate to produce ordinary looking mathematics that hides the fact that it was typeset by a machine. You merely need to load the right files: \begin{itemize} \item Import theory \texttt{LaTeXsugar} in the header of your own theory. You may also want bits of \texttt{OptionalSugar}, which you can copy selectively into your own theory or import as a whole. Both theories live in \texttt{HOL/Library}. \item Should you need additional \LaTeX\ packages (the text will tell you so), you include them at the beginning of your \LaTeX\ document, typically in \texttt{root.tex}. For a start, you should \verb!\usepackage{amssymb}! --- otherwise typesetting @{prop[source]"\(\x. P x)"} will fail because the AMS symbol \\\ is missing. \end{itemize} \section{HOL syntax} \subsection{Logic} The formula @{prop[source]"\(\x. P x)"} is typeset as \<^prop>\\(\x. P x)\. The predefined constructs \if\, \let\ and \case\ are set in sans serif font to distinguish them from other functions. This improves readability: \begin{itemize} \item \<^term>\if b then e\<^sub>1 else e\<^sub>2\ instead of \if b then e\<^sub>1 else e\<^sub>2\. \item \<^term>\let x = e\<^sub>1 in e\<^sub>2\ instead of \let x = e\<^sub>1 in e\<^sub>2\. \item \<^term>\case x of True \ e\<^sub>1 | False \ e\<^sub>2\ instead of\\ \case x of True \ e\<^sub>1 | False \ e\<^sub>2\. \end{itemize} \subsection{Sets} Although set syntax in HOL is already close to standard, we provide a few further improvements: \begin{itemize} \item \<^term>\{x. P}\ instead of \{x. P}\. \item \<^term>\{}\ instead of \{}\, where \<^term>\{}\ is also input syntax. \item \<^term>\insert a (insert b (insert c M))\ instead of \insert a (insert b (insert c M))\. \item \<^term>\card A\ instead of \card A\. \end{itemize} \subsection{Lists} If lists are used heavily, the following notations increase readability: \begin{itemize} \item \<^term>\x # xs\ instead of \x # xs\, where \<^term>\x # xs\ is also input syntax. \item \<^term>\length xs\ instead of \length xs\. \item \<^term>\nth xs n\ instead of \nth xs n\, the $n$th element of \xs\. \item Human readers are good at converting automatically from lists to sets. Hence \texttt{OptionalSugar} contains syntax for suppressing the conversion function \<^const>\set\: for example, @{prop[source]"x \ set xs"} becomes \<^prop>\x \ set xs\. \item The \@\ operation associates implicitly to the right, which leads to unpleasant line breaks if the term is too long for one line. To avoid this, \texttt{OptionalSugar} contains syntax to group \@\-terms to the left before printing, which leads to better line breaking behaviour: @{term[display]"term\<^sub>0 @ term\<^sub>1 @ term\<^sub>2 @ term\<^sub>3 @ term\<^sub>4 @ term\<^sub>5 @ term\<^sub>6 @ term\<^sub>7 @ term\<^sub>8 @ term\<^sub>9 @ term\<^sub>1\<^sub>0"} \end{itemize} \subsection{Numbers} Coercions between numeric types are alien to mathematicians who consider, for example, \<^typ>\nat\ as a subset of \<^typ>\int\. \texttt{OptionalSugar} contains syntax for suppressing numeric coercions such as \<^const>\int\ \::\ \<^typ>\nat \ int\. For example, @{term[source]"int 5"} is printed as \<^term>\int 5\. Embeddings of types \<^typ>\nat\, \<^typ>\int\, \<^typ>\real\ are covered; non-injective coercions such as \<^const>\nat\ \::\ \<^typ>\int \ nat\ are not and should not be hidden. \section{Printing constants and their type} Instead of \verb!@!\verb!{const myconst}! \verb!@!\verb!{text "::"}! \verb!@!\verb!{typeof myconst}!, you can write \verb!@!\verb!{const_typ myconst}! using the new antiquotation \texttt{const\_typ} defined in \texttt{LaTeXsugar}. For example, \verb!@!\verb!{const_typ length}! produces \<^const_typ>\length\ (see below for how to suppress the question mark). This works both for genuine constants and for variables fixed in some context, especially in a locale. \section{Printing theorems} The \<^prop>\P \ Q \ R\ syntax is a bit idiosyncratic. If you would like to avoid it, you can easily print the premises as a conjunction: \<^prop>\P \ Q \ R\. See \texttt{OptionalSugar} for the required ``code''. \subsection{Question marks} If you print anything, especially theorems, containing schematic variables they are prefixed with a question mark: \verb!@!\verb!{thm conjI}! results in @{thm conjI}. Most of the time you would rather not see the question marks. There is an attribute \verb!no_vars! that you can attach to the theorem that turns its schematic into ordinary free variables: \begin{quote} \verb!@!\verb!{thm conjI[no_vars]}!\\ \showout @{thm conjI[no_vars]} \end{quote} This \verb!no_vars! business can become a bit tedious. If you would rather never see question marks, simply put \begin{quote} \verb!options [show_question_marks = false]! \end{quote} into the relevant \texttt{ROOT} file, just before the \texttt{theories} for that session. The rest of this document is produced with this flag set to \texttt{false}. \ (*<*)declare [[show_question_marks = false]](*>*) subsection \Qualified names\ text\If there are multiple declarations of the same name, Isabelle prints the qualified name, for example \T.length\, where \T\ is the theory it is defined in, to distinguish it from the predefined @{const[source] "List.length"}. In case there is no danger of confusion, you can insist on short names (no qualifiers) by setting the \verb!names_short! configuration option in the context. \subsection {Variable names\label{sec:varnames}} It sometimes happens that you want to change the name of a variable in a theorem before printing it. This can easily be achieved with the help of Isabelle's instantiation attribute \texttt{where}: \begin{quote} \verb!@!\verb!{thm conjI[where P = \ and Q = \]}!\\ \showout @{thm conjI[where P = \ and Q = \]} \end{quote} To support the ``\_''-notation for irrelevant variables the constant \texttt{DUMMY} has been introduced: \begin{quote} \verb!@!\verb!{thm fst_conv[of _ DUMMY]}!\\ \showout @{thm fst_conv[of _ DUMMY]} \end{quote} As expected, the second argument has been replaced by ``\_'', but the first argument is the ugly \x1.0\, a schematic variable with suppressed question mark. Schematic variables that end in digits, e.g. \x1\, are still printed with a trailing \.0\, e.g. \x1.0\, their internal index. This can be avoided by turning the last digit into a subscript: write \<^verbatim>\x\<^sub>1\ and obtain the much nicer \x\<^sub>1\. Alternatively, you can display trailing digits of schematic and free variables as subscripts with the \texttt{sub} style: \begin{quote} \verb!@!\verb!{thm (sub) fst_conv[of _ DUMMY]}!\\ \showout @{thm (sub) fst_conv[of _ DUMMY]} \end{quote} The insertion of underscores can be automated with the \verb!dummy_pats! style: \begin{quote} \verb!@!\verb!{thm (dummy_pats,sub) fst_conv}!\\ \showout @{thm (dummy_pats,sub) fst_conv} \end{quote} The theorem must be an equation. Then every schematic variable that occurs on the left-hand but not the right-hand side is replaced by \texttt{DUMMY}. This is convenient for displaying functional programs. Variables that are bound by quantifiers or lambdas can be renamed with the help of the attribute \verb!rename_abs!. It expects a list of names or underscores, similar to the \texttt{of} attribute: \begin{quote} \verb!@!\verb!{thm split_paired_All[rename_abs _ l r]}!\\ \showout @{thm split_paired_All[rename_abs _ l r]} \end{quote} Sometimes Isabelle $\eta$-contracts terms, for example in the following definition: \ fun eta where "eta (x \ xs) = (\y \ set xs. x < y)" text\ \noindent If you now print the defining equation, the result is not what you hoped for: \begin{quote} \verb!@!\verb!{thm eta.simps}!\\ \showout @{thm eta.simps} \end{quote} In such situations you can put the abstractions back by explicitly $\eta$-expanding upon output: \begin{quote} \verb!@!\verb!{thm (eta_expand z) eta.simps}!\\ \showout @{thm (eta_expand z) eta.simps} \end{quote} Instead of a single variable \verb!z! you can give a whole list \verb!x y z! to perform multiple $\eta$-expansions. \subsection{Inference rules} To print theorems as inference rules you need to include Didier -R\'emy's \texttt{mathpartir} package~@{cite mathpartir} +R\'emy's \texttt{mathpartir} package~\<^cite>\mathpartir\ for typesetting inference rules in your \LaTeX\ file. Writing \verb!@!\verb!{thm[mode=Rule] conjI}! produces @{thm[mode=Rule] conjI}, even in the middle of a sentence. If you prefer your inference rule on a separate line, maybe with a name, \begin{center} @{thm[mode=Rule] conjI} {\sc conjI} \end{center} is produced by \begin{quote} \verb!\begin{center}!\\ \verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI}!\\ \verb!\end{center}! \end{quote} It is not recommended to use the standard \texttt{display} option together with \texttt{Rule} because centering does not work and because the line breaking mechanisms of \texttt{display} and \texttt{mathpartir} can clash. Of course you can display multiple rules in this fashion: \begin{quote} \verb!\begin{center}!\\ \verb!@!\verb!{thm[mode=Rule] conjI} {\sc conjI} \\[1ex]!\\ \verb!@!\verb!{thm[mode=Rule] conjE} {\sc disjI$_1$} \qquad!\\ \verb!@!\verb!{thm[mode=Rule] disjE} {\sc disjI$_2$}!\\ \verb!\end{center}! \end{quote} yields \begin{center}\small @{thm[mode=Rule] conjI} {\sc conjI} \\[1ex] @{thm[mode=Rule] disjI1} {\sc disjI$_1$} \qquad @{thm[mode=Rule] disjI2} {\sc disjI$_2$} \end{center} The \texttt{mathpartir} package copes well if there are too many premises for one line: \begin{center} @{prop[mode=Rule] "\ A \ B; B \ C; C \ D; D \ E; E \ F; F \ G; G \ H; H \ I; I \ J; J \ K \ \ A \ K"} \end{center} Limitations: 1. Premises and conclusion must each not be longer than the line. 2. Premises that are \\\-implications are again displayed with a horizontal line, which looks at least unusual. In case you print theorems without premises no rule will be printed by the \texttt{Rule} print mode. However, you can use \texttt{Axiom} instead: \begin{quote} \verb!\begin{center}!\\ \verb!@!\verb!{thm[mode=Axiom] refl} {\sc refl}! \\ \verb!\end{center}! \end{quote} yields \begin{center} @{thm[mode=Axiom] refl} {\sc refl} \end{center} \subsection{Displays and font sizes} When displaying theorems with the \texttt{display} option, for example as in \verb!@!\verb!{thm[display] refl}! @{thm[display] refl} the theorem is set in small font. It uses the \LaTeX-macro \verb!\isastyle!, which is also the style that regular theory text is set in, e.g.\ lemma "t = t" (*<*)oops(*>*) text\\noindent Otherwise \verb!\isastyleminor! is used, which does not modify the font size (assuming you stick to the default \verb!\isabellestyle{it}! in \texttt{root.tex}). If you prefer normal font size throughout your text, include \begin{quote} \verb!\renewcommand{\isastyle}{\isastyleminor}! \end{quote} in \texttt{root.tex}. On the other hand, if you like the small font, just put \verb!\isastyle! in front of the text in question, e.g.\ at the start of one of the center-environments above. The advantage of the display option is that you can display a whole list of theorems in one go. For example, \verb!@!\verb!{thm[display] append.simps}! generates @{thm[display] append.simps} \subsection{If-then} If you prefer a fake ``natural language'' style you can produce the body of \newtheorem{theorem}{Theorem} \begin{theorem} @{thm[mode=IfThen] le_trans} \end{theorem} by typing \begin{quote} \verb!@!\verb!{thm[mode=IfThen] le_trans}! \end{quote} In order to prevent odd line breaks, the premises are put into boxes. At times this is too drastic: \begin{theorem} @{prop[mode=IfThen] "longpremise \ longerpremise \ P(f(f(f(f(f(f(f(f(f(x)))))))))) \ longestpremise \ conclusion"} \end{theorem} In which case you should use \texttt{IfThenNoBox} instead of \texttt{IfThen}: \begin{theorem} @{prop[mode=IfThenNoBox] "longpremise \ longerpremise \ P(f(f(f(f(f(f(f(f(f(x)))))))))) \ longestpremise \ conclusion"} \end{theorem} \subsection{Doing it yourself\label{sec:yourself}} If for some reason you want or need to present theorems your own way, you can extract the premises and the conclusion explicitly and combine them as you like: \begin{itemize} \item \verb!@!\verb!{thm (prem 1)! $thm$\verb!}! prints premise 1 of $thm$. \item \verb!@!\verb!{thm (concl)! $thm$\verb!}! prints the conclusion of $thm$. \end{itemize} For example, ``from @{thm (prem 2) conjI} and @{thm (prem 1) conjI} we conclude @{thm (concl) conjI}'' is produced by \begin{quote} \verb!from !\verb!@!\verb!{thm (prem 2) conjI}! \verb!and !\verb!@!\verb!{thm (prem 1) conjI}!\\ \verb!we conclude !\verb!@!\verb!{thm (concl) conjI}! \end{quote} Thus you can rearrange or hide premises and typeset the theorem as you like. Styles like \verb!(prem 1)! are a general mechanism explained in \S\ref{sec:styles}. \subsection{Patterns} In \S\ref{sec:varnames} we shows how to create patterns containing ``\<^term>\DUMMY\''. You can drive this game even further and extend the syntax of let bindings such that certain functions like \<^term>\fst\, \<^term>\hd\, etc.\ are printed as patterns. \texttt{OptionalSugar} provides the following: \begin{center} \begin{tabular}{l@ {~~produced by~~}l} \<^term>\let x = fst p in t\ & \verb!@!\verb!{term "let x = fst p in t"}!\\ \<^term>\let x = snd p in t\ & \verb!@!\verb!{term "let x = snd p in t"}!\\ \<^term>\let x = hd xs in t\ & \verb!@!\verb!{term "let x = hd xs in t"}!\\ \<^term>\let x = tl xs in t\ & \verb!@!\verb!{term "let x = tl xs in t"}!\\ \<^term>\let x = the y in t\ & \verb!@!\verb!{term "let x = the y in t"}!\\ \end{tabular} \end{center} \section {Styles\label{sec:styles}} The \verb!thm! antiquotation works nicely for single theorems, but sets of equations as used in definitions are more difficult to typeset nicely: people tend to prefer aligned \=\ signs. To deal with such cases where it is desirable to dive into the structure of terms and theorems, Isabelle offers antiquotations featuring ``styles'': \begin{quote} \verb!@!\verb!{thm (style) thm}!\\ \verb!@!\verb!{prop (style) thm}!\\ \verb!@!\verb!{term (style) term}!\\ \verb!@!\verb!{term_type (style) term}!\\ \verb!@!\verb!{typeof (style) term}!\\ \end{quote} A ``style'' is a transformation of a term. There are predefined styles, namely \verb!lhs! and \verb!rhs!, \verb!prem! with one argument, and \verb!concl!. For example, the output \begin{center} \begin{tabular}{l@ {~~\=\~~}l} @{thm (lhs) append_Nil} & @{thm (rhs) append_Nil}\\ @{thm (lhs) append_Cons} & @{thm (rhs) append_Cons} \end{tabular} \end{center} is produced by the following code: \begin{quote} \verb!\begin{center}!\\ \verb!\begin{tabular}{l@ {~~!\verb!@!\verb!{text "="}~~}l}!\\ \verb!@!\verb!{thm (lhs) append_Nil} & @!\verb!{thm (rhs) append_Nil}\\!\\ \verb!@!\verb!{thm (lhs) append_Cons} & @!\verb!{thm (rhs) append_Cons}!\\ \verb!\end{tabular}!\\ \verb!\end{center}! \end{quote} Note the space between \verb!@! and \verb!{! in the tabular argument. It prevents Isabelle from interpreting \verb!@ {~~...~~}! as an antiquotation. The styles \verb!lhs! and \verb!rhs! extract the left hand side (or right hand side respectively) from the conclusion of propositions consisting of a binary operator (e.~g.~\=\, \\\, \<\). Likewise, \verb!concl! may be used as a style to show just the conclusion of a proposition. For example, take \verb!hd_Cons_tl!: \begin{center} @{thm hd_Cons_tl} \end{center} To print just the conclusion, \begin{center} @{thm (concl) hd_Cons_tl} \end{center} type \begin{quote} \verb!\begin{center}!\\ \verb!@!\verb!{thm (concl) hd_Cons_tl}!\\ \verb!\end{center}! \end{quote} Beware that any options must be placed \emph{before} the style, as in this example. Further use cases can be found in \S\ref{sec:yourself}. If you are not afraid of ML, you may also define your own styles. Have a look at module \<^ML_structure>\Term_Style\. \section {Proofs} Full proofs, even if written in beautiful Isar style, are likely to be too long and detailed to be included in conference papers, but some key lemmas might be of interest. It is usually easiest to put them in figures like the one in Fig.\ \ref{fig:proof}. This was achieved with the \isakeyword{text\_raw} command: \ text_raw \ \begin{figure} \begin{center}\begin{minipage}{0.6\textwidth} \isastyleminor\isamarkuptrue \ lemma True proof - \ \pretty trivial\ show True by force qed text_raw \ \end{minipage}\end{center} \caption{Example proof in a figure.}\label{fig:proof} \end{figure} \ text \ \begin{quote} \small \verb!text_raw {!\verb!*!\\ \verb! \begin{figure}!\\ \verb! \begin{center}\begin{minipage}{0.6\textwidth}!\\ \verb! \isastyleminor\isamarkuptrue!\\ \verb!*!\verb!}!\\ \verb!lemma True!\\ \verb!proof -!\\ \verb! -- "pretty trivial"!\\ \verb! show True by force!\\ \verb!qed!\\ \verb!text_raw {!\verb!*!\\ \verb! \end{minipage}\end{center}!\\ \verb! \caption{Example proof in a figure.}\label{fig:proof}!\\ \verb! \end{figure}!\\ \verb!*!\verb!}! \end{quote} Other theory text, e.g.\ definitions, can be put in figures, too. \section{Theory snippets} This section describes how to include snippets of a theory text in some other \LaTeX\ document. The typical scenario is that the description of your theory is not part of the theory text but a separate document that antiquotes bits of the theory. This works well for terms and theorems but there are no antiquotations, for example, for function definitions or proofs. Even if there are antiquotations, the output is usually a reformatted (by Isabelle) version of the input and may not look like you wanted it to look. Here is how to include a snippet of theory text (in \LaTeX\ form) in some other \LaTeX\ document, in 4 easy steps. Beware that these snippets are not processed by any antiquotation mechanism: the resulting \LaTeX\ text is more or less exactly what you wrote in the theory, without any added sugar. \subsection{Theory markup} Include some markers at the beginning and the end of the theory snippet you want to cut out. You have to place the following lines before and after the snippet, where snippets must always be consecutive lines of theory text: \begin{quote} \verb!\text_raw{!\verb!*\snip{!\emph{snippetname}\verb!}{1}{2}{%*!\verb!}!\\ \emph{theory text}\\ \verb!\text_raw{!\verb!*!\verb!}%endsnip*!\verb!}! \end{quote} where \emph{snippetname} should be a unique name for the snippet. The numbers \texttt{1} and \texttt{2} are explained in a moment. \subsection{Generate the \texttt{.tex} file} Run your theory \texttt{T} with the \texttt{isabelle} \texttt{build} tool to generate the \LaTeX-file \texttt{T.tex} which is needed for the next step, extraction of marked snippets. You may also want to process \texttt{T.tex} to generate a pdf document. This requires a definition of \texttt{\char`\\snippet}: \begin{verbatim} \newcommand{\repeatisanl}[1] {\ifnum#1=0\else\isanewline\repeatisanl{\numexpr#1-1}\fi} \newcommand{\snip}[4]{\repeatisanl#2#4\repeatisanl#3} \end{verbatim} Parameter 2 and 3 of \texttt{\char`\\snippet} are numbers (the \texttt{1} and \texttt{2} above) and determine how many newlines are inserted before and after the snippet. Unfortunately \texttt{text\_raw} eats up all preceding and following newlines and they have to be inserted again in this manner. Otherwise the document generated from \texttt{T.tex} will look ugly around the snippets. It can take some iterations to get the number of required newlines exactly right. \subsection{Extract marked snippets} \label{subsec:extract} Extract the marked bits of text with a shell-level script, e.g. \begin{quote} \verb!sed -n '/\\snip{/,/endsnip/p' T.tex > !\emph{snippets}\verb!.tex! \end{quote} File \emph{snippets}\texttt{.tex} (the name is arbitrary) now contains a sequence of blocks like this \begin{quote} \verb!\snip{!\emph{snippetname}\verb!}{1}{2}{%!\\ \emph{theory text}\\ \verb!}%endsnip! \end{quote} \subsection{Including snippets} In the preamble of the document where the snippets are to be used you define \texttt{\char`\\snip} and input \emph{snippets}\texttt{.tex}: \begin{verbatim} \newcommand{\snip}[4] {\expandafter\newcommand\csname #1\endcsname{#4}} \input{snippets} \end{verbatim} This definition of \texttt{\char`\\snip} simply has the effect of defining for each snippet \emph{snippetname} a \LaTeX\ command \texttt{\char`\\}\emph{snippetname} that produces the corresponding snippet text. In the body of your document you can display that text like this: \begin{quote} \verb!\begin{isabelle}!\\ \texttt{\char`\\}\emph{snippetname}\\ \verb!\end{isabelle}! \end{quote} The \texttt{isabelle} environment is the one defined in the standard file \texttt{isabelle.sty} which most likely you are loading anyway. \ (*<*) end (*>*) diff --git a/src/Doc/System/Environment.thy b/src/Doc/System/Environment.thy --- a/src/Doc/System/Environment.thy +++ b/src/Doc/System/Environment.thy @@ -1,490 +1,488 @@ (*:maxLineLen=78:*) theory Environment imports Base begin chapter \The Isabelle system environment\ text \ This manual describes Isabelle together with related tools as seen from a - system oriented view. See also the \<^emph>\Isabelle/Isar Reference Manual\ @{cite - "isabelle-isar-ref"} for the actual Isabelle input language and related - concepts, and \<^emph>\The Isabelle/Isar Implementation Manual\ @{cite - "isabelle-implementation"} for the main concepts of the underlying + system oriented view. See also the \<^emph>\Isabelle/Isar Reference Manual\ \<^cite>\"isabelle-isar-ref"\ for the actual Isabelle input language and related + concepts, and \<^emph>\The Isabelle/Isar Implementation Manual\ \<^cite>\"isabelle-implementation"\ for the main concepts of the underlying implementation in Isabelle/ML. \ section \Isabelle settings \label{sec:settings}\ text \ Isabelle executables may depend on the \<^emph>\Isabelle settings\ within the process environment. This is a statically scoped collection of environment variables, such as @{setting ISABELLE_HOME}, @{setting ML_SYSTEM}, @{setting ML_HOME}. These variables are \<^emph>\not\ intended to be set directly from the shell, but are provided by Isabelle \<^emph>\components\ within their \<^emph>\settings files\, as explained below. \ subsection \Bootstrapping the environment \label{sec:boot}\ text \ Isabelle executables need to be run within a proper settings environment. This is bootstrapped as described below, on the first invocation of one of the outer wrapper scripts (such as @{executable_ref isabelle}). This happens only once for each process tree, i.e.\ the environment is passed to subprocesses according to regular Unix conventions. \<^enum> The special variable @{setting_def ISABELLE_HOME} is determined automatically from the location of the binary that has been run. You should not try to set @{setting ISABELLE_HOME} manually. Also note that the Isabelle executables either have to be run from their original location in the distribution directory, or via the executable objects created by the @{tool install} tool. Symbolic links are admissible, but a plain copy of the \<^dir>\$ISABELLE_HOME/bin\ files will not work! \<^enum> The file \<^file>\$ISABELLE_HOME/etc/settings\ is run as a @{executable_ref bash} shell script with the auto-export option for variables enabled. This file holds a rather long list of shell variable assignments, thus providing the site-wide default settings. The Isabelle distribution already contains a global settings file with sensible defaults for most variables. When installing the system, only a few of these may have to be adapted (probably @{setting ML_SYSTEM} etc.). \<^enum> The file \<^path>\$ISABELLE_HOME_USER/etc/settings\ (if it exists) is run in the same way as the site default settings. Note that the variable @{setting ISABELLE_HOME_USER} has already been set before --- usually to something like \<^verbatim>\$USER_HOME/.isabelle/Isabelle2022\. Thus individual users may override the site-wide defaults. Typically, a user settings file contains only a few lines, with some assignments that are actually changed. Never copy the central \<^file>\$ISABELLE_HOME/etc/settings\ file! Since settings files are regular GNU @{executable_def bash} scripts, one may use complex shell commands, such as \<^verbatim>\if\ or \<^verbatim>\case\ statements to set variables depending on the system architecture or other environment variables. Such advanced features should be added only with great care, though. In particular, external environment references should be kept at a minimum. \<^medskip> A few variables are somewhat special, e.g.\ @{setting_def ISABELLE_TOOL} is set automatically to the absolute path name of the @{executable isabelle} executables. \<^medskip> Note that the settings environment may be inspected with the @{tool getenv} tool. This might help to figure out the effect of complex settings scripts. \ subsection \Common variables\ text \ This is a reference of common Isabelle settings variables. Note that the list is somewhat open-ended. Third-party utilities or interfaces may add their own selection. Variables that are special in some sense are marked with \\<^sup>*\. \<^descr>[@{setting_def USER_HOME}\\<^sup>*\] Is the cross-platform user home directory. On Unix systems this is usually the same as @{setting HOME}, but on Windows it is the regular home directory of the user, not the one of within the Cygwin root file-system.\<^footnote>\Cygwin itself offers another choice whether its HOME should point to the \<^path>\/home\ directory tree or the Windows user home.\ \<^descr>[@{setting_def ISABELLE_HOME}\\<^sup>*\] is the location of the top-level Isabelle distribution directory. This is automatically determined from the Isabelle executable that has been invoked. Do not attempt to set @{setting ISABELLE_HOME} yourself from the shell! \<^descr>[@{setting_def ISABELLE_HOME_USER}] is the user-specific counterpart of @{setting ISABELLE_HOME}. The default value is relative to \<^path>\$USER_HOME/.isabelle\, under rare circumstances this may be changed in the global setting file. Typically, the @{setting ISABELLE_HOME_USER} directory mimics @{setting ISABELLE_HOME} to some extend. In particular, site-wide defaults may be overridden by a private \<^verbatim>\$ISABELLE_HOME_USER/etc/settings\. \<^descr>[@{setting_def ISABELLE_PLATFORM_FAMILY}\\<^sup>*\] is automatically set to the general platform family (\<^verbatim>\linux\, \<^verbatim>\macos\, \<^verbatim>\windows\). Note that platform-dependent tools usually need to refer to the more specific identification according to @{setting ISABELLE_PLATFORM64}, @{setting ISABELLE_WINDOWS_PLATFORM64}, @{setting ISABELLE_APPLE_PLATFORM64}. \<^descr>[@{setting_def ISABELLE_PLATFORM64}\\<^sup>*\] indicates the standard Posix platform (\<^verbatim>\x86_64\, \<^verbatim>\arm64\), together with a symbolic name for the operating system (\<^verbatim>\linux\, \<^verbatim>\darwin\, \<^verbatim>\cygwin\). \<^descr>[@{setting_def ISABELLE_WINDOWS_PLATFORM64}\\<^sup>*\, @{setting_def ISABELLE_WINDOWS_PLATFORM32}\\<^sup>*\] indicate the native Windows platform: both 64\,bit and 32\,bit executables are supported here. In GNU bash scripts, a preference for native Windows platform variants may be specified like this (first 64 bit, second 32 bit): @{verbatim [display] \"${ISABELLE_WINDOWS_PLATFORM64:-${ISABELLE_WINDOWS_PLATFORM32:- $ISABELLE_PLATFORM64}}"\} \<^descr>[@{setting_def ISABELLE_APPLE_PLATFORM64}\\<^sup>*\] indicates the native Apple Silicon platform (\<^verbatim>\arm64-darwin\ if available), instead of Intel emulation via Rosetta (\<^verbatim>\ISABELLE_PLATFORM64=x86_64-darwin\). \<^descr>[@{setting ISABELLE_TOOL}\\<^sup>*\] is automatically set to the full path name of the @{executable isabelle} executable. \<^descr>[@{setting_def ISABELLE_IDENTIFIER}\\<^sup>*\] refers to the name of this Isabelle distribution, e.g.\ ``\<^verbatim>\Isabelle2022\''. \<^descr>[@{setting_def ML_SYSTEM}, @{setting_def ML_HOME}, @{setting_def ML_OPTIONS}, @{setting_def ML_PLATFORM}, @{setting_def ML_IDENTIFIER}\\<^sup>*\] specify the underlying ML system to be used for Isabelle. There is only a fixed set of admissable @{setting ML_SYSTEM} names (see the \<^file>\$ISABELLE_HOME/etc/settings\ file of the distribution). The actual compiler binary will be run from the directory @{setting ML_HOME}, with @{setting ML_OPTIONS} as first arguments on the command line. The optional @{setting ML_PLATFORM} may specify the binary format of ML heap images, which is useful for cross-platform installations. The value of @{setting ML_IDENTIFIER} is automatically obtained by composing the values of @{setting ML_SYSTEM}, @{setting ML_PLATFORM} and the Isabelle version values. \<^descr>[@{setting_def ISABELLE_JDK_HOME}] points to a full JDK (Java Development Kit) installation with \<^verbatim>\javac\ and \<^verbatim>\jar\ executables. Note that conventional \<^verbatim>\JAVA_HOME\ points to the JRE (Java Runtime Environment), not the JDK. \<^descr>[@{setting_def ISABELLE_JAVA_PLATFORM}] identifies the hardware and operating system platform for the Java installation of Isabelle. That is always the (native) 64 bit variant: \<^verbatim>\x86_64-linux\, \<^verbatim>\x86_64-darwin\, \<^verbatim>\x86_64-windows\. \<^descr>[@{setting_def ISABELLE_BROWSER_INFO}] is the directory where HTML and PDF browser information is stored (see also \secref{sec:info}); its default is \<^path>\$ISABELLE_HOME_USER/browser_info\. For ``system build mode'' (see \secref{sec:tool-build}), @{setting_def ISABELLE_BROWSER_INFO_SYSTEM} is used instead; its default is \<^path>\$ISABELLE_HOME/browser_info\. \<^descr>[@{setting_def ISABELLE_HEAPS}] is the directory where session heap images, log files, and build databases are stored; its default is \<^path>\$ISABELLE_HOME_USER/heaps\. If @{system_option system_heaps} is \<^verbatim>\true\, @{setting_def ISABELLE_HEAPS_SYSTEM} is used instead; its default is \<^path>\$ISABELLE_HOME/heaps\. See also \secref{sec:tool-build}. \<^descr>[@{setting_def ISABELLE_LOGIC}] specifies the default logic to load if none is given explicitely by the user. The default value is \<^verbatim>\HOL\. \<^descr>[@{setting_def ISABELLE_LINE_EDITOR}] specifies the line editor for the @{tool_ref console} interface. \<^descr>[@{setting_def ISABELLE_PDFLATEX}, @{setting_def ISABELLE_LUALATEX}, @{setting_def ISABELLE_BIBTEX}, @{setting_def ISABELLE_MAKEINDEX}] refer to {\LaTeX}-related tools for Isabelle document preparation (see also \secref{sec:tool-document}). \<^descr>[@{setting_def ISABELLE_TOOLS}] is a colon separated list of directories that are scanned by @{executable isabelle} for external utility programs (see also \secref{sec:isabelle-tool}). \<^descr>[@{setting_def ISABELLE_DOCS}] is a colon separated list of directories with documentation files. \<^descr>[@{setting_def PDF_VIEWER}] specifies the program to be used for displaying \<^verbatim>\pdf\ files. \<^descr>[@{setting_def ISABELLE_TMP_PREFIX}\\<^sup>*\] is the prefix from which any running Isabelle ML process derives an individual directory for temporary files. \<^descr>[@{setting_def ISABELLE_TOOL_JAVA_OPTIONS}] is passed to the \<^verbatim>\java\ executable when running Isabelle tools (e.g.\ @{tool build}). This is occasionally helpful to provide more heap space, via additional options like \<^verbatim>\-Xms1g -Xmx4g\. \ subsection \Additional components \label{sec:components}\ text \ Any directory may be registered as an explicit \<^emph>\Isabelle component\. The general layout conventions are that of the main Isabelle distribution itself, and the following two files (both optional) have a special meaning: \<^item> \<^verbatim>\etc/settings\ holds additional settings that are initialized when bootstrapping the overall Isabelle environment, cf.\ \secref{sec:boot}. As usual, the content is interpreted as a GNU bash script. It may refer to the component's enclosing directory via the \<^verbatim>\COMPONENT\ shell variable. For example, the following setting allows to refer to files within the component later on, without having to hardwire absolute paths: @{verbatim [display] \MY_COMPONENT_HOME="$COMPONENT"\} Components can also add to existing Isabelle settings such as @{setting_def ISABELLE_TOOLS}, in order to provide component-specific tools that can be invoked by end-users. For example: @{verbatim [display] \ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"\} \<^item> \<^verbatim>\etc/components\ holds a list of further sub-components of the same structure. The directory specifications given here can be either absolute (with leading \<^verbatim>\/\) or relative to the component's main directory. The root of component initialization is @{setting ISABELLE_HOME} itself. After initializing all of its sub-components recursively, @{setting ISABELLE_HOME_USER} is included in the same manner (if that directory exists). This allows to install private components via \<^path>\$ISABELLE_HOME_USER/etc/components\, although it is often more convenient to do that programmatically via the \<^bash_function>\init_component\ shell function in the \<^verbatim>\etc/settings\ script of \<^verbatim>\$ISABELLE_HOME_USER\ (or any other component directory). For example: @{verbatim [display] \init_component "$HOME/screwdriver-2.0"\} This is tolerant wrt.\ missing component directories, but might produce a warning. \<^medskip> More complex situations may be addressed by initializing components listed in a given catalog file, relatively to some base directory: @{verbatim [display] \init_components "$HOME/my_component_store" "some_catalog_file"\} The component directories listed in the catalog file are treated as relative to the given base directory. See also \secref{sec:tool-components} for some tool-support for resolving components that are formally initialized but not installed yet. \ section \The Isabelle tool wrapper \label{sec:isabelle-tool}\ text \ The main \<^emph>\Isabelle tool wrapper\ provides a generic startup environment for Isabelle-related utilities, user interfaces, add-on applications etc. Such tools automatically benefit from the settings mechanism (\secref{sec:settings}). Moreover, this is the standard way to invoke Isabelle/Scala functionality as a separate operating-system process. Isabelle command-line tools are run uniformly via a common wrapper --- @{executable_ref isabelle}: @{verbatim [display] \Usage: isabelle TOOL [ARGS ...] Start Isabelle TOOL with ARGS; pass "-?" for tool-specific help. Available tools: ...\} Tools may be implemented in Isabelle/Scala or as stand-alone executables (usually as GNU bash scripts). In the invocation of ``@{executable isabelle}~\tool\'', the named \tool\ is resolved as follows (and in the given order). \<^enum> An external tool found on the directories listed in the @{setting ISABELLE_TOOLS} settings variable (colon-separated list in standard POSIX notation). It is invoked as stand-alone program with the command-line arguments provided as \<^verbatim>\argv\ array. \<^enum> An internal tool that is declared via class \<^scala_type>\isabelle.Isabelle_Scala_Tools\ and registered via \<^verbatim>\services\ in \<^path>\etc/build.props\. See \secref{sec:scala-build} for more details. There are also various administrative tools that are available from a bare repository clone of Isabelle, but not in regular distributions. \ subsubsection \Examples\ text \ Show the list of available documentation of the Isabelle distribution: @{verbatim [display] \isabelle doc\} View a certain document as follows: @{verbatim [display] \isabelle doc system\} Query the Isabelle settings environment: @{verbatim [display] \isabelle getenv ISABELLE_HOME_USER\} \ section \The raw Isabelle ML process\ subsection \Batch mode \label{sec:tool-process}\ text \ The @{tool_def process} tool runs the raw ML process in batch mode: @{verbatim [display] \Usage: isabelle process [OPTIONS] Options are: -T THEORY load theory -d DIR include session directory -e ML_EXPR evaluate ML expression on startup -f ML_FILE evaluate ML file on startup -l NAME logic session name (default ISABELLE_LOGIC="HOL") -m MODE add print mode for output -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) Run the raw Isabelle ML process in batch mode.\} \<^medskip> Options \<^verbatim>\-e\ and \<^verbatim>\-f\ allow to evaluate ML code, before the ML process is started. The source is either given literally or taken from a file. Multiple \<^verbatim>\-e\ and \<^verbatim>\-f\ options are evaluated in the given order. Errors lead to a premature exit of the ML process with return code 1. \<^medskip> Option \<^verbatim>\-T\ loads a specified theory file. This is a wrapper for \<^verbatim>\-e\ with a suitable \<^ML>\use_thy\ invocation. \<^medskip> Option \<^verbatim>\-l\ specifies the logic session name. Option \<^verbatim>\-d\ specifies additional directories for session roots, see also \secref{sec:tool-build}. \<^medskip> The \<^verbatim>\-m\ option adds identifiers of print modes to be made active for this session. For example, \<^verbatim>\-m ASCII\ prefers ASCII replacement syntax over mathematical Isabelle symbols. \<^medskip> Option \<^verbatim>\-o\ allows to override Isabelle system options for this process, see also \secref{sec:system-options}. \ subsubsection \Examples\ text \ The subsequent example retrieves the \<^verbatim>\Main\ theory value from the theory loader within ML: @{verbatim [display] \isabelle process -e 'Thy_Info.get_theory "Main"'\} Observe the delicate quoting rules for the GNU bash shell vs.\ ML. The Isabelle/ML and Scala libraries provide functions for that, but here we need to do it manually. \<^medskip> This is how to invoke a function body with proper return code and printing of errors, and without printing of a redundant \<^verbatim>\val it = (): unit\ result: @{verbatim [display] \isabelle process -e 'Command_Line.tool (fn () => writeln "OK")'\} @{verbatim [display] \isabelle process -e 'Command_Line.tool (fn () => error "Bad")'\} \ subsection \Interactive mode\ text \ The @{tool_def console} tool runs the raw ML process with interactive console and line editor: @{verbatim [display] \Usage: isabelle console [OPTIONS] Options are: -d DIR include session directory -i NAME include session in name-space of theories -l NAME logic session name (default ISABELLE_LOGIC) -m MODE add print mode for output -n no build of session image on startup -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -r bootstrap from raw Poly/ML Build a logic session image and run the raw Isabelle ML process in interactive mode, with line editor ISABELLE_LINE_EDITOR.\} \<^medskip> Option \<^verbatim>\-l\ specifies the logic session name. By default, its heap image is checked and built on demand, but the option \<^verbatim>\-n\ skips that. Option \<^verbatim>\-i\ includes additional sessions into the name-space of theories: multiple occurrences are possible. Option \<^verbatim>\-r\ indicates a bootstrap from the raw Poly/ML system, which is relevant for Isabelle/Pure development. \<^medskip> Options \<^verbatim>\-d\, \<^verbatim>\-m\, \<^verbatim>\-o\ have the same meaning as for @{tool process} (\secref{sec:tool-process}). \<^medskip> The Isabelle/ML process is run through the line editor that is specified via the settings variable @{setting ISABELLE_LINE_EDITOR} (e.g.\ @{executable_def rlwrap} for GNU readline); the fall-back is to use plain standard input/output. The user is connected to the raw ML toplevel loop: this is neither Isabelle/Isar nor Isabelle/ML within the usual formal context. The most relevant ML commands at this stage are \<^ML>\use\ (for ML files) and \<^ML>\use_thy\ (for theory files). \ section \The raw Isabelle Java process \label{sec:isabelle-java}\ text \ The @{executable_ref isabelle_java} executable allows to run a Java process within the name space of Java and Scala components that are bundled with Isabelle, but \<^emph>\without\ the Isabelle settings environment (\secref{sec:settings}). After such a JVM cold-start, the Isabelle environment can be accessed via \<^verbatim>\Isabelle_System.getenv\ as usual, but the underlying process environment remains clean. This is e.g.\ relevant when invoking other processes that should remain separate from the current Isabelle installation. \<^medskip> Note that under normal circumstances, Isabelle command-line tools are run \<^emph>\within\ the settings environment, as provided by the @{executable isabelle} wrapper (\secref{sec:isabelle-tool} and \secref{sec:tool-java}). \ subsubsection \Example\ text \ The subsequent example creates a raw Java process on the command-line and invokes the main Isabelle application entry point: @{verbatim [display] \isabelle_java -Djava.awt.headless=false isabelle.jedit.JEdit_Main\} \ section \YXML versus XML \label{sec:yxml-vs-xml}\ text \ Isabelle tools often use YXML, which is a simple and efficient syntax for untyped XML trees. The YXML format is defined as follows. \<^enum> The encoding is always UTF-8. \<^enum> Body text is represented verbatim (no escaping, no special treatment of white space, no named entities, no CDATA chunks, no comments). \<^enum> Markup elements are represented via ASCII control characters \\<^bold>X = 5\ and \\<^bold>Y = 6\ as follows: \begin{tabular}{ll} XML & YXML \\\hline \<^verbatim>\<\\name attribute\\<^verbatim>\=\\value \\\<^verbatim>\>\ & \\<^bold>X\<^bold>Yname\<^bold>Yattribute\\<^verbatim>\=\\value\\<^bold>X\ \\ \<^verbatim>\\name\\<^verbatim>\>\ & \\<^bold>X\<^bold>Y\<^bold>X\ \\ \end{tabular} There is no special case for empty body text, i.e.\ \<^verbatim>\\ is treated like \<^verbatim>\\. Also note that \\<^bold>X\ and \\<^bold>Y\ may never occur in well-formed XML documents. Parsing YXML is pretty straight-forward: split the text into chunks separated by \\<^bold>X\, then split each chunk into sub-chunks separated by \\<^bold>Y\. Markup chunks start with an empty sub-chunk, and a second empty sub-chunk indicates close of an element. Any other non-empty chunk consists of plain text. For example, see \<^file>\~~/src/Pure/PIDE/yxml.ML\ or \<^file>\~~/src/Pure/PIDE/yxml.scala\. YXML documents may be detected quickly by checking that the first two characters are \\<^bold>X\<^bold>Y\. \ end diff --git a/src/Doc/System/Scala.thy b/src/Doc/System/Scala.thy --- a/src/Doc/System/Scala.thy +++ b/src/Doc/System/Scala.thy @@ -1,493 +1,493 @@ (*:maxLineLen=78:*) theory Scala imports Base begin chapter \Isabelle/Scala systems programming \label{sec:scala}\ text \ Isabelle/ML and Isabelle/Scala are the two main implementation languages of the Isabelle environment: \<^item> Isabelle/ML is for \<^emph>\mathematics\, to develop tools within the context of symbolic logic, e.g.\ for constructing proofs or defining domain-specific formal languages. See the \<^emph>\Isabelle/Isar implementation - manual\ @{cite "isabelle-implementation"} for more details. + manual\ \<^cite>\"isabelle-implementation"\ for more details. \<^item> Isabelle/Scala is for \<^emph>\physics\, to connect with the world of systems and services, including editors and IDE frameworks. There are various ways to access Isabelle/Scala modules and operations: \<^item> Isabelle command-line tools (\secref{sec:scala-tools}) run in a separate Java process. \<^item> Isabelle/ML antiquotations access Isabelle/Scala functions (\secref{sec:scala-functions}) via the PIDE protocol: execution happens within the running Java process underlying Isabelle/Scala. - \<^item> The \<^verbatim>\Console/Scala\ plugin of Isabelle/jEdit @{cite "isabelle-jedit"} + \<^item> The \<^verbatim>\Console/Scala\ plugin of Isabelle/jEdit \<^cite>\"isabelle-jedit"\ operates on the running Java application, using the Scala read-eval-print-loop (REPL). The main Isabelle/Scala/jEdit functionality is provided by \<^file>\$ISABELLE_HOME/lib/classes/isabelle.jar\. Further underlying Scala and Java libraries are bundled with Isabelle, e.g.\ to access SQLite or PostgreSQL via JDBC. Add-on Isabelle components may augment the system environment by providing suitable configuration in \<^path>\etc/settings\ (GNU bash script). The shell function \<^bash_function>\classpath\ helps to write \<^path>\etc/settings\ in a portable manner: it refers to library \<^verbatim>\jar\ files in standard POSIX path notation. On Windows, this is converted to native platform format, before invoking Java (\secref{sec:scala-tools}). \<^medskip> There is also an implicit build process for Isabelle/Scala/Java modules, based on \<^path>\etc/build.props\ within the component directory (see also \secref{sec:scala-build}). \ section \Command-line tools \label{sec:scala-tools}\ subsection \Java Runtime Environment \label{sec:tool-java}\ text \ The @{tool_def java} tool is a direct wrapper for the Java Runtime Environment, within the regular Isabelle settings environment (\secref{sec:settings}) and Isabelle classpath. The command line arguments are that of the bundled Java distribution: see option \<^verbatim>\-help\ in particular. The \<^verbatim>\java\ executable is taken from @{setting ISABELLE_JDK_HOME}, according to the standard directory layout for regular distributions of OpenJDK. The shell function \<^bash_function>\isabelle_jdk\ allows shell scripts to invoke other Java tools robustly (e.g.\ \<^verbatim>\isabelle_jdk jar\), without depending on accidental operating system installations. \ subsection \Scala toplevel \label{sec:tool-scala}\ text \ The @{tool_def scala} tool is a direct wrapper for the Scala toplevel, similar to @{tool java} above. The command line arguments are that of the bundled Scala distribution: see option \<^verbatim>\-help\ in particular. This allows to interact with Isabelle/Scala interactively. \ subsubsection \Example\ text \ Explore the Isabelle system environment in Scala: @{verbatim [display, indent = 2] \$ isabelle scala\} @{scala [display, indent = 2] \import isabelle._ val isabelle_home = Isabelle_System.getenv("ISABELLE_HOME") val options = Options.init() options.bool("browser_info") options.string("document")\} \ subsection \Scala script wrapper\ text \ The executable @{executable "$ISABELLE_HOME/bin/isabelle_scala_script"} allows to run Isabelle/Scala source files stand-alone programs, by using a suitable ``hash-bang'' line and executable file permissions. For example: @{verbatim [display, indent = 2] \#!/usr/bin/env isabelle_scala_script\} @{scala [display, indent = 2] \val options = isabelle.Options.init() Console.println("browser_info = " + options.bool("browser_info")) Console.println("document = " + options.string("document"))\} This assumes that the executable may be found via the @{setting PATH} from the process environment: this is the case when Isabelle settings are active, e.g.\ in the context of the main Isabelle tool wrapper \secref{sec:isabelle-tool}. Alternatively, the full \<^file>\$ISABELLE_HOME/bin/isabelle_scala_script\ may be specified in expanded form. \ subsection \Scala compiler \label{sec:tool-scalac}\ text \ The @{tool_def scalac} tool is a direct wrapper for the Scala compiler; see also @{tool scala} above. The command line arguments are that of the bundled Scala distribution. This provides a low-level mechanism to compile further Scala modules, depending on existing Isabelle/Scala functionality; the resulting \<^verbatim>\class\ or \<^verbatim>\jar\ files can be added to the Java classpath using the shell function \<^bash_function>\classpath\. A more convenient high-level approach works via \<^path>\etc/build.props\ (see \secref{sec:scala-build}). \ section \Isabelle/Scala/Java modules \label{sec:scala-build}\ subsection \Component configuration via \<^path>\etc/build.props\\ text \ Isabelle components may augment the Isabelle/Scala/Java environment declaratively via properties given in \<^path>\etc/build.props\ (within the component directory). This specifies an output \<^verbatim>\jar\ \<^emph>\module\, based on Scala or Java \<^emph>\sources\, and arbitrary \<^emph>\resources\. Moreover, a module can specify \<^emph>\services\ that are subclasses of \<^scala_type>\isabelle.Isabelle_System.Service\; these have a particular meaning to Isabelle/Scala tools. Before running a Scala or Java process, the Isabelle system implicitly ensures that all provided modules are compiled and packaged (as jars). It is also possible to invoke @{tool scala_build} explicitly, with extra options. \<^medskip> The syntax of \<^path>\etc/build.props\ follows a regular Java properties file\<^footnote>\\<^url>\https://docs.oracle.com/en/java/javase/17/docs/api/java.base/java/util/Properties.html#load(java.io.Reader)\\, but the encoding is \<^verbatim>\UTF-8\, instead of historic \<^verbatim>\ISO 8859-1\ from the API documentation. The subsequent properties are relevant for the Scala/Java build process. Most properties are optional: the default is an empty string (or list). File names are relative to the main component directory and may refer to Isabelle settings variables (e.g. \<^verbatim>\$ISABELLE_HOME\). \<^item> \<^verbatim>\title\ (required) is a human-readable description of the module, used in printed messages. \<^item> \<^verbatim>\module\ specifies a \<^verbatim>\jar\ file name for the output module, as result of the specified sources (and resources). If this is absent (or \<^verbatim>\no_build\ is set, as described below), there is no implicit build process. The contributing sources might be given nonetheless, notably for @{tool scala_project} (\secref{sec:tool-scala-project}), which includes Scala/Java sources of components, while suppressing \<^verbatim>\jar\ modules (to avoid duplication of program content). \<^item> \<^verbatim>\no_build\ is a Boolean property, with default \<^verbatim>\false\. If set to \<^verbatim>\true\, the implicit build process for the given \<^verbatim>\module\ is \<^emph>\omitted\ --- it is assumed to be provided by other means. \<^item> \<^verbatim>\scalac_options\ and \<^verbatim>\javac_options\ augment the default settings @{setting_ref ISABELLE_SCALAC_OPTIONS} and @{setting_ref ISABELLE_JAVAC_OPTIONS} for this component; option syntax follows the regular command-line tools \<^verbatim>\scalac\ and \<^verbatim>\javac\, respectively. \<^item> \<^verbatim>\main\ specifies the main entry point for the \<^verbatim>\jar\ module. This is only relevant for direct invocation like ``\<^verbatim>\java -jar test.jar\''. \<^item> \<^verbatim>\requirements\ is a list of \<^verbatim>\jar\ modules that are needed in the compilation process, but not provided by the regular classpath (notably @{setting ISABELLE_CLASSPATH}). A \<^emph>\normal entry\ refers to a single \<^verbatim>\jar\ file name, possibly with settings variables as usual. E.g. \<^file>\$ISABELLE_SCALA_JAR\ for the main \<^file>\$ISABELLE_HOME/lib/classes/isabelle.jar\ (especially relevant for add-on modules). A \<^emph>\special entry\ is of the form \<^verbatim>\env:\\variable\ and refers to a settings variable from the Isabelle environment: its value may consist of multiple \<^verbatim>\jar\ entries (separated by colons). Environment variables are not expanded recursively. \<^item> \<^verbatim>\resources\ is a list of files that should be included in the resulting \<^verbatim>\jar\ file. Each item consists of a pair separated by colon: \source\\<^verbatim>\:\\target\ means to copy an existing source file (relative to the component directory) to the given target file or directory (relative to the \<^verbatim>\jar\ name space). A \file\ specification without colon abbreviates \file\\<^verbatim>\:\\file\, i.e. the file is copied while retaining its relative path name. \<^item> \<^verbatim>\sources\ is a list of \<^verbatim>\.scala\ or \<^verbatim>\.java\ files that contribute to the specified module. It is possible to use both languages simultaneously: the Scala and Java compiler will be invoked consecutively to make this work. \<^item> \<^verbatim>\services\ is a list of class names to be registered as Isabelle service providers (subclasses of \<^scala_type>\isabelle.Isabelle_System.Service\). Internal class names of the underlying JVM need to be given: e.g. see method @{scala_method (in java.lang.Object) getClass}. Particular services require particular subclasses: instances are filtered according to their dynamic type. For example, class \<^scala_type>\isabelle.Isabelle_Scala_Tools\ collects Scala command-line tools, and class \<^scala_type>\isabelle.Scala.Functions\ collects Scala functions (\secref{sec:scala-functions}). \ subsection \Explicit Isabelle/Scala/Java build \label{sec:tool-scala-build}\ text \ The @{tool_def scala_build} tool explicitly invokes the build process for all registered components. @{verbatim [display] \Usage: isabelle scala_build [OPTIONS] Options are: -f force fresh build -q quiet mode: suppress stdout/stderr Build Isabelle/Scala/Java modules of all registered components (if required). \} For each registered Isabelle component that provides \<^path>\etc/build.props\, the specified output \<^verbatim>\module\ is checked against the corresponding input \<^verbatim>\requirements\, \<^verbatim>\resources\, \<^verbatim>\sources\. If required, there is an automatic build using \<^verbatim>\scalac\ or \<^verbatim>\javac\ (or both). The identity of input files is recorded within the output \<^verbatim>\jar\, using SHA1 digests in \<^verbatim>\META-INF/isabelle/shasum\. \<^medskip> Option \<^verbatim>\-f\ forces a fresh build, regardless of the up-to-date status of input files vs. the output module. \<^medskip> Option \<^verbatim>\-q\ suppresses all output on stdout/stderr produced by the Scala or Java compiler. \<^medskip> Explicit invocation of @{tool scala_build} mainly serves testing or applications with special options: the Isabelle system normally does an automatic the build on demand. \ subsection \Project setup for common Scala IDEs \label{sec:tool-scala-project}\ text \ The @{tool_def scala_project} tool creates a project configuration for all Isabelle/Java/Scala modules specified in components via \<^path>\etc/build.props\, together with additional source files given on the command-line: @{verbatim [display] \Usage: isabelle scala_project [OPTIONS] [MORE_SOURCES ...] Options are: -D DIR project directory (default: "$ISABELLE_HOME_USER/scala_project") -G use Gradle as build tool -L make symlinks to original source files -M use Maven as build tool -f force update of existing directory Setup project for Isabelle/Scala/jEdit --- to support common IDEs such as IntelliJ IDEA. Either option -G or -M is mandatory to specify the build tool.\} The generated configuration is for Gradle\<^footnote>\\<^url>\https://gradle.org\\ or Maven\<^footnote>\\<^url>\https://maven.apache.org\\, but the main purpose is to import it into common IDEs like IntelliJ IDEA\<^footnote>\\<^url>\https://www.jetbrains.com/idea\\. This allows to explore the sources with static analysis and other hints in real-time. The generated files refer to physical file-system locations, using the path notation of the underlying OS platform. Thus the project needs to be recreated whenever the Isabelle installation is changed or moved. \<^medskip> Option \<^verbatim>\-G\ selects Gradle and \<^verbatim>\-M\ selects Maven as Java/Scala build tool: either one needs to be specified explicitly. These tools have a tendency to break down unexpectedly, so supporting both increases the chances that the generated IDE project works properly. \<^medskip> Option \<^verbatim>\-L\ produces \<^emph>\symlinks\ to the original files: this allows to develop Isabelle/Scala/jEdit modules within an external IDE. The default is to \<^emph>\copy\ source files, so editing them within the IDE has no permanent effect on the originals. \<^medskip> Option \<^verbatim>\-D\ specifies an explicit project directory, instead of the default \<^path>\$ISABELLE_HOME_USER/scala_project\. Option \<^verbatim>\-f\ forces an existing project directory to be \<^emph>\purged\ --- after some sanity checks that it has been generated by @{tool "scala_project"} before. \ subsubsection \Examples\ text \ Create a project directory and for editing the original sources: @{verbatim [display] \isabelle scala_project -f -L\} On Windows, this usually requires Administrator rights, in order to create native symlinks. \ section \Registered Isabelle/Scala functions \label{sec:scala-functions}\ subsection \Defining functions in Isabelle/Scala\ text \ The service class \<^scala_type>\isabelle.Scala.Functions\ collects Scala functions of type \<^scala_type>\isabelle.Scala.Fun\: by registering instances via \<^verbatim>\services\ in \<^path>\etc/build.props\ (\secref{sec:scala-build}), it becomes possible to invoke Isabelle/Scala from Isabelle/ML (see below). An example is the predefined collection of \<^scala_type>\isabelle.Scala.Functions\ in \<^file>\$ISABELLE_HOME/etc/build.props\. The overall list of registered functions is accessible in Isabelle/Scala as \<^scala_object>\isabelle.Scala.functions\. The general class \<^scala_type>\isabelle.Scala.Fun\ expects a multi-argument / multi-result function \<^scala_type>\List[isabelle.Bytes] => List[isabelle.Bytes]\; more common are instances of \<^scala_type>\isabelle.Scala.Fun_Strings\ for type \<^scala_type>\List[String] => List[String]\, or \<^scala_type>\isabelle.Scala.Fun_String\ for type \<^scala_type>\String => String\. \ subsection \Invoking functions in Isabelle/ML\ text \ Isabelle/PIDE provides a protocol to invoke registered Scala functions in ML: this works both within the Prover IDE and in batch builds. The subsequent ML antiquotations refer to Scala functions in a formally-checked manner. \begin{matharray}{rcl} @{ML_antiquotation_def "scala_function"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "scala"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ (@{ML_antiquotation scala_function} | @{ML_antiquotation scala}) @{syntax embedded} \ \<^descr> \@{scala_function name}\ inlines the checked function name as ML string literal. \<^descr> \@{scala name}\ and \@{scala_thread name}\ invoke the checked function via the PIDE protocol. In Isabelle/ML this appears as a function of type \<^ML_type>\string list -> string list\ or \<^ML_type>\string -> string\, depending on the definition in Isabelle/Scala. Evaluation is subject to interrupts within the ML runtime environment as usual. A \<^scala>\null\ result in Scala raises an exception \<^ML>\Scala.Null\ in ML. The execution of \@{scala}\ works via a Scala future on a bounded thread farm, while \@{scala_thread}\ always forks a separate Java/VM thread. The standard approach of representing datatypes via strings works via XML in YXML transfer syntax. See Isabelle/ML operations and modules @{ML YXML.string_of_body}, @{ML YXML.parse_body}, @{ML_structure XML.Encode}, @{ML_structure XML.Decode}; similarly for Isabelle/Scala. Isabelle symbols may have to be recoded via Scala operations \<^scala_method>\isabelle.Symbol.decode\ and \<^scala_method>\isabelle.Symbol.encode\. \ subsubsection \Examples\ text \ Invoke the predefined Scala function \<^scala_function>\echo\: \ ML \ val s = "test"; val s' = \<^scala>\echo\ s; \<^assert> (s = s') \ text \ Let the Scala compiler process some toplevel declarations, producing a list of errors: \ ML \ val source = "class A(a: Int, b: Boolean)" val errors = \<^scala>\scala_toplevel\ source |> YXML.parse_body |> let open XML.Decode in list string end; \<^assert> (null errors)\ text \ The above is merely for demonstration. See \<^ML>\Scala_Compiler.toplevel\ for a more convenient version with builtin decoding and treatment of errors. \ section \Documenting Isabelle/Scala entities\ text \ The subsequent document antiquotations help to document Isabelle/Scala entities, with formal checking of names against the Isabelle classpath. \begin{matharray}{rcl} @{antiquotation_def "scala"} & : & \antiquotation\ \\ @{antiquotation_def "scala_object"} & : & \antiquotation\ \\ @{antiquotation_def "scala_type"} & : & \antiquotation\ \\ @{antiquotation_def "scala_method"} & : & \antiquotation\ \\ \end{matharray} \<^rail>\ (@@{antiquotation scala} | @@{antiquotation scala_object}) @{syntax embedded} ; @@{antiquotation scala_type} @{syntax embedded} types ; @@{antiquotation scala_method} class @{syntax embedded} types args ; class: ('(' @'in' @{syntax name} types ')')? ; types: ('[' (@{syntax name} ',' +) ']')? ; args: ('(' (nat | (('_' | @{syntax name}) + ',')) ')')? \ \<^descr> \@{scala s}\ is similar to \@{verbatim s}\, but the given source text is checked by the Scala compiler as toplevel declaration (without evaluation). This allows to write Isabelle/Scala examples that are statically checked. \<^descr> \@{scala_object x}\ checks the given Scala object name (simple value or ground module) and prints the result verbatim. \<^descr> \@{scala_type T[A]}\ checks the given Scala type name (with optional type parameters) and prints the result verbatim. \<^descr> \@{scala_method (in c[A]) m[B](n)}\ checks the given Scala method \m\ in the context of class \c\. The method argument slots are either specified by a number \n\ or by a list of (optional) argument types; this may refer to type variables specified for the class or method: \A\ or \B\ above. Everything except for the method name \m\ is optional. The absence of the class context means that this is a static method. The absence of arguments with types means that the method can be determined uniquely as \<^verbatim>\(\\m\\<^verbatim>\ _)\ in Scala (no overloading). \ subsubsection \Examples\ text \ Miscellaneous Isabelle/Scala entities: \<^item> object: \<^scala_object>\isabelle.Isabelle_Process\ \<^item> type without parameter: @{scala_type isabelle.Console_Progress} \<^item> type with parameter: @{scala_type List[A]} \<^item> static method: \<^scala_method>\isabelle.Isabelle_System.bash\ \<^item> class and method with type parameters: @{scala_method (in List[A]) map[B]("A => B")} \<^item> overloaded method with argument type: @{scala_method (in Int) "+" (Int)} \ end diff --git a/src/Doc/System/Server.thy b/src/Doc/System/Server.thy --- a/src/Doc/System/Server.thy +++ b/src/Doc/System/Server.thy @@ -1,1108 +1,1107 @@ (*:maxLineLen=78:*) theory Server imports Base begin chapter \The Isabelle server\ text \ An Isabelle session requires at least two processes, which are both rather heavy: Isabelle/Scala for the system infrastructure and Isabelle/ML for the logic session (e.g.\ HOL). In principle, these processes can be invoked directly on the command-line, e.g.\ via @{tool java}, @{tool scala}, @{tool process}, @{tool console}, but this approach is inadequate for reactive applications that require quick responses from the prover. In contrast, the Isabelle server exposes Isabelle/Scala as a ``terminate-stay-resident'' application that manages multiple logic \<^emph>\sessions\ and concurrent tasks to use \<^emph>\theories\. This is analogous to \<^ML>\Thy_Info.use_theories\ in Isabelle/ML, with proper support for concurrent invocations. The client/server arrangement via TCP sockets also opens possibilities for remote Isabelle services that are accessed by local applications, e.g.\ via an SSH tunnel. \ section \Command-line tools\ subsection \Server \label{sec:tool-server}\ text \ The @{tool_def server} tool manages resident server processes: @{verbatim [display] \Usage: isabelle server [OPTIONS] Options are: -L FILE logging on FILE -c console interaction with specified server -l list servers (alternative operation) -n NAME explicit server name (default: isabelle) -p PORT explicit server port -s assume existing server, no implicit startup -x exit specified server (alternative operation) Manage resident Isabelle servers.\} The main operation of \<^verbatim>\isabelle server\ is to ensure that a named server is running, either by finding an already running process (according to the central database file \<^path>\$ISABELLE_HOME_USER/servers.db\) or by becoming itself a new server that accepts connections on a particular TCP socket. The server name and its address are printed as initial output line. If another server instance is already running, the current \<^verbatim>\isabelle server\ process will terminate; otherwise, it keeps running as a new server process until an explicit \<^verbatim>\shutdown\ command is received. Further details of the server socket protocol are explained in \secref{sec:server-protocol}. Other server management operations are invoked via options \<^verbatim>\-l\ and \<^verbatim>\-x\ (see below). \<^medskip> Option \<^verbatim>\-n\ specifies an alternative server name: at most one process for each name may run, but each server instance supports multiple connections and logic sessions. \<^medskip> Option \<^verbatim>\-p\ specifies an explicit TCP port for the server socket (which is always on \<^verbatim>\localhost\): the default is to let the operating system assign a free port number. \<^medskip> Option \<^verbatim>\-s\ strictly assumes that the specified server process is already running, skipping the optional server startup phase. \<^medskip> Option \<^verbatim>\-c\ connects the console in/out channels after the initial check for a suitable server process. Also note that the @{tool client} tool (\secref{sec:tool-client}) provides a command-line editor to interact with the server. \<^medskip> Option \<^verbatim>\-L\ specifies a log file for exceptional output of internal server and session operations. \<^medskip> Operation \<^verbatim>\-l\ lists all active server processes with their connection details. \<^medskip> Operation \<^verbatim>\-x\ exits the specified server process by sending it a \<^verbatim>\shutdown\ command. \ subsection \Client \label{sec:tool-client}\ text \ The @{tool_def client} tool provides console interaction for Isabelle servers: @{verbatim [display] \Usage: isabelle client [OPTIONS] Options are: -n NAME explicit server name -p PORT explicit server port Console interaction for Isabelle server (with line-editor).\} This is a wrapper to \<^verbatim>\isabelle server -s -c\ for interactive experimentation, which uses @{setting ISABELLE_LINE_EDITOR} if available. The server name is sufficient for identification, as the client can determine the connection details from the local database of active servers. \<^medskip> Option \<^verbatim>\-n\ specifies an explicit server name as in @{tool server}. \<^medskip> Option \<^verbatim>\-p\ specifies an explicit server port as in @{tool server}. \ subsection \Examples\ text \ Ensure that a particular server instance is running in the background: @{verbatim [display] \isabelle server -n test &\} The first line of output presents the connection details:\<^footnote>\This information may be used in other TCP clients, without access to Isabelle/Scala and the underlying database of running servers.\ @{verbatim [display] \server "test" = 127.0.0.1:4711 (password "XYZ")\} \<^medskip> List available server processes: @{verbatim [display] \isabelle server -l\} \<^medskip> Connect the command-line client to the above test server: @{verbatim [display] \isabelle client -n test\} Interaction now works on a line-by-line basis, with commands like \<^verbatim>\help\ or \<^verbatim>\echo\. For example, some JSON values may be echoed like this: @{verbatim [display] \echo 42 echo [1, 2, 3] echo {"a": "text", "b": true, "c": 42}\} Closing the connection (via CTRL-D) leaves the server running: it is possible to reconnect again, and have multiple connections at the same time. \<^medskip> Exit the named server on the command-line: @{verbatim [display] \isabelle server -n test -x\} \ section \Protocol messages \label{sec:server-protocol}\ text \ The Isabelle server listens on a regular TCP socket, using a line-oriented protocol of structured messages. Input \<^emph>\commands\ and output \<^emph>\results\ (via \<^verbatim>\OK\ or \<^verbatim>\ERROR\) are strictly alternating on the toplevel, but commands may also return a \<^emph>\task\ identifier to indicate an ongoing asynchronous process that is joined later (via \<^verbatim>\FINISHED\ or \<^verbatim>\FAILED\). Asynchronous \<^verbatim>\NOTE\ messages may occur at any time: they are independent of the main command-result protocol. For example, the synchronous \<^verbatim>\echo\ command immediately returns its argument as \<^verbatim>\OK\ result. In contrast, the asynchronous \<^verbatim>\session_build\ command returns \<^verbatim>\OK {"task":\\id\\<^verbatim>\}\ and continues in the background. It will eventually produce \<^verbatim>\FINISHED {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ or \<^verbatim>\FAILED {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ with the final result. Intermediately, it may emit asynchronous messages of the form \<^verbatim>\NOTE {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ to inform about its progress. Due to the explicit task identifier, the client can show these messages in the proper context, e.g.\ a GUI window for this particular session build job. \medskip Subsequently, the protocol message formats are described in further detail. \ subsection \Byte messages \label{sec:byte-messages}\ text \ The client-server connection is a raw byte-channel for bidirectional communication, but the Isabelle server always works with messages of a particular length. Messages are written as a single chunk that is flushed immediately. Message boundaries are determined as follows: \<^item> A \<^emph>\short message\ consists of a single line: it is a sequence of arbitrary bytes excluding CR (13) and LF (10), and terminated by CR-LF or just LF. \<^item> A \<^emph>\long message\ starts with a single line consisting of decimal digits: these are interpreted as length of the subsequent block of arbitrary bytes. A final line-terminator (as above) may be included here, but is not required. Messages in JSON format (see below) always fit on a single line, due to escaping of newline characters within string literals. This is convenient for interactive experimentation, but it can impact performance for very long messages. If the message byte-length is given on the preceding line, the server can read the message more efficiently as a single block. \ subsection \Text messages\ text \ Messages are read and written as byte streams (with byte lengths), but the content is always interpreted as plain text in terms of the UTF-8 encoding.\<^footnote>\See also the ``UTF-8 Everywhere Manifesto'' \<^url>\https://utf8everywhere.org\.\ Note that line-endings and other formatting characters are invariant wrt. UTF-8 representation of text: thus implementations are free to determine the overall message structure before or after applying the text encoding. \ subsection \Input and output messages \label{sec:input-output-messages}\ text \ The uniform format for server input and output messages is \name argument\, such that: \<^item> \name\ is the longest prefix consisting of ASCII letters, digits, ``\<^verbatim>\_\'', ``\<^verbatim>\.\'', \<^item> the separator between \name\ and \argument\ is the longest possible sequence of ASCII blanks (it could be empty, e.g.\ when the argument starts with a quote or bracket), \<^item> \argument\ is the rest of the message without line terminator. \<^medskip> Input messages are sent from the client to the server. Here the \name\ specifies a \<^emph>\server command\: the list of known commands may be retrieved via the \<^verbatim>\help\ command. \<^medskip> Output messages are sent from the server to the client. Here the \name\ specifies the \<^emph>\server reply\, which always has a specific meaning as follows: \<^item> synchronous results: \<^verbatim>\OK\ or \<^verbatim>\ERROR\ \<^item> asynchronous results: \<^verbatim>\FINISHED\ or \<^verbatim>\FAILED\ \<^item> intermediate notifications: \<^verbatim>\NOTE\ \<^medskip> The \argument\ format is uniform for both input and output messages: \<^item> empty argument (Scala type \<^verbatim>\Unit\) \<^item> XML element in YXML notation (Scala type \<^verbatim>\XML.Elem\) \<^item> JSON value (Scala type \<^verbatim>\JSON.T\) JSON values may consist of objects (records), arrays (lists), strings, numbers, bools, or null.\<^footnote>\See also the official specification \<^url>\https://www.json.org\ and unofficial explorations ``Parsing JSON is a Minefield'' \<^url>\http://seriot.ch/parsing_json.php\.\ Since JSON requires explicit quotes and backslash-escapes to represent arbitrary text, the YXML notation for XML trees (\secref{sec:yxml-vs-xml}) works better for large messages with a lot of PIDE markup. Nonetheless, the most common commands use JSON by default: big chunks of text (theory sources etc.) are taken from the underlying file-system and results are pre-formatted for plain-text output, without PIDE markup information. This is a concession to simplicity: the server imitates the appearance of command-line tools on top of the Isabelle/PIDE infrastructure. \ subsection \Initial password exchange\ text \ Whenever a new client opens the server socket, the initial message needs to be its unique password as a single line, without length indication (i.e.\ a ``short message'' in the sense of \secref{sec:byte-messages}). The server replies either with \<^verbatim>\OK\ (and some information about the Isabelle version) or by silent disconnection of what is considered an illegal connection attempt. Note that @{tool client} already presents the correct password internally. Server passwords are created as Universally Unique Identifier (UUID) in Isabelle/Scala and stored in a per-user database, with restricted file-system access only for the current user. The Isabelle/Scala server implementation is careful to expose the password only on private output channels, and not on a process command-line (which is accessible to other users, e.g.\ via the \<^verbatim>\ps\ command). \ subsection \Synchronous commands\ text \ A \<^emph>\synchronous command\ corresponds to regular function application in Isabelle/Scala, with single argument and result (regular or error). Both the argument and the result may consist of type \<^verbatim>\Unit\, \<^verbatim>\XML.Elem\, \<^verbatim>\JSON.T\. An error result typically consists of a JSON object with error message and potentially further result fields (this resembles exceptions in Scala). These are the protocol exchanges for both cases of command execution: \begin{center} \begin{tabular}{rl} \<^bold>\input:\ & \command argument\ \\ (a) regular \<^bold>\output:\ & \<^verbatim>\OK\ \result\ \\ (b) error \<^bold>\output:\ & \<^verbatim>\ERROR\ \result\ \\ \end{tabular} \end{center} \ subsection \Asynchronous commands\ text \ An \<^emph>\asynchronous command\ corresponds to an ongoing process that finishes or fails eventually, while emitting arbitrary notifications in between. Formally, it starts as synchronous command with immediate result \<^verbatim>\OK\ giving the \<^verbatim>\task\ identifier, or an immediate \<^verbatim>\ERROR\ that indicates bad command syntax. For a running task, the termination is indicated later by \<^verbatim>\FINISHED\ or \<^verbatim>\FAILED\, together with its ultimate result value. These are the protocol exchanges for various cases of command task execution: \begin{center} \begin{tabular}{rl} \<^bold>\input:\ & \command argument\ \\ immediate \<^bold>\output:\ & \<^verbatim>\OK {"task":\\id\\<^verbatim>\}\ \\ intermediate \<^bold>\output:\ & \<^verbatim>\NOTE {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ \\ (a) regular \<^bold>\output:\ & \<^verbatim>\FINISHED {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ \\ (b) error \<^bold>\output:\ & \<^verbatim>\FAILED {"task":\\id\\<^verbatim>\,\\\\\<^verbatim>\}\ \\[3ex] \<^bold>\input:\ & \command argument\ \\ immediate \<^bold>\output:\ & \<^verbatim>\ERROR\~\\\ \\ \end{tabular} \end{center} All asynchronous messages are decorated with the task identifier that was revealed in the immediate (synchronous) result. Thus the client can invoke further asynchronous commands and still dispatch the resulting stream of asynchronous messages properly. The synchronous command \<^verbatim>\cancel {"task":\~\id\\<^verbatim>\}\ tells the specified task to terminate prematurely: usually causing a \<^verbatim>\FAILED\ result, but this is not guaranteed: the cancel event may come too late or the running process may just ignore it. \ section \Types for JSON values \label{sec:json-types}\ text \ In order to specify concrete JSON types for command arguments and result messages, the following type definition language shall be used: \<^rail>\ @{syntax type_def}: @'type' @{syntax name} '=' @{syntax type} ; @{syntax type}: @{syntax name} | @{syntax value} | 'any' | 'null' | 'bool' | 'int' | 'long' | 'double' | 'string' | '[' @{syntax type} ']' | '{' (@{syntax field_type} ',' *) '}' | @{syntax type} '\' @{syntax type} | @{syntax type} '|' @{syntax type} | '(' @{syntax type} ')' ; @{syntax field_type}: @{syntax name} ('?'?) ':' @{syntax type} \ This is a simplified variation of TypeScript interfaces.\<^footnote>\\<^url>\https://www.typescriptlang.org/docs/handbook/interfaces.html\\ The meaning of these types is specified wrt. the Isabelle/Scala implementation as follows. \<^item> A \name\ refers to a type defined elsewhere. The environment of type definitions is given informally: put into proper foundational order, it needs to specify a strongly normalizing system of syntactic abbreviations; type definitions may not be recursive. \<^item> A \value\ in JSON notation represents the singleton type of the given item. For example, the string \<^verbatim>\"error"\ can be used as type for a slot that is guaranteed to contain that constant. \<^item> Type \any\ is the super type of all other types: it is an untyped slot in the specification and corresponds to \<^verbatim>\Any\ or \<^verbatim>\JSON.T\ in Isabelle/Scala. \<^item> Type \null\ is the type of the improper value \null\; it corresponds to type \<^verbatim>\Null\ in Scala and is normally not used in Isabelle/Scala.\<^footnote>\See also ``Null References: The Billion Dollar Mistake'' by Tony Hoare \<^url>\https://www.infoq.com/presentations/Null-References-The-Billion-Dollar-Mistake-Tony-Hoare\.\ \<^item> Type \bool\ is the type of the truth values \<^verbatim>\true\ and \<^verbatim>\false\; it corresponds to \<^verbatim>\Boolean\ in Scala. \<^item> Types \int\, \long\, \double\ are specific versions of the generic \number\ type, corresponding to \<^verbatim>\Int\, \<^verbatim>\Long\, \<^verbatim>\Double\ in Scala, but \<^verbatim>\Long\ is limited to 53 bit precision.\<^footnote>\Implementations of JSON typically standardize \number\ to \<^verbatim>\Double\, which can absorb \<^verbatim>\Int\ faithfully, but not all of \<^verbatim>\Long\.\ \<^item> Type \string\ represents Unicode text; it corresponds to type \<^verbatim>\String\ in Scala. \<^item> Type \[t]\ is the array (or list) type over \t\; it corresponds to \<^verbatim>\List[t]\ in Scala. The list type is co-variant as usual (i.e.\ monotonic wrt. the subtype relation). \<^item> Object types describe the possible content of JSON records, with field names and types. A question mark after a field name means that it is optional. In Scala this could refer to an explicit type \<^verbatim>\Option[t]\, e.g.\ \{a: int, b?: string}\ corresponding to a Scala case class with arguments \<^verbatim>\a: Int\, \<^verbatim>\b: Option[String]\. Alternatively, optional fields can have a default value. If nothing else is specified, a standard ``empty value'' is used for each type, i.e.\ \<^verbatim>\0\ for the number types, \<^verbatim>\false\ for \bool\, or the empty string, array, object etc. Object types are \<^emph>\permissive\ in the sense that only the specified field names need to conform to the given types, but unspecified fields may be present as well. \<^item> The type expression \t\<^sub>1 \ t\<^sub>2\ only works for two object types with disjoint field names: it is the concatenation of the respective @{syntax field_type} specifications taken together. For example: \{task: string} \ {ok: bool}\ is the equivalent to \{task: string, ok: bool}\. \<^item> The type expression \t\<^sub>1 | t\<^sub>2\ is the disjoint union of two types, either one of the two cases may occur. \<^item> Parentheses \(t)\ merely group type expressions syntactically. These types correspond to JSON values in an obvious manner, which is not further described here. For example, the JSON array \<^verbatim>\[1, 2, 3]\ conforms to types \[int]\, \[long]\, \[double]\, \[any]\, \any\. Note that JSON objects require field names to be quoted, but the type language omits quotes for clarity. Thus the object \<^verbatim>\{"a": 42, "b": "xyz"}\ conforms to the type \{a: int, b: string}\, for example. \<^medskip> The absence of an argument or result is represented by the Scala type \<^verbatim>\Unit\: it is written as empty text in the message \argument\ (\secref{sec:input-output-messages}). This is not part of the JSON language. Server replies have name tags like \<^verbatim>\OK\, \<^verbatim>\ERROR\: these are used literally together with type specifications to indicate the particular name with the type of its argument, e.g.\ \<^verbatim>\OK\~\[string]\ for a regular result that is a list (JSON array) of strings. \<^bigskip> Here are some common type definitions, for use in particular specifications of command arguments and results. \<^item> \<^bold>\type\~\position = {line?: int, offset?: int, end_offset?: int, file?: string, id?: long}\ describes a source position within Isabelle text. Only the \line\ and \file\ fields make immediate sense to external programs. Detailed \offset\ and \end_offset\ positions are counted according to - Isabelle symbols, see \<^ML_type>\Symbol.symbol\ in Isabelle/ML @{cite - "isabelle-implementation"}. The position \id\ belongs to the representation + Isabelle symbols, see \<^ML_type>\Symbol.symbol\ in Isabelle/ML \<^cite>\"isabelle-implementation"\. The position \id\ belongs to the representation of command transactions in the Isabelle/PIDE protocol: it normally does not occur in externalized positions. \<^item> \<^bold>\type\~\message = {kind: string, message: string, pos?: position}\ where the \kind\ provides some hint about the role and importance of the message. The main message kinds are \<^verbatim>\writeln\ (for regular output), \<^verbatim>\warning\, \<^verbatim>\error\. \<^item> \<^bold>\type\~\error_message = {kind:\~\<^verbatim>\"error"\\, message: string}\ refers to error messages in particular. These occur routinely with \<^verbatim>\ERROR\ or \<^verbatim>\FAILED\ replies, but also as initial command syntax errors (which are omitted in the command specifications below). \<^item> \<^bold>\type\~\theory_progress = {kind:\~\<^verbatim>\"writeln"\\, message: string, theory: string, session: string, percentage?: int}\ reports formal progress in loading theories (e.g.\ when building a session image). Apart from a regular output message, it also reveals the formal theory name (e.g.\ \<^verbatim>\"HOL.Nat"\) and session name (e.g.\ \<^verbatim>\"HOL"\). Note that some rare theory names lack a proper session prefix, e.g.\ theory \<^verbatim>\"Main"\ in session \<^verbatim>\"HOL"\. The optional percentage has the same meaning as in \<^bold>\type\~\node_status\ below. \<^item> \<^bold>\type\~\timing = {elapsed: double, cpu: double, gc: double}\ refers to common Isabelle timing information in seconds, usually with a precision of three digits after the point (whole milliseconds). \<^item> \<^bold>\type\~\uuid = string\ refers to a Universally Unique Identifier (UUID) as plain text.\<^footnote>\See \<^url>\https://www.ietf.org/rfc/rfc4122.txt\ and \<^url>\https://docs.oracle.com/en/java/javase/17/docs/api/java.base/java/util/UUID.html\.\ Such identifiers are created as private random numbers of the server and only revealed to the client that creates a certain resource (e.g.\ task or session). A client may disclose this information for use in a different client connection: this allows to share sessions between multiple connections. Client commands need to provide syntactically wellformed UUIDs: this is trivial to achieve by using only identifiers that have been produced by the server beforehand. \<^item> \<^bold>\type\~\task = {task: uuid}\ identifies a newly created asynchronous task and thus allows the client to control it by the \<^verbatim>\cancel\ command. The same task identification is included in all messages produced by this task. \<^item> \<^bold>\type\ \session_id = {session_id: uuid}\ identifies a newly created PIDE session managed by the server. Sessions are independent of client connections and may be shared by different clients, as long as the internal session identifier is known. \<^item> \<^bold>\type\ \node = {node_name: string, theory_name: string}\ represents the internal node name of a theory. The \node_name\ is derived from the canonical theory file-name (e.g.\ \<^verbatim>\"~~/src/HOL/Examples/Seq.thy"\ after normalization within the file-system). The \theory_name\ is the session-qualified theory name (e.g.\ \<^verbatim>\HOL-Examples.Seq\). \<^item> \<^bold>\type\ \node_status = {ok: bool, total: int, unprocessed: int, running: int, warned: int, failed: int, finished: int, canceled: bool, consolidated: bool, percentage: int}\ represents a formal theory node status of the PIDE document model as follows. \<^item> Fields \total\, \unprocessed\, \running\, \warned\, \failed\, \finished\ account for individual commands within a theory node; \ok\ is an abstraction for \failed = 0\. \<^item> The \canceled\ flag tells if some command in the theory has been spontaneously canceled (by an Interrupt exception that could also indicate resource problems). \<^item> The \consolidated\ flag indicates whether the outermost theory command structure has finished (or failed) and the final \<^theory_text>\end\ command has been checked. \<^item> The \percentage\ field tells how far the node has been processed. It ranges between 0 and 99 in normal operation, and reaches 100 when the node has been formally consolidated as described above. \ section \Server commands and results\ text \ Here follows an overview of particular Isabelle server commands with their results, which are usually represented as JSON values with types according to \secref{sec:json-types}. The general format of input and output messages is described in \secref{sec:input-output-messages}. The relevant Isabelle/Scala source files are: \<^medskip> \begin{tabular}{l} \<^file>\$ISABELLE_HOME/src/Pure/Tools/server_commands.scala\ \\ \<^file>\$ISABELLE_HOME/src/Pure/Tools/server.scala\ \\ \<^file>\$ISABELLE_HOME/src/Pure/General/json.scala\ \\ \end{tabular} \ subsection \Command \<^verbatim>\help\\ text \ \begin{tabular}{ll} regular result: & \<^verbatim>\OK\ \[string]\ \\ \end{tabular} \<^medskip> The \<^verbatim>\help\ command has no argument and returns the list of server command names. This is occasionally useful for interactive experimentation (see also @{tool client} in \secref{sec:tool-client}). \ subsection \Command \<^verbatim>\echo\\ text \ \begin{tabular}{ll} argument: & \any\ \\ regular result: & \<^verbatim>\OK\ \any\ \\ \end{tabular} \<^medskip> The \<^verbatim>\echo\ command is the identity function: it returns its argument as regular result. This is occasionally useful for testing and interactive experimentation (see also @{tool client} in \secref{sec:tool-client}). The Scala type of \<^verbatim>\echo\ is actually more general than given above: \<^verbatim>\Unit\, \<^verbatim>\XML.Elem\, \<^verbatim>\JSON.T\ work uniformly. Note that \<^verbatim>\XML.Elem\ might be difficult to type on the console in its YXML syntax (\secref{sec:yxml-vs-xml}). \ subsection \Command \<^verbatim>\shutdown\\ text \ \begin{tabular}{ll} regular result: & \<^verbatim>\OK\ \\ \end{tabular} \<^medskip> The \<^verbatim>\shutdown\ command has no argument and result value. It forces a shutdown of the connected server process, stopping all open sessions and closing the server socket. This may disrupt pending commands on other connections! \<^medskip> The command-line invocation \<^verbatim>\isabelle server -x\ opens a server connection and issues a \<^verbatim>\shutdown\ command (see also \secref{sec:tool-server}). \ subsection \Command \<^verbatim>\cancel\\ text \ \begin{tabular}{ll} argument: & \task\ \\ regular result: & \<^verbatim>\OK\ \\ \end{tabular} \<^medskip> The command \<^verbatim>\cancel {"task":\~\id\\<^verbatim>\}\ attempts to cancel the specified task. Cancellation is merely a hint that the client prefers an ongoing process to be stopped. The command always succeeds formally, but it may get ignored by a task that is still running; it might also refer to a non-existing or no-longer existing task (without producing an error). Successful cancellation typically leads to an asynchronous failure of type \<^verbatim>\FAILED {\\task: uuid, message:\~\<^verbatim>\"Interrupt"}\. A different message is also possible, depending how the task handles the event. \ subsection \Command \<^verbatim>\session_build\ \label{sec:command-session-build}\ text \ \begin{tabular}{lll} argument: & \session_build_args\ \\ immediate result: & \<^verbatim>\OK\ \task\ \\ notifications: & \<^verbatim>\NOTE\ \task \ (theory_progress | message)\ \\ regular result: & \<^verbatim>\FINISHED\ \task \ session_build_results\ \\ error result: & \<^verbatim>\FAILED\ \task \ error_message \ session_build_results\ \\[2ex] \end{tabular} \begin{tabular}{lll} \<^bold>\type\ \session_build_args =\ \\ \quad\{session: string,\ \\ \quad~~\preferences?: string,\ & \<^bold>\default:\ server preferences \\ \quad~~\options?: [string],\ \\ \quad~~\dirs?: [string],\ \\ \quad~~\include_sessions: [string],\ \\ \quad~~\verbose?: bool}\ \\[2ex] \end{tabular} \begin{tabular}{ll} \<^bold>\type\ \session_build_result =\ \\ \quad\{session: string,\ \\ \quad~~\ok: bool,\ \\ \quad~~\return_code: int,\ \\ \quad~~\timeout: bool,\ \\ \quad~~\timing: timing}\ \\[2ex] \<^bold>\type\ \session_build_results =\ \\ \quad\{ok: bool,\ \\ \quad~~\return_code: int,\ \\ \quad~~\sessions: [session_build_result]}\ \\ \end{tabular} \ text \ The \<^verbatim>\session_build\ command prepares a session image for interactive use of theories. This is a limited version of command-line tool @{tool build} (\secref{sec:tool-build}), with specific options to request a formal context for an interactive PIDE session. The build process is asynchronous, with notifications that inform about the progress of loaded theories. Some further informative messages are output as well. Coordination of independent build processes is at the discretion of the client (or end-user), just as for @{tool build} and @{tool jedit}. There is no built-in coordination of conflicting builds with overlapping hierarchies of session images. In the worst case, a session image produced by one task may get overwritten by another task! \ subsubsection \Arguments\ text \ The \session\ field specifies the target session name. The build process will produce all required ancestor images according to the overall session graph. \<^medskip> The environment of Isabelle system options is determined from \preferences\ that are augmented by \options\, which is a list individual updates of the form the \name\\<^verbatim>\=\\value\ or \name\ (the latter abbreviates \name\\<^verbatim>\=true\); see also command-line option \<^verbatim>\-o\ for @{tool build}. The preferences are loaded from the file \<^path>\$ISABELLE_HOME_USER/etc/preferences\ by default, but the client may provide alternative contents for it (as text, not a file-name). This could be relevant in situations where client and server run in different operating-system contexts. \<^medskip> The \dirs\ field specifies additional directories for session ROOT and ROOTS files (\secref{sec:session-root}). This augments the name space of available sessions; see also option \<^verbatim>\-d\ in @{tool build}. \<^medskip> The \include_sessions\ field specifies sessions whose theories should be included in the overall name space of session-qualified theory names. This corresponds to a \<^bold>\sessions\ specification in ROOT files (\secref{sec:session-root}). It enables the \<^verbatim>\use_theories\ command (\secref{sec:command-use-theories}) to refer to sources from other sessions in a robust manner, instead of relying on directory locations. \<^medskip> The \verbose\ field set to \<^verbatim>\true\ yields extra verbosity. The effect is similar to option \<^verbatim>\-v\ in @{tool build}. \ subsubsection \Intermediate output\ text \ The asynchronous notifications of command \<^verbatim>\session_build\ mainly serve as progress indicator: the output resembles that of the session build window of - Isabelle/jEdit after startup @{cite "isabelle-jedit"}. + Isabelle/jEdit after startup \<^cite>\"isabelle-jedit"\. For the client it is usually sufficient to print the messages in plain text, but note that \theory_progress\ also reveals formal \theory\ and \session\ names directly. \ subsubsection \Results\ text \ The overall \session_build_results\ contain both a summary and an entry \session_build_result\ for each session in the build hierarchy. The result is always provided, independently of overall success (\<^verbatim>\FINISHED\ task) or failure (\<^verbatim>\FAILED\ task). The \ok\ field tells abstractly, whether all required session builds came out as \ok\, i.e.\ with zero \return_code\. A non-zero \return_code\ indicates an error according to usual POSIX conventions for process exit. The individual \session_build_result\ entries provide extra fields: \<^item> \timeout\ tells if the build process was aborted after running too long, \<^item> \timing\ gives the overall process timing in the usual Isabelle format with elapsed, CPU, GC time. \ subsubsection \Examples\ text \ Build of a session image from the Isabelle distribution: @{verbatim [display] \session_build {"session": "HOL-Algebra"}\} Build a session image from the Archive of Formal Proofs: @{verbatim [display] \session_build {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\} \ subsection \Command \<^verbatim>\session_start\ \label{sec:command-session-start}\ text \ \begin{tabular}{lll} argument: & \session_build_args \ {print_mode?: [string]}\ \\ immediate result: & \<^verbatim>\OK\ \task\ \\ notifications: & \<^verbatim>\NOTE\ \task \ (theory_progress | message)\ \\ regular result: & \<^verbatim>\FINISHED\ \task \ session_id \ {tmp_dir: string}\ \\ error result: & \<^verbatim>\FAILED\ \task \ error_message\ \\[2ex] \end{tabular} \<^medskip> The \<^verbatim>\session_start\ command starts a new Isabelle/PIDE session with underlying Isabelle/ML process, based on a session image that it produces on demand using \<^verbatim>\session_build\. Thus it accepts all \session_build_args\ and produces similar notifications, but the detailed \session_build_results\ are omitted. The session build and startup process is asynchronous: when the task is finished, the session remains active for commands, until a \<^verbatim>\session_stop\ or \<^verbatim>\shutdown\ command is sent to the server. Sessions are independent of client connections: it is possible to start a session and later apply \<^verbatim>\use_theories\ on different connections, as long as the internal session identifier is known: shared theory imports will be used only once (and persist until purged explicitly). \ subsubsection \Arguments\ text \ Most arguments are shared with \<^verbatim>\session_build\ (\secref{sec:command-session-build}). \<^medskip> The \print_mode\ field adds identifiers of print modes to be made active for this session. For example, \<^verbatim>\"print_mode": ["ASCII"]\ prefers ASCII replacement syntax over mathematical Isabelle symbols. See also option \<^verbatim>\-m\ in @{tool process} (\secref{sec:tool-process}). \ subsubsection \Results\ text \ The \session_id\ provides the internal identification of the session object within the server process. It can remain active as long as the server is running, independently of the current client connection. \<^medskip> The \tmp_dir\ field refers to a temporary directory that is specifically created for this session and deleted after it has been stopped. This may serve as auxiliary file-space for the \<^verbatim>\use_theories\ command, but concurrent use requires some care in naming temporary files, e.g.\ by using sub-directories with globally unique names. As \tmp_dir\ is the default \master_dir\ for commands \<^verbatim>\use_theories\ and \<^verbatim>\purge_theories\, theory files copied there may be used without further path specification. \ subsubsection \Examples\ text \ Start a default Isabelle/HOL session: @{verbatim [display] \session_start {"session": "HOL"}\} Start a session from the Archive of Formal Proofs: @{verbatim [display] \session_start {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\} \ subsection \Command \<^verbatim>\session_stop\\ text \ \begin{tabular}{ll} argument: & \session_id\ \\ immediate result: & \<^verbatim>\OK\ \task\ \\ regular result: & \<^verbatim>\FINISHED\ \task \ session_stop_result\ \\ error result: & \<^verbatim>\FAILED\ \task \ error_message \ session_stop_result\ \\[2ex] \end{tabular} \begin{tabular}{l} \<^bold>\type\ \session_stop_result = {ok: bool, return_code: int}\ \end{tabular} \<^medskip> The \<^verbatim>\session_stop\ command forces a shutdown of the identified PIDE session. This asynchronous tasks usually finishes quickly. Failure only happens in unusual situations, according to the return code of the underlying Isabelle/ML process. \ subsubsection \Arguments\ text \ The \session_id\ provides the UUID originally created by the server for this session. \ subsubsection \Results\ text \ The \ok\ field tells abstractly, whether the Isabelle/ML process has terminated properly. The \return_code\ field expresses this information according to usual POSIX conventions for process exit. \ subsection \Command \<^verbatim>\use_theories\ \label{sec:command-use-theories}\ text \ \begin{tabular}{ll} argument: & \use_theories_arguments\ \\ immediate result: & \<^verbatim>\OK\ \task\ \\ regular result: & \<^verbatim>\FINISHED\ \use_theories_results\ \\ \end{tabular} \begin{tabular}{ll} \<^bold>\type\ \use_theories_arguments =\ \\ \quad\{session_id: uuid,\ \\ \quad~~\theories: [string],\ \\ \quad~~\master_dir?: string,\ & \<^bold>\default:\ session \tmp_dir\ \\ \quad~~\pretty_margin?: double,\ & \<^bold>\default:\ \<^verbatim>\76\ \\ \quad~~\unicode_symbols?: bool,\ \\ \quad~~\export_pattern?: string,\ \\ \quad~~\check_delay?: double,\ & \<^bold>\default:\ \<^verbatim>\0.5\ \\ \quad~~\check_limit?: int,\ \\ \quad~~\watchdog_timeout?: double,\ & \<^bold>\default:\ \<^verbatim>\600.0\ \\ \quad~~\nodes_status_delay?: double}\ & \<^bold>\default:\ \<^verbatim>\-1.0\ \\ \end{tabular} \begin{tabular}{ll} \<^bold>\type\ \export =\ \\ \quad~~\{name: string, base64: bool, body: string}\ \\ \<^bold>\type\ \node_results =\ \\ \quad~~\{status: node_status, messages: [message], exports: [export]}\ \\ \<^bold>\type\ \nodes_status =\ \\ \quad~~\[node \ {status: node_status}]\ \\ \<^bold>\type\ \use_theories_results =\ \\ \quad\{ok: bool,\ \\ \quad~~\errors: [message],\ \\ \quad~~\nodes: [node \ node_results]}\ \\ \end{tabular} \<^medskip> The \<^verbatim>\use_theories\ command updates the identified session by adding the current version of theory files to it, while dependencies are resolved implicitly. The command succeeds eventually, when all theories have status \<^emph>\terminated\ or \<^emph>\consolidated\ in the sense of \node_status\ (\secref{sec:json-types}). Already used theories persist in the session until purged explicitly (\secref{sec:command-purge-theories}). This also means that repeated invocations of \<^verbatim>\use_theories\ are idempotent: it could make sense to do that with different values for \pretty_margin\ or \unicode_symbols\ to get different formatting for \errors\ or \messages\. \<^medskip> A non-empty \export_pattern\ means that theory \exports\ are retrieved (see \secref{sec:tool-export}). An \export\ \name\ roughly follows file-system standards: ``\<^verbatim>\/\'' separated list of base names (excluding special names like ``\<^verbatim>\.\'' or ``\<^verbatim>\..\''). The \base64\ field specifies the format of the \body\ string: it is true for a byte vector that cannot be represented as plain text in UTF-8 encoding, which means the string needs to be decoded as in \<^verbatim>\java.util.Base64.getDecoder.decode(String)\. \<^medskip> The status of PIDE processing is checked every \check_delay\ seconds, and bounded by \check_limit\ attempts (default: 0, i.e.\ unbounded). A \check_limit > 0\ effectively specifies a global timeout of \check_delay \ check_limit\ seconds. \<^medskip> If \watchdog_timeout\ is greater than 0, it specifies the timespan (in seconds) after the last command status change of Isabelle/PIDE, before finishing with a potentially non-terminating or deadlocked execution. \<^medskip> A non-negative \nodes_status_delay\ enables continuous notifications of kind \nodes_status\, with a field of name and type \nodes_status\. The time interval is specified in seconds; by default it is negative and thus disabled. \ subsubsection \Arguments\ text \ The \session_id\ is the identifier provided by the server, when the session was created (possibly on a different client connection). \<^medskip> The \theories\ field specifies theory names as in theory \<^theory_text>\imports\ or in ROOT \<^bold>\theories\. \<^medskip> The \master_dir\ field specifies the master directory of imported theories: it acts like the ``current working directory'' for locating theory files. This is irrelevant for \theories\ with an absolute path name (e.g.\ \<^verbatim>\"~~/src/HOL/Examples/Seq.thy"\) or session-qualified theory name (e.g.\ \<^verbatim>\"HOL-Examples.Seq"\). \<^medskip> The \pretty_margin\ field specifies the line width for pretty-printing. The default is suitable for classic console output. Formatting happens at the end of \<^verbatim>\use_theories\, when all prover messages are exported to the client. \<^medskip> The \unicode_symbols\ field set to \<^verbatim>\true\ renders message output for direct output on a Unicode capable channel, ideally with the Isabelle fonts as in Isabelle/jEdit. The default is to keep the symbolic representation of Isabelle text, e.g.\ \<^verbatim>\\\ instead of its rendering as \\\. This means the client needs to perform its own rendering before presenting it to the end-user. \ subsubsection \Results\ text \ The \ok\ field indicates overall success of processing the specified theories with all their dependencies. When \ok\ is \<^verbatim>\false\, the \errors\ field lists all errors cumulatively (including imported theories). The messages contain position information for the original theory nodes. \<^medskip> The \nodes\ field provides detailed information about each imported theory node. The individual fields are as follows: \<^item> \node_name\: the canonical name for the theory node, based on its file-system location; \<^item> \theory_name\: the logical theory name; \<^item> \status\: the overall node status, e.g.\ see the visualization in the - \Theories\ panel of Isabelle/jEdit @{cite "isabelle-jedit"}; + \Theories\ panel of Isabelle/jEdit \<^cite>\"isabelle-jedit"\; \<^item> \messages\: the main bulk of prover messages produced in this theory (with kind \<^verbatim>\writeln\, \<^verbatim>\warning\, \<^verbatim>\error\). \ subsubsection \Examples\ text \ Process some example theory from the Isabelle distribution, within the context of an already started session for Isabelle/HOL (see also \secref{sec:command-session-start}): @{verbatim [display] \use_theories {"session_id": ..., "theories": ["~~/src/HOL/Examples/Seq"]}\} \<^medskip> Process some example theories in the context of their (single) parent session: @{verbatim [display] \session_start {"session": "HOL-Library"} use_theories {"session_id": ..., "theories": ["~~/src/HOL/Unix/Unix"]} session_stop {"session_id": ...}\} \<^medskip> Process some example theories that import other theories via session-qualified theory names: @{verbatim [display] \session_start {"session": "HOL", "include_sessions": ["HOL-Unix"]} use_theories {"session_id": ..., "theories": ["HOL-Unix.Unix"]} session_stop {"session_id": ...}\} \ subsection \Command \<^verbatim>\purge_theories\ \label{sec:command-purge-theories}\ text \ \begin{tabular}{ll} argument: & \purge_theories_arguments\ \\ regular result: & \<^verbatim>\OK\ \purge_theories_result\ \\ \end{tabular} \begin{tabular}{lll} \<^bold>\type\ \purge_theories_arguments =\ \\ \quad\{session_id: uuid,\ \\ \quad~~\theories: [string],\ \\ \quad~~\master_dir?: string,\ & \<^bold>\default:\ session \tmp_dir\ \\ \quad~~\all?: bool}\ \\[2ex] \end{tabular} \begin{tabular}{ll} \<^bold>\type\ \purge_theories_result = {purged: [string]}\ \\ \end{tabular} \<^medskip> The \<^verbatim>\purge_theories\ command updates the identified session by removing theories that are no longer required: theories that are used in pending \<^verbatim>\use_theories\ tasks or imported by other theories are retained. \ subsubsection \Arguments\ text \ The \session_id\ is the identifier provided by the server, when the session was created (possibly on a different client connection). \<^medskip> The \theories\ field specifies theory names to be purged: imported dependencies are \<^emph>\not\ completed. Instead it is possible to provide the already completed import graph returned by \<^verbatim>\use_theories\ as \nodes\ / \node_name\. \<^medskip> The \master_dir\ field specifies the master directory as in \<^verbatim>\use_theories\. This is irrelevant, when passing fully-qualified theory node names (e.g.\ \node_name\ from \nodes\ in \use_theories_results\). \<^medskip> The \all\ field set to \<^verbatim>\true\ attempts to purge all presently loaded theories. \ subsubsection \Results\ text \ The \purged\ field gives the theory nodes that were actually removed. \<^medskip> The \retained\ field gives the remaining theory nodes, i.e.\ the complement of \purged\. \ end diff --git a/src/Doc/System/Sessions.thy b/src/Doc/System/Sessions.thy --- a/src/Doc/System/Sessions.thy +++ b/src/Doc/System/Sessions.thy @@ -1,1011 +1,1011 @@ (*:maxLineLen=78:*) theory Sessions imports Base begin chapter \Isabelle sessions and build management \label{ch:session}\ text \ An Isabelle \<^emph>\session\ consists of a collection of related theories that may be associated with formal documents (\chref{ch:present}). There is also a notion of \<^emph>\persistent heap\ image to capture the state of a session, similar to object-code in compiled programming languages. Thus the concept of session resembles that of a ``project'' in common IDE environments, but the specific name emphasizes the connection to interactive theorem proving: the session wraps-up the results of user-interaction with the prover in a persistent form. Application sessions are built on a given parent session, which may be built recursively on other parents. Following this path in the hierarchy eventually leads to some major object-logic session like \HOL\, which itself is based on \Pure\ as the common root of all sessions. Processing sessions may take considerable time. Isabelle build management helps to organize this efficiently. This includes support for parallel build jobs, in addition to the multithreaded theory and proof checking that is already provided by the prover process itself. \ section \Session ROOT specifications \label{sec:session-root}\ text \ Session specifications reside in files called \<^verbatim>\ROOT\ within certain directories, such as the home locations of registered Isabelle components or additional project directories given by the user. The ROOT file format follows the lexical conventions of the \<^emph>\outer syntax\ - of Isabelle/Isar, see also @{cite "isabelle-isar-ref"}. This defines common + of Isabelle/Isar, see also \<^cite>\"isabelle-isar-ref"\. This defines common forms like identifiers, names, quoted strings, verbatim text, nested comments etc. The grammar for @{syntax chapter_def}, @{syntax chapter_entry} and @{syntax session_entry} is given as syntax diagram below. Each ROOT file may contain multiple specifications like this. Chapters help to organize browser info (\secref{sec:info}), but have no formal meaning. The default chapter is ``\Unsorted\''. Chapter definitions, which are optional, allow to associate additional information. - Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple editing mode + Isabelle/jEdit \<^cite>\"isabelle-jedit"\ includes a simple editing mode \<^verbatim>\isabelle-root\ for session ROOT files, which is enabled by default for any file of that name. \<^rail>\ @{syntax_def chapter_def}: @'chapter_definition' @{syntax name} \ groups? description? ; @{syntax_def chapter_entry}: @'chapter' @{syntax name} ; @{syntax_def session_entry}: @'session' @{syntax system_name} groups? dir? '=' \ (@{syntax system_name} '+')? description? options? \ sessions? directories? (theories*) \ (document_theories?) (document_files*) \ (export_files*) (export_classpath?) ; groups: '(' (@{syntax name} +) ')' ; dir: @'in' @{syntax embedded} ; description: @'description' @{syntax text} ; options: @'options' opts ; opts: '[' ( (@{syntax name} '=' value | @{syntax name}) + ',' ) ']' ; value: @{syntax name} | @{syntax real} ; sessions: @'sessions' (@{syntax system_name}+) ; directories: @'directories' (dir+) ; theories: @'theories' opts? (theory_entry+) ; theory_entry: @{syntax system_name} ('(' @'global' ')')? ; document_theories: @'document_theories' (@{syntax name}+) ; document_files: @'document_files' ('(' dir ')')? (@{syntax embedded}+) ; export_files: @'export_files' ('(' dir ')')? ('[' nat ']')? \ (@{syntax embedded}+) ; export_classpath: @'export_classpath' (@{syntax embedded}*) \ \<^descr> \isakeyword{chapter{\isacharunderscorekeyword}definition}~\A (groups)\ associates a collection of groups with chapter \A\. All sessions that belong to this chapter will automatically become members of these groups. \<^descr> \isakeyword{session}~\A = B + body\ defines a new session \A\ based on parent session \B\, with its content given in \body\ (imported sessions and theories). Note that a parent (like \HOL\) is mandatory in practical applications: only Isabelle/Pure can bootstrap itself from nothing. All such session specifications together describe a hierarchy (graph) of sessions, with globally unique names. The new session name \A\ should be sufficiently long and descriptive to stand on its own in a potentially large library. \<^descr> \isakeyword{session}~\A (groups)\ indicates a collection of groups where the new session is a member. Group names are uninterpreted and merely follow certain conventions. For example, the Isabelle distribution tags some important sessions by the group name called ``\main\''. Other projects may invent their own conventions, but this requires some care to avoid clashes within this unchecked name space. \<^descr> \isakeyword{session}~\A\~\isakeyword{in}~\dir\ specifies an explicit directory for this session; by default this is the current directory of the \<^verbatim>\ROOT\ file. All theory files are located relatively to the session directory. The prover process is run within the same as its current working directory. \<^descr> \isakeyword{description}~\text\ is a free-form description for this session (or chapter), e.g. for presentation purposes. \<^descr> \isakeyword{options}~\[x = a, y = b, z]\ defines separate options (\secref{sec:system-options}) that are used when processing this session, but \<^emph>\without\ propagation to child sessions. Note that \z\ abbreviates \z = true\ for Boolean options. \<^descr> \isakeyword{sessions}~\names\ specifies sessions that are \<^emph>\imported\ into the current name space of theories. This allows to refer to a theory \A\ from session \B\ by the qualified name \B.A\ --- although it is loaded again into the current ML process, which is in contrast to a theory that is already present in the \<^emph>\parent\ session. Theories that are imported from other sessions are excluded from the current session document. \<^descr> \isakeyword{directories}~\dirs\ specifies additional directories for import of theory files via \isakeyword{theories} within \<^verbatim>\ROOT\ or \<^theory_text>\imports\ within a theory; \dirs\ are relative to the main session directory (cf.\ \isakeyword{session} \dots \isakeyword{in}~\dir\). These directories need to be exclusively assigned to a unique session, without implicit sharing of file-system locations. \<^descr> \isakeyword{theories}~\options names\ specifies a block of theories that are processed within an environment that is augmented by the given options, in addition to the global session options given before. Any number of blocks of \isakeyword{theories} may be given. Options are only active for each \isakeyword{theories} block separately. A theory name that is followed by \(\\isakeyword{global}\)\ is treated literally in other session specifications or theory imports --- the normal situation is to qualify theory names by the session name; this ensures globally unique names in big session graphs. Global theories are usually the entry points to major logic sessions: \Pure\, \Main\, \Complex_Main\, \HOLCF\, \IFOL\, \FOL\, \ZF\, \ZFC\ etc. Regular Isabelle applications should not claim any global theory names. \<^descr> \isakeyword{document_theories}~\names\ specifies theories from other sessions that should be included in the generated document source directory. These theories need to be explicit imports in the current session, or implicit imports from the underlying hierarchy of parent sessions. The generated \<^verbatim>\session.tex\ file is not affected: the session's {\LaTeX} setup needs to \<^verbatim>\\input{\\\\\<^verbatim>\}\ generated \<^verbatim>\.tex\ files separately. \<^descr> \isakeyword{document_files}~\(\\isakeyword{in}~\base_dir) files\ lists source files for document preparation, typically \<^verbatim>\.tex\ and \<^verbatim>\.sty\ for {\LaTeX}. Only these explicitly given files are copied from the base directory to the document output directory, before formal document processing is started (see also \secref{sec:tool-document}). The local path structure of the \files\ is preserved, which allows to reconstruct the original directory hierarchy of \base_dir\. The default \base_dir\ is \<^verbatim>\document\ within the session root directory. \<^descr> \isakeyword{export_files}~\(\\isakeyword{in}~\target_dir) [number] patterns\ specifies theory exports that may get written to the file-system, e.g. via @{tool_ref build} with option \<^verbatim>\-e\ (\secref{sec:tool-build}). The \target_dir\ specification is relative to the session root directory; its default is \<^verbatim>\export\. Exports are selected via \patterns\ as in @{tool_ref export} (\secref{sec:tool-export}). The number given in brackets (default: 0) specifies the prefix of elements that should be removed from each name: it allows to reduce the resulting directory hierarchy at the danger of overwriting files due to loss of uniqueness. \<^descr> \isakeyword{export_classpath}~\patterns\ specifies export artifacts that should be included into the local Java/Scala classpath of this session context. This is only relevant for tools that allow dynamic loading of service classes (\secref{sec:scala-build}), while most other Isabelle/Scala tools require global configuration during system startup. An empty list of \patterns\ defaults to \<^verbatim>\"*:classpath/*.jar"\, which fits to the naming convention of JAR modules produced by the Isabelle/Isar command - \<^theory_text>\scala_build_generated_files\ @{cite "isabelle-isar-ref"}. + \<^theory_text>\scala_build_generated_files\ \<^cite>\"isabelle-isar-ref"\. \ subsubsection \Examples\ text \ See \<^file>\~~/src/HOL/ROOT\ for a diversity of practically relevant situations, although it uses relatively complex quasi-hierarchic naming conventions like \<^verbatim>\HOL-SPARK\, \<^verbatim>\HOL-SPARK-Examples\. An alternative is to use unqualified names that are relatively long and descriptive, as in the Archive of Formal Proofs (\<^url>\https://isa-afp.org\), for example. \ section \System build options \label{sec:system-options}\ text \ See \<^file>\~~/etc/options\ for the main defaults provided by the Isabelle - distribution. Isabelle/jEdit @{cite "isabelle-jedit"} includes a simple + distribution. Isabelle/jEdit \<^cite>\"isabelle-jedit"\ includes a simple editing mode \<^verbatim>\isabelle-options\ for this file-format. The following options are particularly relevant to build Isabelle sessions, in particular with document preparation (\chref{ch:present}). \<^item> @{system_option_def "browser_info"} controls output of HTML browser info, see also \secref{sec:info}. \<^item> @{system_option_def "document"} controls document output for a particular session or theory; \<^verbatim>\document=pdf\ or \<^verbatim>\document=true\ means enabled, \<^verbatim>\document=""\ or \<^verbatim>\document=false\ means disabled (especially for particular theories). \<^item> @{system_option_def "document_output"} specifies an alternative directory for generated output of the document preparation system; the default is within the @{setting "ISABELLE_BROWSER_INFO"} hierarchy as explained in \secref{sec:info}. See also @{tool mkroot}, which generates a default configuration with output readily available to the author of the document. \<^item> @{system_option_def "document_echo"} informs about document file names during session presentation. \<^item> @{system_option_def "document_variants"} specifies document variants as a colon-separated list of \name=tags\ entries. The default name is \<^verbatim>\document\, without additional tags. Tags are specified as a comma separated list of modifier/name pairs and tell {\LaTeX} how to interpret certain Isabelle command regions: ``\<^verbatim>\+\\foo\'' (or just ``\foo\'') means to keep, ``\<^verbatim>\-\\foo\'' to drop, and ``\<^verbatim>\/\\foo\'' to fold text tagged as \foo\. The builtin default is equivalent to the tag specification ``\<^verbatim>\+document,+theory,+proof,+ML,+visible,-invisible,+important,+unimportant\''; see also the {\LaTeX} macros \<^verbatim>\\isakeeptag\, \<^verbatim>\\isadroptag\, and \<^verbatim>\\isafoldtag\, in \<^file>\~~/lib/texinputs/isabelle.sty\. In contrast, \<^verbatim>\document_variants=document:outline=/proof,/ML\ indicates two documents: the one called \<^verbatim>\document\ with default tags, and the other called \<^verbatim>\outline\ where proofs and ML sections are folded. Document variant names are just a matter of conventions. It is also possible to use different document variant names (without tags) for different document root entries, see also \secref{sec:tool-document}. \<^item> @{system_option_def "document_tags"} specifies alternative command tags as a comma-separated list of items: either ``\command\\<^verbatim>\%\\tag\'' for a specific command, or ``\<^verbatim>\%\\tag\'' as default for all other commands. This is occasionally useful to control the global visibility of commands via session options (e.g.\ in \<^verbatim>\ROOT\). \<^item> @{system_option_def "document_comment_latex"} enables regular {\LaTeX} \<^verbatim>\comment.sty\, instead of the historic version for plain {\TeX} (default). The latter is much faster, but in conflict with {\LaTeX} classes like Dagstuhl LIPIcs\<^footnote>\\<^url>\https://github.com/dagstuhl-publishing/styles\\. \<^item> @{system_option_def "document_bibliography"} explicitly enables the use of \<^verbatim>\bibtex\; the default is to check the presence of \<^verbatim>\root.bib\, but it could have a different name. \<^item> @{system_option_def "document_heading_prefix"} specifies a prefix for the {\LaTeX} macro names generated from Isar commands like \<^theory_text>\chapter\, \<^theory_text>\section\ etc. The default is \<^verbatim>\isamarkup\, e.g. \<^theory_text>\section\ becomes \<^verbatim>\\isamarkupsection\. \<^item> @{system_option_def "threads"} determines the number of worker threads for parallel checking of theories and proofs. The default \0\ means that a sensible maximum value is determined by the underlying hardware. For machines with many cores or with hyperthreading, this sometimes requires manual adjustment (on the command-line or within personal settings or preferences, not within a session \<^verbatim>\ROOT\). \<^item> @{system_option_def "condition"} specifies a comma-separated list of process environment variables (or Isabelle settings) that are required for the subsequent theories to be processed. Conditions are considered ``true'' if the corresponding environment value is defined and non-empty. \<^item> @{system_option_def "timeout"} and @{system_option_def "timeout_scale"} specify a real wall-clock timeout for the session as a whole: the two values are multiplied and taken as the number of seconds. Typically, @{system_option "timeout"} is given for individual sessions, and @{system_option "timeout_scale"} as global adjustment to overall hardware performance. The timer is controlled outside the ML process by the JVM that runs Isabelle/Scala. Thus it is relatively reliable in canceling processes that get out of control, even if there is a deadlock without CPU time usage. \<^item> @{system_option_def "profiling"} specifies a mode for global ML profiling. Possible values are the empty string (disabled), \<^verbatim>\time\ for \<^ML>\profile_time\ and \<^verbatim>\allocations\ for \<^ML>\profile_allocations\. Results appear near the bottom of the session log file. \<^item> @{system_option_def system_log} specifies an optional log file for low-level messages produced by \<^ML>\Output.system_message\ in Isabelle/ML; the standard value ``\<^verbatim>\-\'' refers to console progress of the build job. \<^item> @{system_option_def "system_heaps"} determines the directories for session heap images: \<^path>\$ISABELLE_HEAPS\ is the user directory and \<^path>\$ISABELLE_HEAPS_SYSTEM\ the system directory (usually within the Isabelle application). For \<^verbatim>\system_heaps=false\, heaps are stored in the user directory and may be loaded from both directories. For \<^verbatim>\system_heaps=true\, store and load happens only in the system directory. The @{tool_def options} tool prints Isabelle system options. Its command-line usage is: @{verbatim [display] \Usage: isabelle options [OPTIONS] [MORE_OPTIONS ...] Options are: -b include $ISABELLE_BUILD_OPTIONS -g OPTION get value of OPTION -l list options -x FILE export to FILE in YXML format Report Isabelle system options, augmented by MORE_OPTIONS given as arguments NAME=VAL or NAME.\} The command line arguments provide additional system options of the form \name\\<^verbatim>\=\\value\ or \name\ for Boolean options. Option \<^verbatim>\-b\ augments the implicit environment of system options by the ones of @{setting ISABELLE_BUILD_OPTIONS}, cf.\ \secref{sec:tool-build}. Option \<^verbatim>\-g\ prints the value of the given option. Option \<^verbatim>\-l\ lists all options with their declaration and current value. Option \<^verbatim>\-x\ specifies a file to export the result in YXML format, instead of printing it in human-readable form. \ section \Invoking the build process \label{sec:tool-build}\ text \ The @{tool_def build} tool invokes the build process for Isabelle sessions. It manages dependencies between sessions, related sources of theories and auxiliary files, and target heap images. Accordingly, it runs instances of the prover process with optional document preparation. Its command-line usage is:\<^footnote>\Isabelle/Scala provides the same functionality via \<^scala_method>\isabelle.Build.build\.\ @{verbatim [display] \Usage: isabelle build [OPTIONS] [SESSIONS ...] Options are: -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -N cyclic shuffling of NUMA CPU nodes (performance tuning) -P DIR enable HTML/PDF presentation in directory (":" for default) -R refer to requirements of selected sessions -S soft build: only observe changes of sources, not heap images -X NAME exclude sessions from group NAME and all descendants -a select all sessions -b build heap images -c clean build -d DIR include session directory -e export files from session specification into file-system -f fresh build -g NAME select session group NAME -j INT maximum number of parallel jobs (default 1) -k KEYWORD check theory sources for conflicts with proposed keywords -l list session source files -n no build -- take existing build databases -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose -x NAME exclude session NAME and all descendants Build and manage Isabelle sessions, depending on implicit settings: ISABELLE_TOOL_JAVA_OPTIONS="..." ISABELLE_BUILD_OPTIONS="..." ML_PLATFORM="..." ML_HOME="..." ML_SYSTEM="..." ML_OPTIONS="..."\} \<^medskip> Isabelle sessions are defined via session ROOT files as described in (\secref{sec:session-root}). The totality of sessions is determined by collecting such specifications from all Isabelle component directories (\secref{sec:components}), augmented by more directories given via options \<^verbatim>\-d\~\DIR\ on the command line. Each such directory may contain a session \<^verbatim>\ROOT\ file with several session specifications. Any session root directory may refer recursively to further directories of the same kind, by listing them in a catalog file \<^verbatim>\ROOTS\ line-by-line. This helps to organize large collections of session specifications, or to make \<^verbatim>\-d\ command line options persistent (e.g.\ in \<^verbatim>\$ISABELLE_HOME_USER/ROOTS\). \<^medskip> The subset of sessions to be managed is determined via individual \SESSIONS\ given as command-line arguments, or session groups that are given via one or more options \<^verbatim>\-g\~\NAME\. Option \<^verbatim>\-a\ selects all sessions. The build tool takes session dependencies into account: the set of selected sessions is completed by including all ancestors. \<^medskip> One or more options \<^verbatim>\-B\~\NAME\ specify base sessions to be included (all descendants wrt.\ the session parent or import graph). \<^medskip> One or more options \<^verbatim>\-x\~\NAME\ specify sessions to be excluded (all descendants wrt.\ the session parent or import graph). Option \<^verbatim>\-X\ is analogous to this, but excluded sessions are specified by session group membership. \<^medskip> Option \<^verbatim>\-R\ reverses the selection in the sense that it refers to its requirements: all ancestor sessions excluding the original selection. This allows to prepare the stage for some build process with different options, before running the main build itself (without option \<^verbatim>\-R\). \<^medskip> Option \<^verbatim>\-D\ is similar to \<^verbatim>\-d\, but selects all sessions that are defined in the given directories. \<^medskip> Option \<^verbatim>\-S\ indicates a ``soft build'': the selection is restricted to those sessions that have changed sources (according to actually imported theories). The status of heap images is ignored. \<^medskip> The build process depends on additional options (\secref{sec:system-options}) that are passed to the prover eventually. The settings variable @{setting_ref ISABELLE_BUILD_OPTIONS} allows to provide additional defaults, e.g.\ \<^verbatim>\ISABELLE_BUILD_OPTIONS="document=pdf threads=4"\. Moreover, the environment of system build options may be augmented on the command line via \<^verbatim>\-o\~\name\\<^verbatim>\=\\value\ or \<^verbatim>\-o\~\name\, which abbreviates \<^verbatim>\-o\~\name\\<^verbatim>\=true\ for Boolean or string options. Multiple occurrences of \<^verbatim>\-o\ on the command-line are applied in the given order. \<^medskip> Option \<^verbatim>\-P\ enables PDF/HTML presentation in the given directory, where ``\<^verbatim>\-P:\'' refers to the default @{setting_ref ISABELLE_BROWSER_INFO} (or @{setting_ref ISABELLE_BROWSER_INFO_SYSTEM}). This applies only to explicitly selected sessions; note that option \<^verbatim>\-R\ allows to select all requirements separately. \<^medskip> Option \<^verbatim>\-b\ ensures that heap images are produced for all selected sessions. By default, images are only saved for inner nodes of the hierarchy of sessions, as required for other sessions to continue later on. \<^medskip> Option \<^verbatim>\-c\ cleans the selected sessions (all descendants wrt.\ the session parent or import graph) before performing the specified build operation. \<^medskip> Option \<^verbatim>\-e\ executes the \isakeyword{export_files} directives from the ROOT specification of all explicitly selected sessions: the status of the session build database needs to be OK, but the session could have been built earlier. Using \isakeyword{export_files}, a session may serve as abstract interface for add-on build artefacts, but these are only materialized on explicit request: without option \<^verbatim>\-e\ there is no effect on the physical file-system yet. \<^medskip> Option \<^verbatim>\-f\ forces a fresh build of all selected sessions and their requirements. \<^medskip> Option \<^verbatim>\-n\ omits the actual build process after the preparatory stage (including optional cleanup). The overall return code always the status of the set of selected sessions. Add-on builds (like presentation) are run for successful sessions, i.e.\ already finished ones. \<^medskip> Option \<^verbatim>\-j\ specifies the maximum number of parallel build jobs (prover processes). Each prover process is subject to a separate limit of parallel worker threads, cf.\ system option @{system_option_ref threads}. \<^medskip> Option \<^verbatim>\-N\ enables cyclic shuffling of NUMA CPU nodes. This may help performance tuning on Linux servers with separate CPU/memory modules. \<^medskip> Option \<^verbatim>\-v\ increases the general level of verbosity. \<^medskip> Option \<^verbatim>\-l\ lists the source files that contribute to a session. \<^medskip> Option \<^verbatim>\-k\ specifies a newly proposed keyword for outer syntax. It is possible to use option \<^verbatim>\-k\ repeatedly to check multiple keywords. The theory sources are checked for conflicts wrt.\ this hypothetical change of syntax, e.g.\ to reveal occurrences of identifiers that need to be quoted. \ subsubsection \Examples\ text \ Build a specific logic image: @{verbatim [display] \ isabelle build -b HOLCF\} \<^smallskip> Build the main group of logic images: @{verbatim [display] \ isabelle build -b -g main\} \<^smallskip> Build all descendants (and requirements) of \<^verbatim>\FOL\ and \<^verbatim>\ZF\: @{verbatim [display] \ isabelle build -B FOL -B ZF\} \<^smallskip> Build all sessions where sources have changed (ignoring heaps): @{verbatim [display] \ isabelle build -a -S\} \<^smallskip> Provide a general overview of the status of all Isabelle sessions, without building anything: @{verbatim [display] \ isabelle build -a -n -v\} \<^smallskip> Build all sessions with HTML browser info and PDF document preparation: @{verbatim [display] \ isabelle build -a -o browser_info -o document\} \<^smallskip> Build all sessions with a maximum of 8 parallel prover processes and 4 worker threads each (on a machine with many cores): @{verbatim [display] \ isabelle build -a -j8 -o threads=4\} \<^smallskip> Build some session images with cleanup of their descendants, while retaining their ancestry: @{verbatim [display] \ isabelle build -b -c HOL-Library HOL-Algebra\} \<^smallskip> HTML/PDF presentation for sessions that happen to be properly built already, without rebuilding anything except the missing browser info: @{verbatim [display] \ isabelle build -a -n -o browser_info\} \<^smallskip> Clean all sessions without building anything: @{verbatim [display] \ isabelle build -a -n -c\} \<^smallskip> Build all sessions from some other directory hierarchy, according to the settings variable \<^verbatim>\AFP\ that happens to be defined inside the Isabelle environment: @{verbatim [display] \ isabelle build -D '$AFP'\} \<^smallskip> Inform about the status of all sessions required for AFP, without building anything yet: @{verbatim [display] \ isabelle build -D '$AFP' -R -v -n\} \ section \Print messages from build database \label{sec:tool-log}\ text \ The @{tool_def "log"} tool prints prover messages from the build database of the given session. Its command-line usage is: @{verbatim [display] \Usage: isabelle log [OPTIONS] [SESSIONS ...] Options are: -H REGEX filter messages by matching against head -M REGEX filter messages by matching against body -T NAME restrict to given theories (multiple options possible) -U output Unicode symbols -m MARGIN margin for pretty printing (default: 76.0) -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v print all messages, including information etc. Print messages from the build database of the given sessions, without any checks against current sources nor session structure: results from old sessions or failed builds can be printed as well. Multiple options -H and -M are conjunctive: all given patterns need to match. Patterns match any substring, but ^ or $ may be used to match the start or end explicitly.\} The specified session databases are taken as is, with formal checking against current sources: There is \<^emph>\no\ implicit build process involved, so it is possible to retrieve error messages from a failed session as well. The order of messages follows the source positions of source files; thus the result is mostly deterministic, independent of the somewhat erratic evaluation of parallel processing. \<^medskip> Option \<^verbatim>\-o\ allows to change system options, as in @{tool build} (\secref{sec:tool-build}). This may affect the storage space for the build database, notably via @{system_option system_heaps}, or @{system_option build_database_server} and its relatives. \<^medskip> Option \<^verbatim>\-T\ restricts output to given theories: multiple entries are possible by repeating this option on the command-line. The default is to refer to \<^emph>\all\ theories used in the original session build process. \<^medskip> Options \<^verbatim>\-m\ and \<^verbatim>\-U\ modify pretty printing and output of Isabelle symbols. The default is for an old-fashioned ASCII terminal at 80 characters per line (76 + 4 characters to prefix warnings or errors). \<^medskip> Option \<^verbatim>\-v\ prints all messages from the session database that are normally inlined into the source text, including information messages etc. \<^medskip> Options \<^verbatim>\-H\ and \<^verbatim>\-M\ filter messages according to their header or body content, respectively. The header follows a very basic format that makes it easy to match message kinds (e.g. \<^verbatim>\Warning\ or \<^verbatim>\Error\) and file names (e.g. \<^verbatim>\src/HOL/Nat.thy\). The body is usually pretty-printed, but for matching it is treated like one long line: blocks are ignored and breaks are turned into plain spaces (according to their formal width). The syntax for patters follows regular expressions of the Java platform.\<^footnote>\\<^url>\https://docs.oracle.com/en/java/javase/17/docs/api/java.base/java/util/regex/Pattern.html\\ \ subsubsection \Examples\ text \ Print messages from theory \<^verbatim>\HOL.Nat\ of session \<^verbatim>\HOL\, using Unicode rendering of Isabelle symbols and a margin of 100 characters: @{verbatim [display] \ isabelle log -T HOL.Nat -U -m 100 HOL\} Print warnings about ambiguous input (inner syntax) of session \<^verbatim>\HOL-Library\, which is built beforehand: @{verbatim [display] \ isabelle build HOL-Library isabelle log -H "Warning" -M "Ambiguous input" HOL-Library\} Print all errors from all sessions, e.g. from a partial build of Isabelle/AFP: @{verbatim [display] \ isabelle log -H "Error" $(isabelle sessions -a -d AFP/thys)\} \ section \Retrieve theory exports \label{sec:tool-export}\ text \ The @{tool_def "export"} tool retrieves theory exports from the session database. Its command-line usage is: @{verbatim [display] \Usage: isabelle export [OPTIONS] SESSION Options are: -O DIR output directory for exported files (default: "export") -d DIR include session directory -l list exports -n no build of session -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -p NUM prune path of exported files by NUM elements -x PATTERN extract files matching pattern (e.g. "*:**" for all) List or export theory exports for SESSION: named blobs produced by isabelle build. Option -l or -x is required; option -x may be repeated. The PATTERN language resembles glob patterns in the shell, with ? and * (both excluding ":" and "/"), ** (excluding ":"), and [abc] or [^abc], and variants {pattern1,pattern2,pattern3}.\} \<^medskip> The specified session is updated via @{tool build} (\secref{sec:tool-build}), with the same options \<^verbatim>\-d\, \<^verbatim>\-o\. The option \<^verbatim>\-n\ suppresses the implicit build process: it means that a potentially outdated session database is used! \<^medskip> Option \<^verbatim>\-l\ lists all stored exports, with compound names \theory\\<^verbatim>\:\\name\. \<^medskip> Option \<^verbatim>\-x\ extracts stored exports whose compound name matches the given pattern. Note that wild cards ``\<^verbatim>\?\'' and ``\<^verbatim>\*\'' do not match the separators ``\<^verbatim>\:\'' and ``\<^verbatim>\/\''; the wild card \<^verbatim>\**\ matches over directory name hierarchies separated by ``\<^verbatim>\/\''. Thus the pattern ``\<^verbatim>\*:**\'' matches \<^emph>\all\ theory exports. Multiple options \<^verbatim>\-x\ refer to the union of all specified patterns. Option \<^verbatim>\-O\ specifies an alternative output directory for option \<^verbatim>\-x\: the default is \<^verbatim>\export\ within the current directory. Each theory creates its own sub-directory hierarchy, using the session-qualified theory name. Option \<^verbatim>\-p\ specifies the number of elements that should be pruned from each name: it allows to reduce the resulting directory hierarchy at the danger of overwriting files due to loss of uniqueness. \ section \Dump PIDE session database \label{sec:tool-dump}\ text \ The @{tool_def "dump"} tool dumps information from the cumulative PIDE session database (which is processed on the spot). Its command-line usage is: @{verbatim [display] \Usage: isabelle dump [OPTIONS] [SESSIONS ...] Options are: -A NAMES dump named aspects (default: ...) -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -O DIR output directory for dumped files (default: "dump") -R refer to requirements of selected sessions -X NAME exclude sessions from group NAME and all descendants -a select all sessions -b NAME base logic image (default "Pure") -d DIR include session directory -g NAME select session group NAME -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose -x NAME exclude session NAME and all descendants Dump cumulative PIDE session database, with the following aspects: ...\} \<^medskip> Options \<^verbatim>\-B\, \<^verbatim>\-D\, \<^verbatim>\-R\, \<^verbatim>\-X\, \<^verbatim>\-a\, \<^verbatim>\-d\, \<^verbatim>\-g\, \<^verbatim>\-x\ and the remaining command-line arguments specify sessions as in @{tool build} (\secref{sec:tool-build}): the cumulative PIDE database of all their loaded theories is dumped to the output directory of option \<^verbatim>\-O\ (default: \<^verbatim>\dump\ in the current directory). \<^medskip> Option \<^verbatim>\-b\ specifies an optional base logic image, for improved scalability of the PIDE session. Its theories are only processed if it is included in the overall session selection. \<^medskip> Option \<^verbatim>\-o\ overrides Isabelle system options as for @{tool build} (\secref{sec:tool-build}). \<^medskip> Option \<^verbatim>\-v\ increases the general level of verbosity. \<^medskip> Option \<^verbatim>\-A\ specifies named aspects of the dump, as a comma-separated list. The default is to dump all known aspects, as given in the command-line usage of the tool. The underlying Isabelle/Scala operation \<^scala_method>\isabelle.Dump.dump\ takes aspects as user-defined operations on the final PIDE state and document version. This allows to imitate Prover IDE rendering under program control. \ subsubsection \Examples\ text \ Dump all Isabelle/ZF sessions (which are rather small): @{verbatim [display] \ isabelle dump -v -B ZF\} \<^smallskip> Dump the quite substantial \<^verbatim>\HOL-Analysis\ session, with full bootstrap from Isabelle/Pure: @{verbatim [display] \ isabelle dump -v HOL-Analysis\} \<^smallskip> Dump all sessions connected to HOL-Analysis, using main Isabelle/HOL as basis: @{verbatim [display] \ isabelle dump -v -b HOL -B HOL-Analysis\} This results in uniform PIDE markup for everything, except for the Isabelle/Pure bootstrap process itself. Producing that on the spot requires several GB of heap space, both for the Isabelle/Scala and Isabelle/ML process (in 64bit mode). Here are some relevant settings (\secref{sec:boot}) for such ambitious applications: @{verbatim [display] \ ISABELLE_TOOL_JAVA_OPTIONS="-Xms4g -Xmx32g -Xss16m" ML_OPTIONS="--minheap 4G --maxheap 32G" \} \ section \Update theory sources based on PIDE markup \label{sec:tool-update}\ text \ The @{tool_def "update"} tool updates theory sources based on markup that is produced by the regular @{tool build} process \secref{sec:tool-build}). Its command-line usage is: @{verbatim [display] \Usage: isabelle update [OPTIONS] [SESSIONS ...] Options are: -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -R refer to requirements of selected sessions -X NAME exclude sessions from group NAME and all descendants -a select all sessions -b build heap images -c clean build -d DIR include session directory -f fresh build -g NAME select session group NAME -j INT maximum number of parallel jobs (default 1) -l NAME additional base logic -n no build -- take existing build databases -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -u OPT override "update" option for selected sessions -v verbose -x NAME exclude session NAME and all descendants Update theory sources based on PIDE markup produced by "isabelle build".\} \<^medskip> Most options are the same as for @{tool build} (\secref{sec:tool-build}). \<^medskip> Option \<^verbatim>\-l\ specifies one or more base logics: these sessions and their ancestors are \<^emph>\excluded\ from the update. \<^medskip> Option \<^verbatim>\-u\ refers to specific \<^verbatim>\update\ options, by relying on naming convention: ``\<^verbatim>\-u\~\OPT\'' is a shortcut for ``\<^verbatim>\-o\~\<^verbatim>\update_\\OPT\''. \<^medskip> The following \<^verbatim>\update\ options are supported: \<^item> @{system_option_ref update_inner_syntax_cartouches} to update inner syntax (types, terms, etc.)~to use cartouches, instead of double-quoted strings or atomic identifiers. For example, ``\<^theory_text>\lemma \x = x\\'' is replaced by ``\<^theory_text>\lemma \x = x\\'', and ``\<^theory_text>\assume A\'' is replaced by ``\<^theory_text>\assume \A\\''. \<^item> @{system_option update_mixfix_cartouches} to update mixfix templates to use cartouches instead of double-quoted strings. For example, ``\<^theory_text>\(infixl \+\ 65)\'' is replaced by ``\<^theory_text>\(infixl \+\ 65)\''. \<^item> @{system_option_ref update_control_cartouches} to update antiquotations to use the compact form with control symbol and cartouche argument. For example, ``\@{term \x + y\}\'' is replaced by ``\\<^term>\x + y\\'' (the control symbol is literally \<^verbatim>\\<^term>\.) \<^item> @{system_option_ref update_path_cartouches} to update file-system paths to use cartouches: this depends on language markup provided by semantic processing of parsed input. \<^item> @{system_option_ref update_cite} to update {\LaTeX} \<^verbatim>\\cite\ commands and old-style \<^verbatim>\@{cite "name"}\ document antiquotations. It is also possible to produce custom updates in Isabelle/ML, by reporting \<^ML>\Markup.update\ with the precise source position and a replacement text. This operation should be made conditional on specific system options, similar to the ones above. Searching the above option names in ML sources of \<^dir>\$ISABELLE_HOME/src/Pure\ provides some examples. Updates can be in conflict by producing nested or overlapping edits: this may require to run @{tool update} multiple times. \ subsubsection \Examples\ text \ Update some cartouche notation in all theory sources required for session \<^verbatim>\HOL-Analysis\ (and ancestors): @{verbatim [display] \ isabelle update -u mixfix_cartouches HOL-Analysis\} \<^smallskip> Update the same for all application sessions based on \<^verbatim>\HOL-Analysis\, but do not change the underlying \<^verbatim>\HOL\ (and \<^verbatim>\Pure\) session: @{verbatim [display] \ isabelle update -u mixfix_cartouches -l HOL -B HOL-Analysis\} \<^smallskip> Update all sessions that happen to be properly built beforehand: @{verbatim [display] \ isabelle update -u mixfix_cartouches -n -a\} \ section \Explore sessions structure\ text \ The @{tool_def "sessions"} tool explores the sessions structure. Its command-line usage is: @{verbatim [display] \Usage: isabelle sessions [OPTIONS] [SESSIONS ...] Options are: -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -R refer to requirements of selected sessions -X NAME exclude sessions from group NAME and all descendants -a select all sessions -b follow session build dependencies (default: source imports) -d DIR include session directory -g NAME select session group NAME -x NAME exclude session NAME and all descendants Explore the structure of Isabelle sessions and print result names in topological order (on stdout).\} Arguments and options for session selection resemble @{tool build} (\secref{sec:tool-build}). \ subsubsection \Examples\ text \ All sessions of the Isabelle distribution: @{verbatim [display] \ isabelle sessions -a\} \<^medskip> Sessions that are imported by \<^verbatim>\ZF\: @{verbatim [display] \ isabelle sessions ZF\} \<^medskip> Sessions that are required to build \<^verbatim>\ZF\: @{verbatim [display] \ isabelle sessions -b ZF\} \<^medskip> Sessions that are based on \<^verbatim>\ZF\ (and imported by it): @{verbatim [display] \ isabelle sessions -B ZF\} \<^medskip> All sessions of Isabelle/AFP (based in directory \<^path>\AFP\): @{verbatim [display] \ isabelle sessions -D AFP/thys\} \<^medskip> Sessions required by Isabelle/AFP (based in directory \<^path>\AFP\): @{verbatim [display] \ isabelle sessions -R -D AFP/thys\} \ section \Synchronize source repositories and session images for Isabelle and AFP\ text \ The @{tool_def sync} tool synchronizes a local Isabelle and AFP source repository, possibly with prebuilt \<^verbatim>\.jar\ files and session images. Its command-line usage is: @{verbatim [display] \Usage: isabelle sync [OPTIONS] TARGET Options are: -A ROOT include AFP with given root directory (":" for $AFP_BASE) -H purge heaps directory on target -I NAME include session heap image and build database (based on accidental local state) -J preserve *.jar files -P protect spaces in target file names: more robust, less portable -S PATH SSH control path for connection multiplexing -T thorough treatment of file content and directory times -a REV explicit AFP revision (default: state of working directory) -n no changes: dry-run -p PORT SSH port -r REV explicit revision (default: state of working directory) -v verbose Synchronize Isabelle + AFP repositories, based on "isabelle hg_sync".\} The approach is to apply @{tool hg_sync} (see \secref{sec:tool-hg-sync}) on the underlying Isabelle repository, and an optional AFP repository. Consequently, the Isabelle installation needs to be a Mercurial repository clone: a regular download of the Isabelle distribution is not sufficient! On the target side, AFP is placed into @{setting ISABELLE_HOME} as immediate sub-directory with the literal name \<^verbatim>\AFP\; thus it can be easily included elsewhere, e.g. @{tool build}~\<^verbatim>\-d\~\<^verbatim>\'~~/AFP'\ on the remote side. \<^medskip> Options \<^verbatim>\-P\, \<^verbatim>\-S\, \<^verbatim>\-T\, \<^verbatim>\-n\, \<^verbatim>\-p\, \<^verbatim>\-v\ are the same as the underlying @{tool hg_sync}. \<^medskip> Options \<^verbatim>\-r\ and \<^verbatim>\-a\ are the same as option \<^verbatim>\-r\ for @{tool hg_sync}, but for the Isabelle and AFP repositories, respectively. The AFP version is only used if a corresponding repository is given via option \<^verbatim>\-A\, either with explicit root directory, or as \<^verbatim>\-A:\ to refer to \<^verbatim>\$AFP_BASE\ (this assumes AFP as component of the local Isabelle installation). If no AFP repository is given, an existing \<^verbatim>\AFP\ directory on the target remains unchanged. \<^medskip> Option \<^verbatim>\-J\ uploads existing \<^verbatim>\.jar\ files from the working directory, which are usually Isabelle/Scala/Java modules under control of @{tool scala_build} via \<^verbatim>\etc/build.props\ (see also \secref{sec:scala-build}). Thus the dependency management is accurate: bad uploads will be rebuilt eventually (or ignored). This might fail for very old Isabelle versions, when going into the past via option \<^verbatim>\-r\: here it is better to omit option \<^verbatim>\-J\ and thus purge \<^verbatim>\.jar\ files on the target (because they do not belong to the repository). \<^medskip> Option \<^verbatim>\-I\ uploads a collection of session images. The set of \<^verbatim>\-I\ options specifies the end-points in the session build graph, including all required ancestors. The result collection is uploaded using the underlying \<^verbatim>\rsync\ policies, so unchanged images are not sent again. Session images are assembled within the target \<^verbatim>\heaps\ directory: this scheme fits together with @{tool build}~\<^verbatim>\-o system_heaps\. Images are taken as-is from the local Isabelle installation, regardless of option \<^verbatim>\-r\. Upload of bad images could waste time and space, but running e.g. @{tool build} on the target will check dependencies accurately and rebuild outdated images on demand. \<^medskip> Option \<^verbatim>\-H\ tells the underlying \<^verbatim>\rsync\ process to purge the \<^verbatim>\heaps\ directory on the target, before uploading new images via option \<^verbatim>\-I\. The default is to work monotonically: old material that is not overwritten remains unchanged. Over time, this may lead to unused garbage, due to changes in session names or the Poly/ML version. Option \<^verbatim>\-H\ helps to avoid wasting file-system space. \ subsubsection \Examples\ text \ For quick testing of Isabelle + AFP on a remote machine, upload changed sources, jars, and local sessions images for \<^verbatim>\HOL\: @{verbatim [display] \ isabelle sync -A: -I HOL -J testmachine:test/isabelle_afp\} Assuming that the local \<^verbatim>\HOL\ hierarchy has been up-to-date, and the local and remote ML platforms coincide, a remote @{tool build} will proceed without building \<^verbatim>\HOL\ again. \<^medskip> Here is a variation for extra-clean testing of Isabelle + AFP: no option \<^verbatim>\-J\, but option \<^verbatim>\-T\ to disable the default ``quick check'' of \<^verbatim>\rsync\ (which only inspects file sizes and date stamps); existing heaps are deleted: @{verbatim [display] \ isabelle sync -A: -T -H testmachine:test/isabelle_afp\} \ end diff --git a/src/Doc/Tutorial/Advanced/simp2.thy b/src/Doc/Tutorial/Advanced/simp2.thy --- a/src/Doc/Tutorial/Advanced/simp2.thy +++ b/src/Doc/Tutorial/Advanced/simp2.thy @@ -1,188 +1,188 @@ (*<*) theory simp2 imports Main begin (*>*) section\Simplification\ text\\label{sec:simplification-II}\index{simplification|(} This section describes features not covered until now. It also outlines the simplification process itself, which can be helpful when the simplifier does not do what you expect of it. \ subsection\Advanced Features\ subsubsection\Congruence Rules\ text\\label{sec:simp-cong} While simplifying the conclusion $Q$ of $P \Imp Q$, it is legal to use the assumption $P$. For $\Imp$ this policy is hardwired, but contextual information can also be made available for other operators. For example, \<^prop>\xs = [] --> xs@xs = xs\ simplifies to \<^term>\True\ because we may use \<^prop>\xs = []\ when simplifying \<^prop>\xs@xs = xs\. The generation of contextual information during simplification is controlled by so-called \bfindex{congruence rules}. This is the one for \\\: @{thm[display]imp_cong[no_vars]} It should be read as follows: In order to simplify \<^prop>\P-->Q\ to \<^prop>\P'-->Q'\, simplify \<^prop>\P\ to \<^prop>\P'\ and assume \<^prop>\P'\ when simplifying \<^prop>\Q\ to \<^prop>\Q'\. Here are some more examples. The congruence rules for bounded quantifiers supply contextual information about the bound variable: @{thm[display,eta_contract=false,margin=60]ball_cong[no_vars]} One congruence rule for conditional expressions supplies contextual information for simplifying the \then\ and \else\ cases: @{thm[display]if_cong[no_vars]} An alternative congruence rule for conditional expressions actually \emph{prevents} simplification of some arguments: @{thm[display]if_weak_cong[no_vars]} Only the first argument is simplified; the others remain unchanged. This makes simplification much faster and is faithful to the evaluation strategy in programming languages, which is why this is the default congruence rule for \if\. Analogous rules control the evaluation of \case\ expressions. You can declare your own congruence rules with the attribute \attrdx{cong}, either globally, in the usual manner, \begin{quote} \isacommand{declare} \textit{theorem-name} \[cong]\ \end{quote} or locally in a \simp\ call by adding the modifier \begin{quote} \cong:\ \textit{list of theorem names} \end{quote} The effect is reversed by \cong del\ instead of \cong\. \begin{warn} The congruence rule @{thm[source]conj_cong} @{thm[display]conj_cong[no_vars]} \par\noindent is occasionally useful but is not a default rule; you have to declare it explicitly. \end{warn} \ subsubsection\Permutative Rewrite Rules\ text\ \index{rewrite rules!permutative|bold}% An equation is a \textbf{permutative rewrite rule} if the left-hand side and right-hand side are the same up to renaming of variables. The most common permutative rule is commutativity: \<^prop>\x+y = y+x\. Other examples include \<^prop>\(x-y)-z = (x-z)-y\ in arithmetic and \<^prop>\insert x (insert y A) = insert y (insert x A)\ for sets. Such rules are problematic because once they apply, they can be used forever. The simplifier is aware of this danger and treats permutative rules by means of a special strategy, called \bfindex{ordered rewriting}: a permutative rewrite rule is only applied if the term becomes smaller with respect to a fixed lexicographic ordering on terms. For example, commutativity rewrites \<^term>\b+a\ to \<^term>\a+b\, but then stops because \<^term>\a+b\ is strictly smaller than \<^term>\b+a\. Permutative rewrite rules can be turned into simplification rules in the usual manner via the \simp\ attribute; the simplifier recognizes their special status automatically. Permutative rewrite rules are most effective in the case of associative-commutative functions. (Associativity by itself is not permutative.) When dealing with an AC-function~$f$, keep the following points in mind: \begin{itemize}\index{associative-commutative function} \item The associative law must always be oriented from left to right, namely $f(f(x,y),z) = f(x,f(y,z))$. The opposite orientation, if used with commutativity, can lead to nontermination. \item To complete your set of rewrite rules, you must add not just associativity~(A) and commutativity~(C) but also a derived rule, {\bf left-com\-mut\-ativ\-ity} (LC): $f(x,f(y,z)) = f(y,f(x,z))$. \end{itemize} Ordered rewriting with the combination of A, C, and LC sorts a term lexicographically: \[\def\maps#1{~\stackrel{#1}{\leadsto}~} f(f(b,c),a) \maps{A} f(b,f(c,a)) \maps{C} f(b,f(a,c)) \maps{LC} f(a,f(b,c)) \] Note that ordered rewriting for \+\ and \*\ on numbers is rarely necessary because the built-in arithmetic prover often succeeds without such tricks. \ subsection\How the Simplifier Works\ text\\label{sec:SimpHow} Roughly speaking, the simplifier proceeds bottom-up: subterms are simplified first. A conditional equation is only applied if its condition can be proved, again by simplification. Below we explain some special features of the rewriting process. \ subsubsection\Higher-Order Patterns\ text\\index{simplification rule|(} So far we have pretended the simplifier can deal with arbitrary rewrite rules. This is not quite true. For reasons of feasibility, the simplifier expects the left-hand side of each rule to be a so-called \emph{higher-order -pattern}~@{cite "nipkow-patterns"}\indexbold{patterns!higher-order}. +pattern}~\<^cite>\"nipkow-patterns"\\indexbold{patterns!higher-order}. This restricts where unknowns may occur. Higher-order patterns are terms in $\beta$-normal form. (This means there are no subterms of the form $(\lambda x. M)(N)$.) Each occurrence of an unknown is of the form $\Var{f}~x@1~\dots~x@n$, where the $x@i$ are distinct bound variables. Thus all ordinary rewrite rules, where all unknowns are of base type, for example @{thm add.assoc}, are acceptable: if an unknown is of base type, it cannot have any arguments. Additionally, the rule \(\x. ?P x \ ?Q x) = ((\x. ?P x) \ (\x. ?Q x))\ is also acceptable, in both directions: all arguments of the unknowns \?P\ and \?Q\ are distinct bound variables. If the left-hand side is not a higher-order pattern, all is not lost. The simplifier will still try to apply the rule provided it matches directly: without much $\lambda$-calculus hocus pocus. For example, \(?f ?x \ range ?f) = True\ rewrites \<^term>\g a \ range g\ to \<^const>\True\, but will fail to match \g(h b) \ range(\x. g(h x))\. However, you can eliminate the offending subterms --- those that are not patterns --- by adding new variables and conditions. In our example, we eliminate \?f ?x\ and obtain \?y = ?f ?x \ (?y \ range ?f) = True\, which is fine as a conditional rewrite rule since conditions can be arbitrary terms. However, this trick is not a panacea because the newly introduced conditions may be hard to solve. There is no restriction on the form of the right-hand sides. They may not contain extraneous term or type variables, though. \ subsubsection\The Preprocessor\ text\\label{sec:simp-preprocessor} When a theorem is declared a simplification rule, it need not be a conditional equation already. The simplifier will turn it into a set of conditional equations automatically. For example, \<^prop>\f x = g x & h x = k x\ becomes the two separate simplification rules \<^prop>\f x = g x\ and \<^prop>\h x = k x\. In general, the input theorem is converted as follows: \begin{eqnarray} \neg P &\mapsto& P = \hbox{\isa{False}} \nonumber\\ P \longrightarrow Q &\mapsto& P \Longrightarrow Q \nonumber\\ P \land Q &\mapsto& P,\ Q \nonumber\\ \forall x.~P~x &\mapsto& P~\Var{x}\nonumber\\ \forall x \in A.\ P~x &\mapsto& \Var{x} \in A \Longrightarrow P~\Var{x} \nonumber\\ \if\\ P\ \then\\ Q\ \else\\ R &\mapsto& P \Longrightarrow Q,\ \neg P \Longrightarrow R \nonumber \end{eqnarray} Once this conversion process is finished, all remaining non-equations $P$ are turned into trivial equations $P =\isa{True}$. For example, the formula \begin{center}\<^prop>\(p \ t=u \ ~r) \ s\\end{center} is converted into the three rules \begin{center} \<^prop>\p \ t = u\,\quad \<^prop>\p \ r = False\,\quad \<^prop>\s = True\. \end{center} \index{simplification rule|)} \index{simplification|)} \ (*<*) end (*>*) diff --git a/src/Doc/Tutorial/CTL/Base.thy b/src/Doc/Tutorial/CTL/Base.thy --- a/src/Doc/Tutorial/CTL/Base.thy +++ b/src/Doc/Tutorial/CTL/Base.thy @@ -1,91 +1,91 @@ (*<*)theory Base imports Main begin(*>*) section\Case Study: Verified Model Checking\ text\\label{sec:VMC} This chapter ends with a case study concerning model checking for Computation Tree Logic (CTL), a temporal logic. Model checking is a popular technique for the verification of finite state systems (implementations) with respect to temporal logic formulae -(specifications) @{cite "ClarkeGP-book" and "Huth-Ryan-book"}. Its foundations are set theoretic +(specifications) \<^cite>\"ClarkeGP-book" and "Huth-Ryan-book"\. Its foundations are set theoretic and this section will explore them in HOL\@. This is done in two steps. First we consider a simple modal logic called propositional dynamic logic (PDL)\@. We then proceed to the temporal logic CTL, which is used in many real model checkers. In each case we give both a traditional semantics (\\\) and a recursive function \<^term>\mc\ that maps a formula into the set of all states of the system where the formula is valid. If the system has a finite number of states, \<^term>\mc\ is directly executable: it is a model checker, albeit an inefficient one. The main proof obligation is to show that the semantics and the model checker agree. \underscoreon Our models are \emph{transition systems}:\index{transition systems} sets of \emph{states} with transitions between them. Here is a simple example: \begin{center} \unitlength.5mm \thicklines \begin{picture}(100,60) \put(50,50){\circle{20}} \put(50,50){\makebox(0,0){$p,q$}} \put(61,55){\makebox(0,0)[l]{$s_0$}} \put(44,42){\vector(-1,-1){26}} \put(16,18){\vector(1,1){26}} \put(57,43){\vector(1,-1){26}} \put(10,10){\circle{20}} \put(10,10){\makebox(0,0){$q,r$}} \put(-1,15){\makebox(0,0)[r]{$s_1$}} \put(20,10){\vector(1,0){60}} \put(90,10){\circle{20}} \put(90,10){\makebox(0,0){$r$}} \put(98, 5){\line(1,0){10}} \put(108, 5){\line(0,1){10}} \put(108,15){\vector(-1,0){10}} \put(91,21){\makebox(0,0)[bl]{$s_2$}} \end{picture} \end{center} Each state has a unique name or number ($s_0,s_1,s_2$), and in each state certain \emph{atomic propositions} ($p,q,r$) hold. The aim of temporal logic is to formalize statements such as ``there is no path starting from $s_2$ leading to a state where $p$ or $q$ holds,'' which is true, and ``on all paths starting from $s_0$, $q$ always holds,'' which is false. Abstracting from this concrete example, we assume there is a type of states: \ typedecl state text\\noindent Command \commdx{typedecl} merely declares a new type but without defining it (see \S\ref{sec:typedecl}). Thus we know nothing about the type other than its existence. That is exactly what we need because \<^typ>\state\ really is an implicit parameter of our model. Of course it would have been more generic to make \<^typ>\state\ a type parameter of everything but declaring \<^typ>\state\ globally as above reduces clutter. Similarly we declare an arbitrary but fixed transition system, i.e.\ a relation between states: \ consts M :: "(state \ state)set" text\\noindent This is Isabelle's way of declaring a constant without defining it. Finally we introduce a type of atomic propositions \ typedecl "atom" text\\noindent and a \emph{labelling function} \ consts L :: "state \ atom set" text\\noindent telling us which atomic propositions are true in each state. \ (*<*)end(*>*) diff --git a/src/Doc/Tutorial/CTL/CTL.thy b/src/Doc/Tutorial/CTL/CTL.thy --- a/src/Doc/Tutorial/CTL/CTL.thy +++ b/src/Doc/Tutorial/CTL/CTL.thy @@ -1,449 +1,449 @@ (*<*)theory CTL imports Base begin(*>*) subsection\Computation Tree Logic --- CTL\ text\\label{sec:CTL} \index{CTL|(}% The semantics of PDL only needs reflexive transitive closure. Let us be adventurous and introduce a more expressive temporal operator. We extend the datatype \formula\ by a new constructor \ (*<*) datatype formula = Atom "atom" | Neg formula | And formula formula | AX formula | EF formula(*>*) | AF formula text\\noindent which stands for ``\emph{A}lways in the \emph{F}uture'': on all infinite paths, at some point the formula holds. Formalizing the notion of an infinite path is easy in HOL: it is simply a function from \<^typ>\nat\ to \<^typ>\state\. \ definition Paths :: "state \ (nat \ state)set" where "Paths s \ {p. s = p 0 \ (\i. (p i, p(i+1)) \ M)}" text\\noindent This definition allows a succinct statement of the semantics of \<^const>\AF\: \footnote{Do not be misled: neither datatypes nor recursive functions can be extended by new constructors or equations. This is just a trick of the presentation (see \S\ref{sec:doc-prep-suppress}). In reality one has to define a new datatype and a new function.} \ (*<*) primrec valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) where "s \ Atom a = (a \ L s)" | "s \ Neg f = (~(s \ f))" | "s \ And f g = (s \ f \ s \ g)" | "s \ AX f = (\t. (s,t) \ M \ t \ f)" | "s \ EF f = (\t. (s,t) \ M\<^sup>* \ t \ f)" | (*>*) "s \ AF f = (\p \ Paths s. \i. p i \ f)" text\\noindent Model checking \<^const>\AF\ involves a function which is just complicated enough to warrant a separate definition: \ definition af :: "state set \ state set \ state set" where "af A T \ A \ {s. \t. (s, t) \ M \ t \ T}" text\\noindent Now we define \<^term>\mc(AF f)\ as the least set \<^term>\T\ that includes \<^term>\mc f\ and all states all of whose direct successors are in \<^term>\T\: \ (*<*) primrec mc :: "formula \ state set" where "mc(Atom a) = {s. a \ L s}" | "mc(Neg f) = -mc f" | "mc(And f g) = mc f \ mc g" | "mc(AX f) = {s. \t. (s,t) \ M \ t \ mc f}" | "mc(EF f) = lfp(\T. mc f \ M\ `` T)"|(*>*) "mc(AF f) = lfp(af(mc f))" text\\noindent Because \<^const>\af\ is monotone in its second argument (and also its first, but that is irrelevant), \<^term>\af A\ has a least fixed point: \ lemma mono_af: "mono(af A)" apply(simp add: mono_def af_def) apply blast done (*<*) lemma mono_ef: "mono(\T. A \ M\ `` T)" apply(rule monoI) by(blast) lemma EF_lemma: "lfp(\T. A \ M\ `` T) = {s. \t. (s,t) \ M\<^sup>* \ t \ A}" apply(rule equalityI) apply(rule subsetI) apply(simp) apply(erule lfp_induct_set) apply(rule mono_ef) apply(simp) apply(blast intro: rtrancl_trans) apply(rule subsetI) apply(simp, clarify) apply(erule converse_rtrancl_induct) apply(subst lfp_unfold[OF mono_ef]) apply(blast) apply(subst lfp_unfold[OF mono_ef]) by(blast) (*>*) text\ All we need to prove now is \<^prop>\mc(AF f) = {s. s \ AF f}\, which states that \<^term>\mc\ and \\\ agree for \<^const>\AF\\@. This time we prove the two inclusions separately, starting with the easy one: \ theorem AF_lemma1: "lfp(af A) \ {s. \p \ Paths s. \i. p i \ A}" txt\\noindent In contrast to the analogous proof for \<^const>\EF\, and just for a change, we do not use fixed point induction. Park-induction, named after David Park, is weaker but sufficient for this proof: \begin{center} @{thm lfp_lowerbound[of _ "S",no_vars]} \hfill (@{thm[source]lfp_lowerbound}) \end{center} The instance of the premise \<^prop>\f S \ S\ is proved pointwise, a decision that \isa{auto} takes for us: \ apply(rule lfp_lowerbound) apply(auto simp add: af_def Paths_def) txt\ @{subgoals[display,indent=0,margin=70,goals_limit=1]} In this remaining case, we set \<^term>\t\ to \<^term>\p(1::nat)\. The rest is automatic, which is surprising because it involves finding the instantiation \<^term>\\i::nat. p(i+1)\ for \\p\. \ apply(erule_tac x = "p 1" in allE) apply(auto) done text\ The opposite inclusion is proved by contradiction: if some state \<^term>\s\ is not in \<^term>\lfp(af A)\, then we can construct an infinite \<^term>\A\-avoiding path starting from~\<^term>\s\. The reason is that by unfolding \<^const>\lfp\ we find that if \<^term>\s\ is not in \<^term>\lfp(af A)\, then \<^term>\s\ is not in \<^term>\A\ and there is a direct successor of \<^term>\s\ that is again not in \mbox{\<^term>\lfp(af A)\}. Iterating this argument yields the promised infinite \<^term>\A\-avoiding path. Let us formalize this sketch. The one-step argument in the sketch above is proved by a variant of contraposition: \ lemma not_in_lfp_afD: "s \ lfp(af A) \ s \ A \ (\ t. (s,t) \ M \ t \ lfp(af A))" apply(erule contrapos_np) apply(subst lfp_unfold[OF mono_af]) apply(simp add: af_def) done text\\noindent We assume the negation of the conclusion and prove \<^term>\s \ lfp(af A)\. Unfolding \<^const>\lfp\ once and simplifying with the definition of \<^const>\af\ finishes the proof. Now we iterate this process. The following construction of the desired path is parameterized by a predicate \<^term>\Q\ that should hold along the path: \ primrec path :: "state \ (state \ bool) \ (nat \ state)" where "path s Q 0 = s" | "path s Q (Suc n) = (SOME t. (path s Q n,t) \ M \ Q t)" text\\noindent Element \<^term>\n+1::nat\ on this path is some arbitrary successor \<^term>\t\ of element \<^term>\n\ such that \<^term>\Q t\ holds. Remember that \SOME t. R t\ is some arbitrary but fixed \<^term>\t\ such that \<^prop>\R t\ holds (see \S\ref{sec:SOME}). Of course, such a \<^term>\t\ need not exist, but that is of no concern to us since we will only use \<^const>\path\ when a suitable \<^term>\t\ does exist. Let us show that if each state \<^term>\s\ that satisfies \<^term>\Q\ has a successor that again satisfies \<^term>\Q\, then there exists an infinite \<^term>\Q\-path: \ lemma infinity_lemma: "\ Q s; \s. Q s \ (\ t. (s,t) \ M \ Q t) \ \ \p\Paths s. \i. Q(p i)" txt\\noindent First we rephrase the conclusion slightly because we need to prove simultaneously both the path property and the fact that \<^term>\Q\ holds: \ apply(subgoal_tac "\p. s = p 0 \ (\i::nat. (p i, p(i+1)) \ M \ Q(p i))") txt\\noindent From this proposition the original goal follows easily: \ apply(simp add: Paths_def, blast) txt\\noindent The new subgoal is proved by providing the witness \<^term>\path s Q\ for \<^term>\p\: \ apply(rule_tac x = "path s Q" in exI) apply(clarsimp) txt\\noindent After simplification and clarification, the subgoal has the following form: @{subgoals[display,indent=0,margin=70,goals_limit=1]} It invites a proof by induction on \<^term>\i\: \ apply(induct_tac i) apply(simp) txt\\noindent After simplification, the base case boils down to @{subgoals[display,indent=0,margin=70,goals_limit=1]} The conclusion looks exceedingly trivial: after all, \<^term>\t\ is chosen such that \<^prop>\(s,t)\M\ holds. However, we first have to show that such a \<^term>\t\ actually exists! This reasoning is embodied in the theorem @{thm[source]someI2_ex}: @{thm[display,eta_contract=false]someI2_ex} When we apply this theorem as an introduction rule, \?P x\ becomes \<^prop>\(s, x) \ M \ Q x\ and \?Q x\ becomes \<^prop>\(s,x) \ M\ and we have to prove two subgoals: \<^prop>\\a. (s, a) \ M \ Q a\, which follows from the assumptions, and \<^prop>\(s, x) \ M \ Q x \ (s,x) \ M\, which is trivial. Thus it is not surprising that \fast\ can prove the base case quickly: \ apply(fast intro: someI2_ex) txt\\noindent What is worth noting here is that we have used \methdx{fast} rather than \blast\. The reason is that \blast\ would fail because it cannot cope with @{thm[source]someI2_ex}: unifying its conclusion with the current subgoal is non-trivial because of the nested schematic variables. For efficiency reasons \blast\ does not even attempt such unifications. Although \fast\ can in principle cope with complicated unification problems, in practice the number of unifiers arising is often prohibitive and the offending rule may need to be applied explicitly rather than automatically. This is what happens in the step case. The induction step is similar, but more involved, because now we face nested occurrences of \SOME\. As a result, \fast\ is no longer able to solve the subgoal and we apply @{thm[source]someI2_ex} by hand. We merely show the proof commands but do not describe the details: \ apply(simp) apply(rule someI2_ex) apply(blast) apply(rule someI2_ex) apply(blast) apply(blast) done text\ Function \<^const>\path\ has fulfilled its purpose now and can be forgotten. It was merely defined to provide the witness in the proof of the @{thm[source]infinity_lemma}. Aficionados of minimal proofs might like to know that we could have given the witness without having to define a new function: the term @{term[display]"rec_nat s (\n t. SOME u. (t,u)\M \ Q u)"} is extensionally equal to \<^term>\path s Q\, where \<^term>\rec_nat\ is the predefined primitive recursor on \<^typ>\nat\. \ (*<*) lemma "\ Q s; \ s. Q s \ (\ t. (s,t)\M \ Q t) \ \ \ p\Paths s. \ i. Q(p i)" apply(subgoal_tac "\ p. s = p 0 \ (\ i. (p i,p(Suc i))\M \ Q(p i))") apply(simp add: Paths_def) apply(blast) apply(rule_tac x = "rec_nat s (\n t. SOME u. (t,u)\M \ Q u)" in exI) apply(simp) apply(intro strip) apply(induct_tac i) apply(simp) apply(fast intro: someI2_ex) apply(simp) apply(rule someI2_ex) apply(blast) apply(rule someI2_ex) apply(blast) by(blast) (*>*) text\ At last we can prove the opposite direction of @{thm[source]AF_lemma1}: \ theorem AF_lemma2: "{s. \p \ Paths s. \i. p i \ A} \ lfp(af A)" txt\\noindent The proof is again pointwise and then by contraposition: \ apply(rule subsetI) apply(erule contrapos_pp) apply simp txt\ @{subgoals[display,indent=0,goals_limit=1]} Applying the @{thm[source]infinity_lemma} as a destruction rule leaves two subgoals, the second premise of @{thm[source]infinity_lemma} and the original subgoal: \ apply(drule infinity_lemma) txt\ @{subgoals[display,indent=0,margin=65]} Both are solved automatically: \ apply(auto dest: not_in_lfp_afD) done text\ If you find these proofs too complicated, we recommend that you read \S\ref{sec:CTL-revisited}, where we show how inductive definitions lead to simpler arguments. The main theorem is proved as for PDL, except that we also derive the necessary equality \lfp(af A) = ...\ by combining @{thm[source]AF_lemma1} and @{thm[source]AF_lemma2} on the spot: \ theorem "mc f = {s. s \ f}" apply(induct_tac f) apply(auto simp add: EF_lemma equalityI[OF AF_lemma1 AF_lemma2]) done text\ The language defined above is not quite CTL\@. The latter also includes an until-operator \<^term>\EU f g\ with semantics ``there \emph{E}xists a path where \<^term>\f\ is true \emph{U}ntil \<^term>\g\ becomes true''. We need an auxiliary function: \ primrec until:: "state set \ state set \ state \ state list \ bool" where "until A B s [] = (s \ B)" | "until A B s (t#p) = (s \ A \ (s,t) \ M \ until A B t p)" (*<*)definition eusem :: "state set \ state set \ state set" where "eusem A B \ {s. \p. until A B s p}"(*>*) text\\noindent Expressing the semantics of \<^term>\EU\ is now straightforward: @{prop[display]"s \ EU f g = (\p. until {t. t \ f} {t. t \ g} s p)"} Note that \<^term>\EU\ is not definable in terms of the other operators! Model checking \<^term>\EU\ is again a least fixed point construction: @{text[display]"mc(EU f g) = lfp(\T. mc g \ mc f \ (M\ `` T))"} \begin{exercise} Extend the datatype of formulae by the above until operator and prove the equivalence between semantics and model checking, i.e.\ that @{prop[display]"mc(EU f g) = {s. s \ EU f g}"} %For readability you may want to annotate {term EU} with its customary syntax %{text[display]"| EU formula formula E[_ U _]"} %which enables you to read and write {text"E[f U g]"} instead of {term"EU f g"}. \end{exercise} -For more CTL exercises see, for example, Huth and Ryan @{cite "Huth-Ryan-book"}. +For more CTL exercises see, for example, Huth and Ryan \<^cite>\"Huth-Ryan-book"\. \ (*<*) definition eufix :: "state set \ state set \ state set \ state set" where "eufix A B T \ B \ A \ (M\ `` T)" lemma "lfp(eufix A B) \ eusem A B" apply(rule lfp_lowerbound) apply(auto simp add: eusem_def eufix_def) apply(rule_tac x = "[]" in exI) apply simp apply(rule_tac x = "xa#xb" in exI) apply simp done lemma mono_eufix: "mono(eufix A B)" apply(simp add: mono_def eufix_def) apply blast done lemma "eusem A B \ lfp(eufix A B)" apply(clarsimp simp add: eusem_def) apply(erule rev_mp) apply(rule_tac x = x in spec) apply(induct_tac p) apply(subst lfp_unfold[OF mono_eufix]) apply(simp add: eufix_def) apply(clarsimp) apply(subst lfp_unfold[OF mono_eufix]) apply(simp add: eufix_def) apply blast done (* definition eusem :: "state set \ state set \ state set" where "eusem A B \ {s. \p\Paths s. \j. p j \ B \ (\i < j. p i \ A)}" axiomatization where M_total: "\t. (s,t) \ M" consts apath :: "state \ (nat \ state)" primrec "apath s 0 = s" "apath s (Suc i) = (SOME t. (apath s i,t) \ M)" lemma [iff]: "apath s \ Paths s"; apply(simp add: Paths_def); apply(blast intro: M_total[THEN someI_ex]) done definition pcons :: "state \ (nat \ state) \ (nat \ state)" where "pcons s p == \i. case i of 0 \ s | Suc j \ p j" lemma pcons_PathI: "[| (s,t) : M; p \ Paths t |] ==> pcons s p \ Paths s"; by(simp add: Paths_def pcons_def split: nat.split); lemma "lfp(eufix A B) \ eusem A B" apply(rule lfp_lowerbound) apply(clarsimp simp add: eusem_def eufix_def); apply(erule disjE); apply(rule_tac x = "apath x" in bexI); apply(rule_tac x = 0 in exI); apply simp; apply simp; apply(clarify); apply(rule_tac x = "pcons xb p" in bexI); apply(rule_tac x = "j+1" in exI); apply (simp add: pcons_def split: nat.split); apply (simp add: pcons_PathI) done *) (*>*) text\Let us close this section with a few words about the executability of our model checkers. It is clear that if all sets are finite, they can be represented as lists and the usual set operations are easily implemented. Only \<^const>\lfp\ requires a little thought. Fortunately, theory -\While_Combinator\ in the Library~@{cite "HOL-Library"} provides a +\While_Combinator\ in the Library~\<^cite>\"HOL-Library"\ provides a theorem stating that in the case of finite sets and a monotone function~\<^term>\F\, the value of \mbox{\<^term>\lfp F\} can be computed by iterated application of \<^term>\F\ to~\<^term>\{}\ until a fixed point is reached. It is actually possible to generate executable functional programs from HOL definitions, but that is beyond the scope of the tutorial.% \index{CTL|)}\ (*<*)end(*>*) diff --git a/src/Doc/Tutorial/CTL/PDL.thy b/src/Doc/Tutorial/CTL/PDL.thy --- a/src/Doc/Tutorial/CTL/PDL.thy +++ b/src/Doc/Tutorial/CTL/PDL.thy @@ -1,204 +1,204 @@ (*<*)theory PDL imports Base begin(*>*) subsection\Propositional Dynamic Logic --- PDL\ text\\index{PDL|(} The formulae of PDL are built up from atomic propositions via negation and conjunction and the two temporal connectives \AX\ and \EF\\@. Since formulae are essentially syntax trees, they are naturally modelled as a datatype:% \footnote{The customary definition of PDL -@{cite "HarelKT-DL"} looks quite different from ours, but the two are easily +\<^cite>\"HarelKT-DL"\ looks quite different from ours, but the two are easily shown to be equivalent.} \ datatype formula = Atom "atom" | Neg formula | And formula formula | AX formula | EF formula text\\noindent This resembles the boolean expression case study in \S\ref{sec:boolex}. A validity relation between states and formulae specifies the semantics. The syntax annotation allows us to write \s \ f\ instead of \hbox{\valid s f\}. The definition is by recursion over the syntax: \ primrec valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) where "s \ Atom a = (a \ L s)" | "s \ Neg f = (\(s \ f))" | "s \ And f g = (s \ f \ s \ g)" | "s \ AX f = (\t. (s,t) \ M \ t \ f)" | "s \ EF f = (\t. (s,t) \ M\<^sup>* \ t \ f)" text\\noindent The first three equations should be self-explanatory. The temporal formula \<^term>\AX f\ means that \<^term>\f\ is true in \emph{A}ll ne\emph{X}t states whereas \<^term>\EF f\ means that there \emph{E}xists some \emph{F}uture state in which \<^term>\f\ is true. The future is expressed via \\<^sup>*\, the reflexive transitive closure. Because of reflexivity, the future includes the present. Now we come to the model checker itself. It maps a formula into the set of states where the formula is true. It too is defined by recursion over the syntax:\ primrec mc :: "formula \ state set" where "mc(Atom a) = {s. a \ L s}" | "mc(Neg f) = -mc f" | "mc(And f g) = mc f \ mc g" | "mc(AX f) = {s. \t. (s,t) \ M \ t \ mc f}" | "mc(EF f) = lfp(\T. mc f \ (M\ `` T))" text\\noindent Only the equation for \<^term>\EF\ deserves some comments. Remember that the postfix \\\ and the infix \``\ are predefined and denote the converse of a relation and the image of a set under a relation. Thus \<^term>\M\ `` T\ is the set of all predecessors of \<^term>\T\ and the least fixed point (\<^term>\lfp\) of \<^term>\\T. mc f \ M\ `` T\ is the least set \<^term>\T\ containing \<^term>\mc f\ and all predecessors of \<^term>\T\. If you find it hard to see that \<^term>\mc(EF f)\ contains exactly those states from which there is a path to a state where \<^term>\f\ is true, do not worry --- this will be proved in a moment. First we prove monotonicity of the function inside \<^term>\lfp\ in order to make sure it really has a least fixed point. \ lemma mono_ef: "mono(\T. A \ (M\ `` T))" apply(rule monoI) apply blast done text\\noindent Now we can relate model checking and semantics. For the \EF\ case we need a separate lemma: \ lemma EF_lemma: "lfp(\T. A \ (M\ `` T)) = {s. \t. (s,t) \ M\<^sup>* \ t \ A}" txt\\noindent The equality is proved in the canonical fashion by proving that each set includes the other; the inclusion is shown pointwise: \ apply(rule equalityI) apply(rule subsetI) apply(simp)(*<*)apply(rename_tac s)(*>*) txt\\noindent Simplification leaves us with the following first subgoal @{subgoals[display,indent=0,goals_limit=1]} which is proved by \<^term>\lfp\-induction: \ apply(erule lfp_induct_set) apply(rule mono_ef) apply(simp) txt\\noindent Having disposed of the monotonicity subgoal, simplification leaves us with the following goal: \begin{isabelle} \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ A\ {\isasymor}\isanewline \ \ \ \ \ \ \ \ \ x\ {\isasymin}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharparenleft}lfp\ {\isacharparenleft}\dots{\isacharparenright}\ {\isasyminter}\ {\isacharbraceleft}x{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isacharparenright}\isanewline \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A \end{isabelle} It is proved by \blast\, using the transitivity of \isa{M\isactrlsup {\isacharasterisk}}. \ apply(blast intro: rtrancl_trans) txt\ We now return to the second set inclusion subgoal, which is again proved pointwise: \ apply(rule subsetI) apply(simp, clarify) txt\\noindent After simplification and clarification we are left with @{subgoals[display,indent=0,goals_limit=1]} This goal is proved by induction on \<^term>\(s,t)\M\<^sup>*\. But since the model checker works backwards (from \<^term>\t\ to \<^term>\s\), we cannot use the induction theorem @{thm[source]rtrancl_induct}: it works in the forward direction. Fortunately the converse induction theorem @{thm[source]converse_rtrancl_induct} already exists: @{thm[display,margin=60]converse_rtrancl_induct[no_vars]} It says that if \<^prop>\(a,b)\r\<^sup>*\ and we know \<^prop>\P b\ then we can infer \<^prop>\P a\ provided each step backwards from a predecessor \<^term>\z\ of \<^term>\b\ preserves \<^term>\P\. \ apply(erule converse_rtrancl_induct) txt\\noindent The base case @{subgoals[display,indent=0,goals_limit=1]} is solved by unrolling \<^term>\lfp\ once \ apply(subst lfp_unfold[OF mono_ef]) txt\ @{subgoals[display,indent=0,goals_limit=1]} and disposing of the resulting trivial subgoal automatically: \ apply(blast) txt\\noindent The proof of the induction step is identical to the one for the base case: \ apply(subst lfp_unfold[OF mono_ef]) apply(blast) done text\ The main theorem is proved in the familiar manner: induction followed by \auto\ augmented with the lemma as a simplification rule. \ theorem "mc f = {s. s \ f}" apply(induct_tac f) apply(auto simp add: EF_lemma) done text\ \begin{exercise} \<^term>\AX\ has a dual operator \<^term>\EN\ (``there exists a next state such that'')% \footnote{We cannot use the customary \EX\: it is reserved as the \textsc{ascii}-equivalent of \\\.} with the intended semantics @{prop[display]"(s \ EN f) = (\t. (s,t) \ M \ t \ f)"} Fortunately, \<^term>\EN f\ can already be expressed as a PDL formula. How? Show that the semantics for \<^term>\EF\ satisfies the following recursion equation: @{prop[display]"(s \ EF f) = (s \ f | s \ EN(EF f))"} \end{exercise} \index{PDL|)} \ (*<*) theorem main: "mc f = {s. s \ f}" apply(induct_tac f) apply(auto simp add: EF_lemma) done lemma aux: "s \ f = (s \ mc f)" apply(simp add: main) done lemma "(s \ EF f) = (s \ f | s \ Neg(AX(Neg(EF f))))" apply(simp only: aux) apply(simp) apply(subst lfp_unfold[OF mono_ef], fast) done end (*>*) diff --git a/src/Doc/Tutorial/Datatype/Nested.thy b/src/Doc/Tutorial/Datatype/Nested.thy --- a/src/Doc/Tutorial/Datatype/Nested.thy +++ b/src/Doc/Tutorial/Datatype/Nested.thy @@ -1,158 +1,158 @@ (*<*) theory Nested imports ABexpr begin (*>*) text\ \index{datatypes!and nested recursion}% So far, all datatypes had the property that on the right-hand side of their definition they occurred only at the top-level: directly below a constructor. Now we consider \emph{nested recursion}, where the recursive datatype occurs nested in some other datatype (but not inside itself!). Consider the following model of terms where function symbols can be applied to a list of arguments: \ (*<*)hide_const Var(*>*) datatype ('v,'f)"term" = Var 'v | App 'f "('v,'f)term list" text\\noindent Note that we need to quote \term\ on the left to avoid confusion with the Isabelle command \isacommand{term}. Parameter \<^typ>\'v\ is the type of variables and \<^typ>\'f\ the type of function symbols. A mathematical term like $f(x,g(y))$ becomes \<^term>\App f [Var x, App g [Var y]]\, where \<^term>\f\, \<^term>\g\, \<^term>\x\, \<^term>\y\ are suitable values, e.g.\ numbers or strings. What complicates the definition of \term\ is the nested occurrence of \term\ inside \list\ on the right-hand side. In principle, nested recursion can be eliminated in favour of mutual recursion by unfolding the offending datatypes, here \list\. The result for \term\ would be something like \medskip \input{unfoldnested.tex} \medskip \noindent Although we do not recommend this unfolding to the user, it shows how to simulate nested recursion by mutual recursion. Now we return to the initial definition of \term\ using nested recursion. Let us define a substitution function on terms. Because terms involve term lists, we need to define two substitution functions simultaneously: \ primrec subst :: "('v\('v,'f)term) \ ('v,'f)term \ ('v,'f)term" and substs:: "('v\('v,'f)term) \ ('v,'f)term list \ ('v,'f)term list" where "subst s (Var x) = s x" | subst_App: "subst s (App f ts) = App f (substs s ts)" | "substs s [] = []" | "substs s (t # ts) = subst s t # substs s ts" text\\noindent Individual equations in a \commdx{primrec} definition may be named as shown for @{thm[source]subst_App}. The significance of this device will become apparent below. Similarly, when proving a statement about terms inductively, we need to prove a related statement about term lists simultaneously. For example, the fact that the identity substitution does not change a term needs to be strengthened and proved as follows: \ lemma subst_id(*<*)(*referred to from ABexpr*)(*>*): "subst Var t = (t ::('v,'f)term) \ substs Var ts = (ts::('v,'f)term list)" apply(induct_tac t and ts rule: subst.induct substs.induct, simp_all) done text\\noindent Note that \<^term>\Var\ is the identity substitution because by definition it leaves variables unchanged: \<^prop>\subst Var (Var x) = Var x\. Note also that the type annotations are necessary because otherwise there is nothing in the goal to enforce that both halves of the goal talk about the same type parameters \('v,'f)\. As a result, induction would fail because the two halves of the goal would be unrelated. \begin{exercise} The fact that substitution distributes over composition can be expressed roughly as follows: @{text[display]"subst (f \ g) t = subst f (subst g t)"} Correct this statement (you will find that it does not type-check), strengthen it, and prove it. (Note: \\\ is function composition; its definition is found in theorem @{thm[source]o_def}). \end{exercise} \begin{exercise}\label{ex:trev-trev} Define a function \<^term>\trev\ of type \<^typ>\('v,'f)term => ('v,'f)term\ that recursively reverses the order of arguments of all function symbols in a term. Prove that \<^prop>\trev(trev t) = t\. \end{exercise} The experienced functional programmer may feel that our definition of \<^term>\subst\ is too complicated in that \<^const>\substs\ is unnecessary. The \<^term>\App\-case can be defined directly as @{term[display]"subst s (App f ts) = App f (map (subst s) ts)"} where \<^term>\map\ is the standard list function such that \map f [x1,...,xn] = [f x1,...,f xn]\. This is true, but Isabelle insists on the conjunctive format. Fortunately, we can easily \emph{prove} that the suggested equation holds: \ (*<*) (* Exercise 1: *) lemma "subst ((subst f) \ g) t = subst f (subst g t) \ substs ((subst f) \ g) ts = substs f (substs g ts)" apply (induct_tac t and ts rule: subst.induct substs.induct) apply (simp_all) done (* Exercise 2: *) primrec trev :: "('v,'f) term \ ('v,'f) term" and trevs:: "('v,'f) term list \ ('v,'f) term list" where "trev (Var v) = Var v" | "trev (App f ts) = App f (trevs ts)" | "trevs [] = []" | "trevs (t#ts) = (trevs ts) @ [(trev t)]" lemma [simp]: "\ ys. trevs (xs @ ys) = (trevs ys) @ (trevs xs)" apply (induct_tac xs, auto) done lemma "trev (trev t) = (t::('v,'f)term) \ trevs (trevs ts) = (ts::('v,'f)term list)" apply (induct_tac t and ts rule: trev.induct trevs.induct, simp_all) done (*>*) lemma [simp]: "subst s (App f ts) = App f (map (subst s) ts)" apply(induct_tac ts, simp_all) done text\\noindent What is more, we can now disable the old defining equation as a simplification rule: \ declare subst_App [simp del] text\\noindent The advantage is that now we have replaced \<^const>\substs\ by \<^const>\map\, we can profit from the large number of pre-proved lemmas about \<^const>\map\. Unfortunately, inductive proofs about type \term\ are still awkward because they expect a conjunction. One could derive a new induction principle as well (see \S\ref{sec:derive-ind}), but simpler is to stop using \isacommand{primrec} and to define functions with \isacommand{fun} instead. Simple uses of \isacommand{fun} are described in \S\ref{sec:fun} below. Advanced applications, including functions over nested datatypes like \term\, are discussed in a -separate tutorial~@{cite "isabelle-function"}. +separate tutorial~\<^cite>\"isabelle-function"\. Of course, you may also combine mutual and nested recursion of datatypes. For example, constructor \Sum\ in \S\ref{sec:datatype-mut-rec} could take a list of expressions as its argument: \Sum\~@{typ[quotes]"'a aexp list"}. \ (*<*)end(*>*) diff --git a/src/Doc/Tutorial/Documents/Documents.thy b/src/Doc/Tutorial/Documents/Documents.thy --- a/src/Doc/Tutorial/Documents/Documents.thy +++ b/src/Doc/Tutorial/Documents/Documents.thy @@ -1,725 +1,725 @@ (*<*) theory Documents imports Main begin (*>*) section \Concrete Syntax \label{sec:concrete-syntax}\ text \ The core concept of Isabelle's framework for concrete syntax is that of \bfindex{mixfix annotations}. Associated with any kind of constant declaration, mixfixes affect both the grammar productions for the parser and output templates for the pretty printer. In full generality, parser and pretty printer configuration is a - subtle affair~@{cite "isabelle-isar-ref"}. Your syntax specifications need + subtle affair~\<^cite>\"isabelle-isar-ref"\. Your syntax specifications need to interact properly with the existing setup of Isabelle/Pure and Isabelle/HOL\@. To avoid creating ambiguities with existing elements, it is particularly important to give new syntactic constructs the right precedence. Below we introduce a few simple syntax declaration forms that already cover many common situations fairly well. \ subsection \Infix Annotations\ text \ Syntax annotations may be included wherever constants are declared, such as \isacommand{definition} and \isacommand{primrec} --- and also \isacommand{datatype}, which declares constructor operations. Type-constructors may be annotated as well, although this is less frequently encountered in practice (the infix type \\\ comes to mind). Infix declarations\index{infix annotations} provide a useful special case of mixfixes. The following example of the exclusive-or operation on boolean values illustrates typical infix declarations. \ definition xor :: "bool \ bool \ bool" (infixl "[+]" 60) where "A [+] B \ (A \ \ B) \ (\ A \ B)" text \ \noindent Now \xor A B\ and \A [+] B\ refer to the same expression internally. Any curried function with at least two arguments may be given infix syntax. For partial applications with fewer than two operands, the operator is enclosed in parentheses. For instance, \xor\ without arguments is represented as \([+])\; together with ordinary function application, this turns \xor A\ into \([+]) A\. The keyword \isakeyword{infixl} seen above specifies an infix operator that is nested to the \emph{left}: in iterated applications the more complex expression appears on the left-hand side, and \<^term>\A [+] B [+] C\ stands for \(A [+] B) [+] C\. Similarly, \isakeyword{infixr} means nesting to the \emph{right}, reading \<^term>\A [+] B [+] C\ as \A [+] (B [+] C)\. A \emph{non-oriented} declaration via \isakeyword{infix} would render \<^term>\A [+] B [+] C\ illegal, but demand explicit parentheses to indicate the intended grouping. The string @{text [source] "[+]"} in our annotation refers to the concrete syntax to represent the operator (a literal token), while the number \60\ determines the precedence of the construct: the syntactic priorities of the arguments and result. Isabelle/HOL already uses up many popular combinations of ASCII symbols for its own use, including both \+\ and \++\. Longer character combinations are more likely to be still available for user extensions, such as our~\[+]\. Operator precedences have a range of 0--1000. Very low or high priorities are reserved for the meta-logic. HOL syntax mainly uses the range of 10--100: the equality infix \=\ is centered at 50; logical connectives (like \\\ and \\\) are below 50; algebraic ones (like \+\ and \*\) are above 50. User syntax should strive to coexist with common HOL forms, or use the mostly unused range 100--900. \ subsection \Mathematical Symbols \label{sec:syntax-symbols}\ text \ Concrete syntax based on ASCII characters has inherent limitations. Mathematical notation demands a larger repertoire of glyphs. Several standards of extended character sets have been proposed over decades, but none has become universally available so far. Isabelle has its own notion of \bfindex{symbols} as the smallest entities of source text, without referring to internal encodings. There are three kinds of such ``generalized characters'': \begin{enumerate} \item 7-bit ASCII characters \item named symbols: \verb,\,\verb,<,$ident$\verb,>, \item named control symbols: \verb,\,\verb,<^,$ident$\verb,>, \end{enumerate} Here $ident$ is any sequence of letters. This results in an infinite store of symbols, whose interpretation is left to further front-end tools. For example, the Isabelle document processor (see \S\ref{sec:document-preparation}) display the \verb,\,\verb,, symbol as~\\\. A list of standard Isabelle symbols is given in - @{cite "isabelle-isar-ref"}. You may introduce your own + \<^cite>\"isabelle-isar-ref"\. You may introduce your own interpretation of further symbols by configuring the appropriate front-end tool accordingly, e.g.\ by defining certain {\LaTeX} macros (see also \S\ref{sec:doc-prep-symbols}). There are also a few predefined control symbols, such as \verb,\,\verb,<^sub>, and \verb,\,\verb,<^sup>, for sub- and superscript of the subsequent printable symbol, respectively. For example, \<^verbatim>\A\<^sup>\\, is output as \A\<^sup>\\. A number of symbols are considered letters by the Isabelle lexer and can be used as part of identifiers. These are the greek letters \\\ (\verb+\+\verb++), \\\ (\verb+\+\verb++), etc. (excluding \\\), special letters like \\\ (\verb+\+\verb++) and \\\ (\verb+\+\verb++). Moreover the control symbol \verb+\+\verb+<^sub>+ may be used to subscript a single letter or digit in the trailing part of an identifier. This means that the input \medskip {\small\noindent \<^verbatim>\\\\<^sub>1. \\<^sub>1 = \\<^sub>\\} \medskip \noindent is recognized as the term \<^term>\\\\<^sub>1. \\<^sub>1 = \\<^sub>\\ by Isabelle. Replacing our previous definition of \xor\ by the following specifies an Isabelle symbol for the new operator: \ (*<*) hide_const xor setup \Sign.add_path "version1"\ (*>*) definition xor :: "bool \ bool \ bool" (infixl "\" 60) where "A \ B \ (A \ \ B) \ (\ A \ B)" (*<*) setup \Sign.local_path\ (*>*) text \ It is possible to provide alternative syntax forms - through the \bfindex{print mode} concept~@{cite "isabelle-isar-ref"}. By + through the \bfindex{print mode} concept~\<^cite>\"isabelle-isar-ref"\. By convention, the mode of ``$xsymbols$'' is enabled whenever Proof~General's X-Symbol mode or {\LaTeX} output is active. Now consider the following hybrid declaration of \xor\: \ (*<*) hide_const xor setup \Sign.add_path "version2"\ (*>*) definition xor :: "bool \ bool \ bool" (infixl "[+]\" 60) where "A [+]\ B \ (A \ \ B) \ (\ A \ B)" notation (xsymbols) xor (infixl "\\" 60) (*<*) setup \Sign.local_path\ (*>*) text \\noindent The \commdx{notation} command associates a mixfix annotation with a known constant. The print mode specification, here \(xsymbols)\, is optional. We may now write \A [+] B\ or \A \ B\ in input, while output uses the nicer syntax of $xsymbols$ whenever that print mode is active. Such an arrangement is particularly useful for interactive development, where users may type ASCII text and see mathematical symbols displayed during proofs.\ subsection \Prefix Annotations\ text \ Prefix syntax annotations\index{prefix annotation} are another form - of mixfixes @{cite "isabelle-isar-ref"}, without any template arguments or + of mixfixes \<^cite>\"isabelle-isar-ref"\, without any template arguments or priorities --- just some literal syntax. The following example associates common symbols with the constructors of a datatype. \ datatype currency = Euro nat ("\") | Pounds nat ("\") | Yen nat ("\") | Dollar nat ("$") text \ \noindent Here the mixfix annotations on the rightmost column happen to consist of a single Isabelle symbol each: \verb,\,\verb,,, \verb,\,\verb,,, \verb,\,\verb,,, and \verb,$,. Recall that a constructor like \Euro\ actually is a function \<^typ>\nat \ currency\. The expression \Euro 10\ will be printed as \<^term>\\ 10\; only the head of the application is subject to our concrete syntax. This rather simple form already achieves conformance with notational standards of the European Commission. Prefix syntax works the same way for other commands that introduce new constants, e.g. \isakeyword{primrec}. \ subsection \Abbreviations \label{sec:abbreviations}\ text\Mixfix syntax annotations merely decorate particular constant application forms with concrete syntax, for instance replacing \xor A B\ by \A \ B\. Occasionally, the relationship between some piece of notation and its internal form is more complicated. Here we need \emph{abbreviations}. Command \commdx{abbreviation} introduces an uninterpreted notational constant as an abbreviation for a complex term. Abbreviations are unfolded upon parsing and re-introduced upon printing. This provides a simple mechanism for syntactic macros. A typical use of abbreviations is to introduce relational notation for membership in a set of pairs, replacing \(x, y) \ sim\ by \x \ y\. We assume that a constant \sim\ of type \<^typ>\('a \ 'a) set\ has been introduced at this point.\ (*<*)consts sim :: "('a \ 'a) set"(*>*) abbreviation sim2 :: "'a \ 'a \ bool" (infix "\" 50) where "x \ y \ (x, y) \ sim" text \\noindent The given meta-equality is used as a rewrite rule after parsing (replacing \mbox{\<^prop>\x \ y\} by \(x,y) \ sim\) and before printing (turning \(x,y) \ sim\ back into \mbox{\<^prop>\x \ y\}). The name of the dummy constant \sim2\ does not matter, as long as it is unique. Another common application of abbreviations is to provide variant versions of fundamental relational expressions, such as \\\ for negated equalities. The following declaration stems from Isabelle/HOL itself: \ abbreviation not_equal :: "'a \ 'a \ bool" (infixl "~=\" 50) where "x ~=\ y \ \ (x = y)" notation (xsymbols) not_equal (infix "\\" 50) text \\noindent The notation \\\ is introduced separately to restrict it to the \emph{xsymbols} mode. Abbreviations are appropriate when the defined concept is a simple variation on an existing one. But because of the automatic folding and unfolding of abbreviations, they do not scale up well to large hierarchies of concepts. Abbreviations do not replace definitions. Abbreviations are a simplified form of the general concept of \emph{syntax translations}; even heavier transformations may be -written in ML @{cite "isabelle-isar-ref"}. +written in ML \<^cite>\"isabelle-isar-ref"\. \ section \Document Preparation \label{sec:document-preparation}\ text \ Isabelle/Isar is centered around the concept of \bfindex{formal proof documents}\index{documents|bold}. The outcome of a formal development effort is meant to be a human-readable record, presented as browsable PDF file or printed on paper. The overall document structure follows traditional mathematical articles, with sections, intermediate explanations, definitions, theorems and proofs. \medskip The Isabelle document preparation system essentially acts as a front-end to {\LaTeX}. After checking specifications and proofs formally, the theory sources are turned into typesetting instructions in a schematic manner. This lets you write authentic reports on theory developments with little effort: many technical consistency checks are handled by the system. Here is an example to illustrate the idea of Isabelle document preparation. \ text_raw \\begin{quotation}\ text \ The following datatype definition of \'a bintree\ models binary trees with nodes being decorated by elements of type \<^typ>\'a\. \ datatype 'a bintree = Leaf | Branch 'a "'a bintree" "'a bintree" text \ \noindent The datatype induction rule generated here is of the form @{thm [indent = 1, display] bintree.induct [no_vars]} \ text_raw \\end{quotation}\ text \ \noindent The above document output has been produced as follows: \begin{ttbox} text {\ttlbrace}* The following datatype definition of {\at}{\ttlbrace}text "'a bintree"{\ttrbrace} models binary trees with nodes being decorated by elements of type {\at}{\ttlbrace}typ 'a{\ttrbrace}. *{\ttrbrace} datatype 'a bintree = Leaf | Branch 'a "'a bintree" "'a bintree" \end{ttbox} \begin{ttbox} text {\ttlbrace}* {\ttback}noindent The datatype induction rule generated here is of the form {\at}{\ttlbrace}thm [display] bintree.induct [no_vars]{\ttrbrace} *{\ttrbrace} \end{ttbox}\vspace{-\medskipamount} \noindent Here we have augmented the theory by formal comments (using \isakeyword{text} blocks), the informal parts may again refer to formal entities by means of ``antiquotations'' (such as \texttt{\at}\verb,{text "'a bintree"}, or \texttt{\at}\verb,{typ 'a},), see also \S\ref{sec:doc-prep-text}. \ subsection \Isabelle Sessions\ text \ In contrast to the highly interactive mode of Isabelle/Isar theory development, the document preparation stage essentially works in batch-mode. An Isabelle \bfindex{session} consists of a collection of source files that may contribute to an output document. Each session is derived from a single parent, usually an object-logic image like \texttt{HOL}. This results in an overall tree structure, which is reflected by the output location in the file system (the root directory is determined by the Isabelle settings variable \verb,ISABELLE_BROWSER_INFO,). \medskip The easiest way to manage Isabelle sessions is via \texttt{isabelle mkroot} (to generate an initial session source setup) and \texttt{isabelle build} (to run sessions as specified in the corresponding \texttt{ROOT} file). These Isabelle tools are described in further detail in the \emph{Isabelle System Manual} - @{cite "isabelle-system"}. + \<^cite>\"isabelle-system"\. For example, a new session \texttt{MySession} (with document preparation) may be produced as follows: \begin{verbatim} isabelle mkroot MySession isabelle build -D MySession \end{verbatim} The \texttt{isabelle build} job also informs about the file-system location of the ultimate results. The above dry run should be able to produce some \texttt{document.pdf} (with dummy title, empty table of contents etc.). Any failure at this stage usually indicates technical problems of the {\LaTeX} installation. \medskip The detailed arrangement of the session sources is as follows. \begin{itemize} \item Directory \texttt{MySession} holds the required theory files $T@1$\texttt{.thy}, \dots, $T@n$\texttt{.thy}. \item File \texttt{MySession/ROOT} specifies the session options and content, with declarations for all wanted theories; it is sufficient to specify the terminal nodes of the theory dependency graph. \item Directory \texttt{MySession/document} contains everything required for the {\LaTeX} stage; only \texttt{root.tex} needs to be provided initially. The latter file holds appropriate {\LaTeX} code to commence a document (\verb,\documentclass, etc.), and to include the generated files $T@i$\texttt{.tex} for each theory. Isabelle will generate a file \texttt{session.tex} holding {\LaTeX} commands to include all generated theory output files in topologically sorted order, so \verb,\input{session}, in the body of \texttt{root.tex} does the job in most situations. \end{itemize} One may now start to populate the directory \texttt{MySession} and its \texttt{ROOT} file accordingly. The file \texttt{MySession/document/root.tex} should also be adapted at some point; the default version is mostly self-explanatory. Note that \verb,\isabellestyle, enables fine-tuning of the general appearance of characters and mathematical symbols (see also \S\ref{sec:doc-prep-symbols}). Especially observe the included {\LaTeX} packages \texttt{isabelle} (mandatory), \texttt{isabellesym} (required for mathematical symbols), and the final \texttt{pdfsetup} (provides sane defaults for \texttt{hyperref}, including URL markup). All three are distributed with Isabelle. Further packages may be required in particular applications, say for unusual mathematical symbols. \medskip Any additional files for the {\LaTeX} stage go into the \texttt{MySession/document} directory as well. In particular, adding a file named \texttt{root.bib} causes an automatic run of \texttt{bibtex} to process a bibliographic database; see also - \texttt{isabelle document} @{cite "isabelle-system"}. + \texttt{isabelle document} \<^cite>\"isabelle-system"\. \medskip Any failure of the document preparation phase in an Isabelle batch session leaves the generated sources in their target location, identified by the accompanying error message. This lets you trace {\LaTeX} problems with the generated files at hand. \ subsection \Structure Markup\ text \ The large-scale structure of Isabelle documents follows existing {\LaTeX} conventions, with chapters, sections, subsubsections etc. The Isar language includes separate \bfindex{markup commands}, which do not affect the formal meaning of a theory (or proof), but result in corresponding {\LaTeX} elements. From the Isabelle perspective, each markup command takes a single $text$ argument (delimited by \verb,",~\\\~\verb,", or \verb,{,\verb,*,~\\\~\verb,*,\verb,},). After stripping any surrounding white space, the argument is passed to a {\LaTeX} macro \verb,\isamarkupXYZ, for command \isakeyword{XYZ}. These macros are defined in \verb,isabelle.sty, according to the meaning given in the rightmost column above. \medskip The following source fragment illustrates structure markup of a theory. Note that {\LaTeX} labels may be included inside of section headings as well. \begin{ttbox} section {\ttlbrace}* Some properties of Foo Bar elements *{\ttrbrace} theory Foo_Bar imports Main begin subsection {\ttlbrace}* Basic definitions *{\ttrbrace} definition foo :: \dots definition bar :: \dots subsection {\ttlbrace}* Derived rules *{\ttrbrace} lemma fooI: \dots lemma fooE: \dots subsection {\ttlbrace}* Main theorem {\ttback}label{\ttlbrace}sec:main-theorem{\ttrbrace} *{\ttrbrace} theorem main: \dots end \end{ttbox} \ subsection \Formal Comments and Antiquotations \label{sec:doc-prep-text}\ text \ Isabelle \bfindex{source comments}, which are of the form \verb,(,\verb,*,~\\\~\verb,*,\verb,),, essentially act like white space and do not really contribute to the content. They mainly serve technical purposes to mark certain oddities in the raw input text. In contrast, \bfindex{formal comments} are portions of text that are associated with formal Isabelle/Isar commands (\bfindex{marginal comments}), or as standalone paragraphs within a theory or proof context (\bfindex{text blocks}). \medskip Marginal comments are part of each command's concrete - syntax @{cite "isabelle-isar-ref"}; the common form is ``\verb,--,~$text$'' + syntax \<^cite>\"isabelle-isar-ref"\; the common form is ``\verb,--,~$text$'' where $text$ is delimited by \verb,",\\\\verb,", or \verb,{,\verb,*,~\\\~\verb,*,\verb,}, as before. Multiple marginal comments may be given at the same time. Here is a simple example: \ lemma "A --> A" \ \a triviality of propositional logic\ \ \(should not really bother)\ by (rule impI) \ \implicit assumption step involved here\ text \ \noindent The above output has been produced as follows: \begin{verbatim} lemma "A --> A" -- "a triviality of propositional logic" -- "(should not really bother)" by (rule impI) -- "implicit assumption step involved here" \end{verbatim} From the {\LaTeX} viewpoint, ``\verb,--,'' acts like a markup command, associated with the macro \verb,\isamarkupcmt, (taking a single argument). \medskip Text blocks are introduced by the commands \bfindex{text} and \bfindex{txt}. Each takes again a single $text$ argument, which is interpreted as a free-form paragraph in {\LaTeX} (surrounded by some additional vertical space). The typesetting may be changed by redefining the {\LaTeX} environments of \verb,isamarkuptext, or \verb,isamarkuptxt,, respectively (via \verb,\renewenvironment,). \medskip The $text$ part of Isabelle markup commands essentially inserts \emph{quoted material} into a formal text, mainly for instruction of the reader. An \bfindex{antiquotation} is again a formal object embedded into such an informal portion. The interpretation of antiquotations is limited to some well-formedness checks, with the result being pretty printed to the resulting document. Quoted text blocks together with antiquotations provide an attractive means of referring to formal entities, with good confidence in getting the technical details right (especially syntax and types). The general syntax of antiquotations is as follows: \texttt{{\at}{\ttlbrace}$name$ $arguments${\ttrbrace}}, or \texttt{{\at}{\ttlbrace}$name$ [$options$] $arguments${\ttrbrace}} for a comma-separated list of options consisting of a $name$ or \texttt{$name$=$value$} each. The syntax of $arguments$ depends on the kind of antiquotation, it generally follows the same conventions for types, terms, or theorems as in the formal part of a theory. \medskip This sentence demonstrates quotations and antiquotations: \<^term>\%x y. x\ is a well-typed term. \medskip\noindent The output above was produced as follows: \begin{ttbox} text {\ttlbrace}* This sentence demonstrates quotations and antiquotations: {\at}{\ttlbrace}term "%x y. x"{\ttrbrace} is a well-typed term. *{\ttrbrace} \end{ttbox}\vspace{-\medskipamount} The notational change from the ASCII character~\verb,%, to the symbol~\\\ reveals that Isabelle printed this term, after parsing and type-checking. Document preparation enables symbolic output by default. \medskip The next example includes an option to show the type of all variables. The antiquotation \texttt{{\at}}\verb,{term [show_types] "%x y. x"}, produces the output @{term [show_types] "%x y. x"}. Type inference has figured out the most general typings in the present theory context. Terms may acquire different typings due to constraints imposed by their environment; within a proof, for example, variables are given the same types as they have in the main goal statement. \medskip Several further kinds of antiquotations and options are - available @{cite "isabelle-isar-ref"}. Here are a few commonly used + available \<^cite>\"isabelle-isar-ref"\. Here are a few commonly used combinations: \medskip \begin{tabular}{ll} \texttt{\at}\verb,{typ,~$\tau$\verb,}, & print type $\tau$ \\ \texttt{\at}\verb,{const,~$c$\verb,}, & check existence of $c$ and print it \\ \texttt{\at}\verb,{term,~$t$\verb,}, & print term $t$ \\ \texttt{\at}\verb,{prop,~$\phi$\verb,}, & print proposition $\phi$ \\ \texttt{\at}\verb,{prop [display],~$\phi$\verb,}, & print large proposition $\phi$ (with linebreaks) \\ \texttt{\at}\verb,{prop [source],~$\phi$\verb,}, & check proposition $\phi$, print its input \\ \texttt{\at}\verb,{thm,~$a$\verb,}, & print fact $a$ \\ \texttt{\at}\verb,{thm,~$a$~\verb,[no_vars]}, & print fact $a$, fixing schematic variables \\ \texttt{\at}\verb,{thm [source],~$a$\verb,}, & check availability of fact $a$, print its name \\ \texttt{\at}\verb,{text,~$s$\verb,}, & print uninterpreted text $s$ \\ \end{tabular} \medskip Note that \attrdx{no_vars} given above is \emph{not} an antiquotation option, but an attribute of the theorem argument given here. This might be useful with a diagnostic command like \isakeyword{thm}, too. \medskip The \texttt{\at}\verb,{text, $s$\verb,}, antiquotation is particularly interesting. Embedding uninterpreted text within an informal body might appear useless at first sight. Here the key virtue is that the string $s$ is processed as Isabelle output, interpreting Isabelle symbols appropriately. For example, \texttt{\at}\verb,{text "\\"}, produces \\\\, according to the standard interpretation of these symbol (cf.\ \S\ref{sec:doc-prep-symbols}). Thus we achieve consistent mathematical notation in both the formal and informal parts of the document very easily, independently of the term language of Isabelle. Manual {\LaTeX} code would leave more control over the typesetting, but is also slightly more tedious. \ subsection \Interpretation of Symbols \label{sec:doc-prep-symbols}\ text \ As has been pointed out before (\S\ref{sec:syntax-symbols}), Isabelle symbols are the smallest syntactic entities --- a straightforward generalization of ASCII characters. While Isabelle does not impose any interpretation of the infinite collection of named symbols, {\LaTeX} documents use canonical glyphs for certain - standard symbols @{cite "isabelle-isar-ref"}. + standard symbols \<^cite>\"isabelle-isar-ref"\. The {\LaTeX} code produced from Isabelle text follows a simple scheme. You can tune the final appearance by redefining certain macros, say in \texttt{root.tex} of the document. \begin{enumerate} \item 7-bit ASCII characters: letters \texttt{A\dots Z} and \texttt{a\dots z} are output directly, digits are passed as an argument to the \verb,\isadigit, macro, other characters are replaced by specifically named macros of the form \verb,\isacharXYZ,. \item Named symbols: \verb,\,\verb,, is turned into \verb,{\isasymXYZ},; note the additional braces. \item Named control symbols: \verb,\,\verb,<^XYZ>, is turned into \verb,\isactrlXYZ,; subsequent symbols may act as arguments if the control macro is defined accordingly. \end{enumerate} You may occasionally wish to give new {\LaTeX} interpretations of named symbols. This merely requires an appropriate definition of \verb,\isasymXYZ,, for \verb,\,\verb,, (see \texttt{isabelle.sty} for working examples). Control symbols are slightly more difficult to get right, though. \medskip The \verb,\isabellestyle, macro provides a high-level interface to tune the general appearance of individual symbols. For example, \verb,\isabellestyle{it}, uses the italics text style to mimic the general appearance of the {\LaTeX} math mode; double quotes are not printed at all. The resulting quality of typesetting is quite good, so this should be the default style for work that gets distributed to a broader audience. \ subsection \Suppressing Output \label{sec:doc-prep-suppress}\ text \ By default, Isabelle's document system generates a {\LaTeX} file for each theory that gets loaded while running the session. The generated \texttt{session.tex} will include all of these in order of appearance, which in turn gets included by the standard \texttt{root.tex}. Certainly one may change the order or suppress unwanted theories by ignoring \texttt{session.tex} and load individual files directly in \texttt{root.tex}. On the other hand, such an arrangement requires additional maintenance whenever the collection of theories changes. Alternatively, one may tune the theory loading process in \texttt{ROOT} itself: some sequential order of \textbf{theories} sections may enforce a certain traversal of the dependency graph, although this could degrade parallel processing. The nodes of each sub-graph that is specified here are presented in some topological order of their formal dependencies. Moreover, the system build option \verb,document=false, allows to disable document generation for some theories. Its usage in the session \texttt{ROOT} is like this: \begin{verbatim} theories [document = false] T \end{verbatim} \medskip Theory output may be suppressed more selectively, either via \bfindex{tagged command regions} or \bfindex{ignored material}. Tagged command regions works by annotating commands with named tags, which correspond to certain {\LaTeX} markup that tells how to treat particular parts of a document when doing the actual type-setting. By default, certain Isabelle/Isar commands are implicitly marked up using the predefined tags ``\emph{theory}'' (for theory begin and end), ``\emph{proof}'' (for proof commands), and ``\emph{ML}'' (for commands involving ML code). Users may add their own tags using the \verb,%,\emph{tag} notation right after a command name. In the subsequent example we hide a particularly irrelevant proof: \ lemma "x = x" by %invisible (simp) text \ The original source has been ``\verb,lemma "x = x" by %invisible (simp),''. Tags observe the structure of proofs; adjacent commands with the same tag are joined into a single region. The Isabelle document preparation system allows the user to specify how to interpret a tagged region, in order to keep, drop, or fold the corresponding parts of the document. See the \emph{Isabelle System Manual} - @{cite "isabelle-system"} for further details, especially on + \<^cite>\"isabelle-system"\ for further details, especially on \texttt{isabelle build} and \texttt{isabelle document}. Ignored material is specified by delimiting the original formal source with special source comments \verb,(,\verb,*,\verb,<,\verb,*,\verb,), and \verb,(,\verb,*,\verb,>,\verb,*,\verb,),. These parts are stripped before the type-setting phase, without affecting the formal checking of the theory, of course. For example, we may hide parts of a proof that seem unfit for general public inspection. The following ``fully automatic'' proof is actually a fake: \ lemma "x \ (0::int) \ 0 < x * x" by (auto(*<*)simp add: zero_less_mult_iff(*>*)) text \ \noindent The real source of the proof has been as follows: \begin{verbatim} by (auto(*<*)simp add: zero_less_mult_iff(*>*)) \end{verbatim} %(* \medskip Suppressing portions of printed text demands care. You should not misrepresent the underlying theory development. It is easy to invalidate the visible text by hiding references to questionable axioms, for example. \ (*<*) end (*>*) diff --git a/src/Doc/Tutorial/Fun/fun0.thy b/src/Doc/Tutorial/Fun/fun0.thy --- a/src/Doc/Tutorial/Fun/fun0.thy +++ b/src/Doc/Tutorial/Fun/fun0.thy @@ -1,260 +1,260 @@ (*<*) theory fun0 imports Main begin (*>*) text\ \subsection{Definition} \label{sec:fun-examples} Here is a simple example, the \rmindex{Fibonacci function}: \ fun fib :: "nat \ nat" where "fib 0 = 0" | "fib (Suc 0) = 1" | "fib (Suc(Suc x)) = fib x + fib (Suc x)" text\\noindent This resembles ordinary functional programming languages. Note the obligatory \isacommand{where} and \isa{|}. Command \isacommand{fun} declares and defines the function in one go. Isabelle establishes termination automatically because \<^const>\fib\'s argument decreases in every recursive call. Slightly more interesting is the insertion of a fixed element between any two elements of a list: \ fun sep :: "'a \ 'a list \ 'a list" where "sep a [] = []" | "sep a [x] = [x]" | "sep a (x#y#zs) = x # a # sep a (y#zs)" text\\noindent This time the length of the list decreases with the recursive call; the first argument is irrelevant for termination. Pattern matching\index{pattern matching!and \isacommand{fun}} need not be exhaustive and may employ wildcards: \ fun last :: "'a list \ 'a" where "last [x] = x" | "last (_#y#zs) = last (y#zs)" text\ Overlapping patterns are disambiguated by taking the order of equations into account, just as in functional programming: \ fun sep1 :: "'a \ 'a list \ 'a list" where "sep1 a (x#y#zs) = x # a # sep1 a (y#zs)" | "sep1 _ xs = xs" text\\noindent To guarantee that the second equation can only be applied if the first one does not match, Isabelle internally replaces the second equation by the two possibilities that are left: \<^prop>\sep1 a [] = []\ and \<^prop>\sep1 a [x] = [x]\. Thus the functions \<^const>\sep\ and \<^const>\sep1\ are identical. Because of its pattern matching syntax, \isacommand{fun} is also useful for the definition of non-recursive functions: \ fun swap12 :: "'a list \ 'a list" where "swap12 (x#y#zs) = y#x#zs" | "swap12 zs = zs" text\ After a function~$f$ has been defined via \isacommand{fun}, its defining equations (or variants derived from them) are available under the name $f$\.simps\ as theorems. For example, look (via \isacommand{thm}) at @{thm[source]sep.simps} and @{thm[source]sep1.simps} to see that they define the same function. What is more, those equations are automatically declared as simplification rules. \subsection{Termination} Isabelle's automatic termination prover for \isacommand{fun} has a fixed notion of the \emph{size} (of type \<^typ>\nat\) of an argument. The size of a natural number is the number itself. The size of a list is its length. For the general case see \S\ref{sec:general-datatype}. A recursive function is accepted if \isacommand{fun} can show that the size of one fixed argument becomes smaller with each recursive call. More generally, \isacommand{fun} allows any \emph{lexicographic combination} of size measures in case there are multiple arguments. For example, the following version of \rmindex{Ackermann's function} is accepted:\ fun ack2 :: "nat \ nat \ nat" where "ack2 n 0 = Suc n" | "ack2 0 (Suc m) = ack2 (Suc 0) m" | "ack2 (Suc n) (Suc m) = ack2 (ack2 n (Suc m)) m" text\The order of arguments has no influence on whether \isacommand{fun} can prove termination of a function. For more details -see elsewhere~@{cite bulwahnKN07}. +see elsewhere~\<^cite>\bulwahnKN07\. \subsection{Simplification} \label{sec:fun-simplification} Upon a successful termination proof, the recursion equations become simplification rules, just as with \isacommand{primrec}. In most cases this works fine, but there is a subtle problem that must be mentioned: simplification may not terminate because of automatic splitting of \if\. \index{*if expressions!splitting of} Let us look at an example: \ fun gcd :: "nat \ nat \ nat" where "gcd m n = (if n=0 then m else gcd n (m mod n))" text\\noindent The second argument decreases with each recursive call. The termination condition @{prop[display]"n ~= (0::nat) ==> m mod n < n"} is proved automatically because it is already present as a lemma in HOL\@. Thus the recursion equation becomes a simplification rule. Of course the equation is nonterminating if we are allowed to unfold the recursive call inside the \else\ branch, which is why programming languages and our simplifier don't do that. Unfortunately the simplifier does something else that leads to the same problem: it splits each \if\-expression unless its condition simplifies to \<^term>\True\ or \<^term>\False\. For example, simplification reduces @{prop[display]"gcd m n = k"} in one step to @{prop[display]"(if n=0 then m else gcd n (m mod n)) = k"} where the condition cannot be reduced further, and splitting leads to @{prop[display]"(n=0 --> m=k) & (n ~= 0 --> gcd n (m mod n)=k)"} Since the recursive call \<^term>\gcd n (m mod n)\ is no longer protected by an \if\, it is unfolded again, which leads to an infinite chain of simplification steps. Fortunately, this problem can be avoided in many different ways. The most radical solution is to disable the offending theorem @{thm[source]if_split}, as shown in \S\ref{sec:AutoCaseSplits}. However, we do not recommend this approach: you will often have to invoke the rule explicitly when \if\ is involved. If possible, the definition should be given by pattern matching on the left rather than \if\ on the right. In the case of \<^term>\gcd\ the following alternative definition suggests itself: \ fun gcd1 :: "nat \ nat \ nat" where "gcd1 m 0 = m" | "gcd1 m n = gcd1 n (m mod n)" text\\noindent The order of equations is important: it hides the side condition \<^prop>\n ~= (0::nat)\. Unfortunately, not all conditionals can be expressed by pattern matching. A simple alternative is to replace \if\ by \case\, which is also available for \<^typ>\bool\ and is not split automatically: \ fun gcd2 :: "nat \ nat \ nat" where "gcd2 m n = (case n=0 of True \ m | False \ gcd2 n (m mod n))" text\\noindent This is probably the neatest solution next to pattern matching, and it is always available. A final alternative is to replace the offending simplification rules by derived conditional ones. For \<^term>\gcd\ it means we have to prove these lemmas: \ lemma [simp]: "gcd m 0 = m" apply(simp) done lemma [simp]: "n \ 0 \ gcd m n = gcd n (m mod n)" apply(simp) done text\\noindent Simplification terminates for these proofs because the condition of the \if\ simplifies to \<^term>\True\ or \<^term>\False\. Now we can disable the original simplification rule: \ declare gcd.simps [simp del] text\ \index{induction!recursion|(} \index{recursion induction|(} \subsection{Induction} \label{sec:fun-induction} Having defined a function we might like to prove something about it. Since the function is recursive, the natural proof principle is again induction. But this time the structural form of induction that comes with datatypes is unlikely to work well --- otherwise we could have defined the function by \isacommand{primrec}. Therefore \isacommand{fun} automatically proves a suitable induction rule $f$\.induct\ that follows the recursion pattern of the particular function $f$. We call this \textbf{recursion induction}. Roughly speaking, it requires you to prove for each \isacommand{fun} equation that the property you are trying to establish holds for the left-hand side provided it holds for all recursive calls on the right-hand side. Here is a simple example involving the predefined \<^term>\map\ functional on lists: \ lemma "map f (sep x xs) = sep (f x) (map f xs)" txt\\noindent Note that \<^term>\map f xs\ is the result of applying \<^term>\f\ to all elements of \<^term>\xs\. We prove this lemma by recursion induction over \<^term>\sep\: \ apply(induct_tac x xs rule: sep.induct) txt\\noindent The resulting proof state has three subgoals corresponding to the three clauses for \<^term>\sep\: @{subgoals[display,indent=0]} The rest is pure simplification: \ apply simp_all done text\\noindent The proof goes smoothly because the induction rule follows the recursion of \<^const>\sep\. Try proving the above lemma by structural induction, and you find that you need an additional case distinction. In general, the format of invoking recursion induction is \begin{quote} \isacommand{apply}\(induct_tac\ $x@1 \dots x@n$ \rule:\ $f$\.induct)\ \end{quote}\index{*induct_tac (method)}% where $x@1~\dots~x@n$ is a list of free variables in the subgoal and $f$ the name of a function that takes $n$ arguments. Usually the subgoal will contain the term $f x@1 \dots x@n$ but this need not be the case. The induction rules do not mention $f$ at all. Here is @{thm[source]sep.induct}: \begin{isabelle} {\isasymlbrakk}~{\isasymAnd}a.~P~a~[];\isanewline ~~{\isasymAnd}a~x.~P~a~[x];\isanewline ~~{\isasymAnd}a~x~y~zs.~P~a~(y~\#~zs)~{\isasymLongrightarrow}~P~a~(x~\#~y~\#~zs){\isasymrbrakk}\isanewline {\isasymLongrightarrow}~P~u~v% \end{isabelle} It merely says that in order to prove a property \<^term>\P\ of \<^term>\u\ and \<^term>\v\ you need to prove it for the three cases where \<^term>\v\ is the empty list, the singleton list, and the list with at least two elements. The final case has an induction hypothesis: you may assume that \<^term>\P\ holds for the tail of that list. \index{induction!recursion|)} \index{recursion induction|)} \ (*<*) end (*>*) diff --git a/src/Doc/Tutorial/Inductive/AB.thy b/src/Doc/Tutorial/Inductive/AB.thy --- a/src/Doc/Tutorial/Inductive/AB.thy +++ b/src/Doc/Tutorial/Inductive/AB.thy @@ -1,304 +1,304 @@ (*<*)theory AB imports Main begin(*>*) section\Case Study: A Context Free Grammar\ text\\label{sec:CFG} \index{grammars!defining inductively|(}% Grammars are nothing but shorthands for inductive definitions of nonterminals which represent sets of strings. For example, the production $A \to B c$ is short for \[ w \in B \Longrightarrow wc \in A \] This section demonstrates this idea with an example due to Hopcroft and Ullman, a grammar for generating all words with an equal number of $a$'s and~$b$'s: \begin{eqnarray} S &\to& \epsilon \mid b A \mid a B \nonumber\\ A &\to& a S \mid b A A \nonumber\\ B &\to& b S \mid a B B \nonumber \end{eqnarray} At the end we say a few words about the relationship between -the original proof @{cite \p.\ts81\ HopcroftUllman} and our formal version. +the original proof \<^cite>\\p.\ts81\ in HopcroftUllman\ and our formal version. We start by fixing the alphabet, which consists only of \<^term>\a\'s and~\<^term>\b\'s: \ datatype alfa = a | b text\\noindent For convenience we include the following easy lemmas as simplification rules: \ lemma [simp]: "(x \ a) = (x = b) \ (x \ b) = (x = a)" by (case_tac x, auto) text\\noindent Words over this alphabet are of type \<^typ>\alfa list\, and the three nonterminals are declared as sets of such words. The productions above are recast as a \emph{mutual} inductive definition\index{inductive definition!simultaneous} of \<^term>\S\, \<^term>\A\ and~\<^term>\B\: \ inductive_set S :: "alfa list set" and A :: "alfa list set" and B :: "alfa list set" where "[] \ S" | "w \ A \ b#w \ S" | "w \ B \ a#w \ S" | "w \ S \ a#w \ A" | "\ v\A; w\A \ \ b#v@w \ A" | "w \ S \ b#w \ B" | "\ v \ B; w \ B \ \ a#v@w \ B" text\\noindent First we show that all words in \<^term>\S\ contain the same number of \<^term>\a\'s and \<^term>\b\'s. Since the definition of \<^term>\S\ is by mutual induction, so is the proof: we show at the same time that all words in \<^term>\A\ contain one more \<^term>\a\ than \<^term>\b\ and all words in \<^term>\B\ contain one more \<^term>\b\ than \<^term>\a\. \ lemma correctness: "(w \ S \ size[x\w. x=a] = size[x\w. x=b]) \ (w \ A \ size[x\w. x=a] = size[x\w. x=b] + 1) \ (w \ B \ size[x\w. x=b] = size[x\w. x=a] + 1)" txt\\noindent These propositions are expressed with the help of the predefined \<^term>\filter\ function on lists, which has the convenient syntax \[x\xs. P x]\, the list of all elements \<^term>\x\ in \<^term>\xs\ such that \<^prop>\P x\ holds. Remember that on lists \size\ and \length\ are synonymous. The proof itself is by rule induction and afterwards automatic: \ by (rule S_A_B.induct, auto) text\\noindent This may seem surprising at first, and is indeed an indication of the power of inductive definitions. But it is also quite straightforward. For example, consider the production $A \to b A A$: if $v,w \in A$ and the elements of $A$ contain one more $a$ than~$b$'s, then $bvw$ must again contain one more $a$ than~$b$'s. As usual, the correctness of syntactic descriptions is easy, but completeness is hard: does \<^term>\S\ contain \emph{all} words with an equal number of \<^term>\a\'s and \<^term>\b\'s? It turns out that this proof requires the following lemma: every string with two more \<^term>\a\'s than \<^term>\b\'s can be cut somewhere such that each half has one more \<^term>\a\ than \<^term>\b\. This is best seen by imagining counting the difference between the number of \<^term>\a\'s and \<^term>\b\'s starting at the left end of the word. We start with 0 and end (at the right end) with 2. Since each move to the right increases or decreases the difference by 1, we must have passed through 1 on our way from 0 to 2. Formally, we appeal to the following discrete intermediate value theorem @{thm[source]nat0_intermed_int_val} @{thm[display,margin=60]nat0_intermed_int_val[no_vars]} where \<^term>\f\ is of type \<^typ>\nat \ int\, \<^typ>\int\ are the integers, \\.\\ is the absolute value function\footnote{See Table~\ref{tab:ascii} in the Appendix for the correct \textsc{ascii} syntax.}, and \<^term>\1::int\ is the integer 1 (see \S\ref{sec:numbers}). First we show that our specific function, the difference between the numbers of \<^term>\a\'s and \<^term>\b\'s, does indeed only change by 1 in every move to the right. At this point we also start generalizing from \<^term>\a\'s and \<^term>\b\'s to an arbitrary property \<^term>\P\. Otherwise we would have to prove the desired lemma twice, once as stated above and once with the roles of \<^term>\a\'s and \<^term>\b\'s interchanged. \ lemma step1: "\i < size w. \(int(size[x\take (i+1) w. P x])-int(size[x\take (i+1) w. \P x])) - (int(size[x\take i w. P x])-int(size[x\take i w. \P x]))\ \ 1" txt\\noindent The lemma is a bit hard to read because of the coercion function \int :: nat \ int\. It is required because \<^term>\size\ returns a natural number, but subtraction on type~\<^typ>\nat\ will do the wrong thing. Function \<^term>\take\ is predefined and \<^term>\take i xs\ is the prefix of length \<^term>\i\ of \<^term>\xs\; below we also need \<^term>\drop i xs\, which is what remains after that prefix has been dropped from \<^term>\xs\. The proof is by induction on \<^term>\w\, with a trivial base case, and a not so trivial induction step. Since it is essentially just arithmetic, we do not discuss it. \ apply(induct_tac w) apply(auto simp add: abs_if take_Cons split: nat.split) done text\ Finally we come to the above-mentioned lemma about cutting in half a word with two more elements of one sort than of the other sort: \ lemma part1: "size[x\w. P x] = size[x\w. \P x]+2 \ \i\size w. size[x\take i w. P x] = size[x\take i w. \P x]+1" txt\\noindent This is proved by \force\ with the help of the intermediate value theorem, instantiated appropriately and with its first premise disposed of by lemma @{thm[source]step1}: \ apply(insert nat0_intermed_int_val[OF step1, of "P" "w" "1"]) by force text\\noindent Lemma @{thm[source]part1} tells us only about the prefix \<^term>\take i w\. An easy lemma deals with the suffix \<^term>\drop i w\: \ lemma part2: "\size[x\take i w @ drop i w. P x] = size[x\take i w @ drop i w. \P x]+2; size[x\take i w. P x] = size[x\take i w. \P x]+1\ \ size[x\drop i w. P x] = size[x\drop i w. \P x]+1" by(simp del: append_take_drop_id) text\\noindent In the proof we have disabled the normally useful lemma \begin{isabelle} @{thm append_take_drop_id[no_vars]} \rulename{append_take_drop_id} \end{isabelle} to allow the simplifier to apply the following lemma instead: @{text[display]"[x\xs@ys. P x] = [x\xs. P x] @ [x\ys. P x]"} To dispose of trivial cases automatically, the rules of the inductive definition are declared simplification rules: \ declare S_A_B.intros[simp] text\\noindent This could have been done earlier but was not necessary so far. The completeness theorem tells us that if a word has the same number of \<^term>\a\'s and \<^term>\b\'s, then it is in \<^term>\S\, and similarly for \<^term>\A\ and \<^term>\B\: \ theorem completeness: "(size[x\w. x=a] = size[x\w. x=b] \ w \ S) \ (size[x\w. x=a] = size[x\w. x=b] + 1 \ w \ A) \ (size[x\w. x=b] = size[x\w. x=a] + 1 \ w \ B)" txt\\noindent The proof is by induction on \<^term>\w\. Structural induction would fail here because, as we can see from the grammar, we need to make bigger steps than merely appending a single letter at the front. Hence we induct on the length of \<^term>\w\, using the induction rule @{thm[source]length_induct}: \ apply(induct_tac w rule: length_induct) apply(rename_tac w) txt\\noindent The \rule\ parameter tells \induct_tac\ explicitly which induction rule to use. For details see \S\ref{sec:complete-ind} below. In this case the result is that we may assume the lemma already holds for all words shorter than \<^term>\w\. Because the induction step renames the induction variable we rename it back to \w\. The proof continues with a case distinction on \<^term>\w\, on whether \<^term>\w\ is empty or not. \ apply(case_tac w) apply(simp_all) (*<*)apply(rename_tac x v)(*>*) txt\\noindent Simplification disposes of the base case and leaves only a conjunction of two step cases to be proved: if \<^prop>\w = a#v\ and @{prop[display]"size[x\v. x=a] = size[x\v. x=b]+2"} then \<^prop>\b#v \ A\, and similarly for \<^prop>\w = b#v\. We only consider the first case in detail. After breaking the conjunction up into two cases, we can apply @{thm[source]part1} to the assumption that \<^term>\w\ contains two more \<^term>\a\'s than \<^term>\b\'s. \ apply(rule conjI) apply(clarify) apply(frule part1[of "\x. x=a", simplified]) apply(clarify) txt\\noindent This yields an index \<^prop>\i \ length v\ such that @{prop[display]"length [x\take i v . x = a] = length [x\take i v . x = b] + 1"} With the help of @{thm[source]part2} it follows that @{prop[display]"length [x\drop i v . x = a] = length [x\drop i v . x = b] + 1"} \ apply(drule part2[of "\x. x=a", simplified]) apply(assumption) txt\\noindent Now it is time to decompose \<^term>\v\ in the conclusion \<^prop>\b#v \ A\ into \<^term>\take i v @ drop i v\, \ apply(rule_tac n1=i and t=v in subst[OF append_take_drop_id]) txt\\noindent (the variables \<^term>\n1\ and \<^term>\t\ are the result of composing the theorems @{thm[source]subst} and @{thm[source]append_take_drop_id}) after which the appropriate rule of the grammar reduces the goal to the two subgoals \<^prop>\take i v \ A\ and \<^prop>\drop i v \ A\: \ apply(rule S_A_B.intros) txt\ Both subgoals follow from the induction hypothesis because both \<^term>\take i v\ and \<^term>\drop i v\ are shorter than \<^term>\w\: \ apply(force simp add: min_less_iff_disj) apply(force split: nat_diff_split) txt\ The case \<^prop>\w = b#v\ is proved analogously: \ apply(clarify) apply(frule part1[of "\x. x=b", simplified]) apply(clarify) apply(drule part2[of "\x. x=b", simplified]) apply(assumption) apply(rule_tac n1=i and t=v in subst[OF append_take_drop_id]) apply(rule S_A_B.intros) apply(force simp add: min_less_iff_disj) by(force simp add: min_less_iff_disj split: nat_diff_split) text\ We conclude this section with a comparison of our proof with Hopcroft\index{Hopcroft, J. E.} and Ullman's\index{Ullman, J. D.} -@{cite \p.\ts81\ HopcroftUllman}. +\<^cite>\\p.\ts81\ in HopcroftUllman\. For a start, the textbook grammar, for no good reason, excludes the empty word, thus complicating matters just a little bit: they have 8 instead of our 7 productions. More importantly, the proof itself is different: rather than separating the two directions, they perform one induction on the length of a word. This deprives them of the beauty of rule induction, and in the easy direction (correctness) their reasoning is more detailed than our \auto\. For the hard part (completeness), they consider just one of the cases that our \simp_all\ disposes of automatically. Then they conclude the proof by saying about the remaining cases: ``We do this in a manner similar to our method of proof for part (1); this part is left to the reader''. But this is precisely the part that requires the intermediate value theorem and thus is not at all similar to the other cases (which are automatic in Isabelle). The authors are at least cavalier about this point and may even have overlooked the slight difficulty lurking in the omitted cases. Such errors are found in many pen-and-paper proofs when they are scrutinized formally.% \index{grammars!defining inductively|)} \ (*<*)end(*>*) diff --git a/src/Doc/Tutorial/Protocol/NS_Public.thy b/src/Doc/Tutorial/Protocol/NS_Public.thy --- a/src/Doc/Tutorial/Protocol/NS_Public.thy +++ b/src/Doc/Tutorial/Protocol/NS_Public.thy @@ -1,400 +1,400 @@ (* Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1996 University of Cambridge Inductive relation "ns_public" for the Needham-Schroeder Public-Key protocol. Version incorporating Lowe's fix (inclusion of B's identity in round 2). *)(*<*) theory NS_Public imports Public begin(*>*) section\Modelling the Protocol \label{sec:modelling}\ text_raw \ \begin{figure} \begin{isabelle} \ inductive_set ns_public :: "event list set" where Nil: "[] \ ns_public" | Fake: "\evsf \ ns_public; X \ synth (analz (knows Spy evsf))\ \ Says Spy B X # evsf \ ns_public" | NS1: "\evs1 \ ns_public; Nonce NA \ used evs1\ \ Says A B (Crypt (pubK B) \Nonce NA, Agent A\) # evs1 \ ns_public" | NS2: "\evs2 \ ns_public; Nonce NB \ used evs2; Says A' B (Crypt (pubK B) \Nonce NA, Agent A\) \ set evs2\ \ Says B A (Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\) # evs2 \ ns_public" | NS3: "\evs3 \ ns_public; Says A B (Crypt (pubK B) \Nonce NA, Agent A\) \ set evs3; Says B' A (Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\) \ set evs3\ \ Says A B (Crypt (pubK B) (Nonce NB)) # evs3 \ ns_public" text_raw \ \end{isabelle} \caption{An Inductive Protocol Definition}\label{fig:ns_public} \end{figure} \ text \ Let us formalize the Needham-Schroeder public-key protocol, as corrected by Lowe: \begin{alignat*% }{2} &1.&\quad A\to B &: \comp{Na,A}\sb{Kb} \\ &2.&\quad B\to A &: \comp{Na,Nb,B}\sb{Ka} \\ &3.&\quad A\to B &: \comp{Nb}\sb{Kb} \end{alignat*% } Each protocol step is specified by a rule of an inductive definition. An event trace has type \event list\, so we declare the constant \ns_public\ to be a set of such traces. Figure~\ref{fig:ns_public} presents the inductive definition. The \Nil\ rule introduces the empty trace. The \Fake\ rule models the adversary's sending a message built from components taken from past traffic, expressed using the functions \synth\ and \analz\. The next three rules model how honest agents would perform the three protocol steps. Here is a detailed explanation of rule \NS2\. A trace containing an event of the form @{term [display,indent=5] "Says A' B (Crypt (pubK B) \Nonce NA, Agent A\)"} may be extended by an event of the form @{term [display,indent=5] "Says B A (Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\)"} where \NB\ is a fresh nonce: \<^term>\Nonce NB \ used evs2\. Writing the sender as \A'\ indicates that \B\ does not know who sent the message. Calling the trace variable \evs2\ rather than simply \evs\ helps us know where we are in a proof after many case-splits: every subgoal mentioning \evs2\ involves message~2 of the protocol. Benefits of this approach are simplicity and clarity. The semantic model is set theory, proofs are by induction and the translation from the informal notation to the inductive rules is straightforward. \ section\Proving Elementary Properties \label{sec:regularity}\ (*<*) declare knows_Spy_partsEs [elim] declare analz_subset_parts [THEN subsetD, dest] declare Fake_parts_insert [THEN subsetD, dest] (*A "possibility property": there are traces that reach the end*) lemma "\NB. \evs \ ns_public. Says A B (Crypt (pubK B) (Nonce NB)) \ set evs" apply (intro exI bexI) apply (rule_tac [2] ns_public.Nil [THEN ns_public.NS1, THEN ns_public.NS2, THEN ns_public.NS3]) by possibility (**** Inductive proofs about ns_public ****) (** Theorems of the form X \ parts (knows Spy evs) imply that NOBODY sends messages containing X! **) (*Spy never sees another agent's private key! (unless it's bad at start)*) (*>*) text \ Secrecy properties can be hard to prove. The conclusion of a typical secrecy theorem is \<^term>\X \ analz (knows Spy evs)\. The difficulty arises from having to reason about \analz\, or less formally, showing that the spy can never learn~\X\. Much easier is to prove that \X\ can never occur at all. Such \emph{regularity} properties are typically expressed using \parts\ rather than \analz\. The following lemma states that \A\'s private key is potentially known to the spy if and only if \A\ belongs to the set \bad\ of compromised agents. The statement uses \parts\: the very presence of \A\'s private key in a message, whether protected by encryption or not, is enough to confirm that \A\ is compromised. The proof, like nearly all protocol proofs, is by induction over traces. \ lemma Spy_see_priK [simp]: "evs \ ns_public \ (Key (priK A) \ parts (knows Spy evs)) = (A \ bad)" apply (erule ns_public.induct, simp_all) txt \ The induction yields five subgoals, one for each rule in the definition of \ns_public\. The idea is to prove that the protocol property holds initially (rule \Nil\), is preserved by each of the legitimate protocol steps (rules \NS1\--\3\), and even is preserved in the face of anything the spy can do (rule \Fake\). The proof is trivial. No legitimate protocol rule sends any keys at all, so only \Fake\ is relevant. Indeed, simplification leaves only the \Fake\ case, as indicated by the variable name \evsf\: @{subgoals[display,indent=0,margin=65]} \ by blast (*<*) lemma Spy_analz_priK [simp]: "evs \ ns_public \ (Key (priK A) \ analz (knows Spy evs)) = (A \ bad)" by auto (*>*) text \ The \Fake\ case is proved automatically. If \<^term>\priK A\ is in the extended trace then either (1) it was already in the original trace or (2) it was generated by the spy, who must have known this key already. Either way, the induction hypothesis applies. \emph{Unicity} lemmas are regularity lemmas stating that specified items can occur only once in a trace. The following lemma states that a nonce cannot be used both as $Na$ and as $Nb$ unless it is known to the spy. Intuitively, it holds because honest agents always choose fresh values as nonces; only the spy might reuse a value, and he doesn't know this particular value. The proof script is short: induction, simplification, \blast\. The first line uses the rule \rev_mp\ to prepare the induction by moving two assumptions into the induction formula. \ lemma no_nonce_NS1_NS2: "\Crypt (pubK C) \NA', Nonce NA, Agent D\ \ parts (knows Spy evs); Crypt (pubK B) \Nonce NA, Agent A\ \ parts (knows Spy evs); evs \ ns_public\ \ Nonce NA \ analz (knows Spy evs)" apply (erule rev_mp, erule rev_mp) apply (erule ns_public.induct, simp_all) apply (blast intro: analz_insertI)+ done text \ The following unicity lemma states that, if \isa{NA} is secret, then its appearance in any instance of message~1 determines the other components. The proof is similar to the previous one. \ lemma unique_NA: "\Crypt(pubK B) \Nonce NA, Agent A \ \ parts(knows Spy evs); Crypt(pubK B') \Nonce NA, Agent A'\ \ parts(knows Spy evs); Nonce NA \ analz (knows Spy evs); evs \ ns_public\ \ A=A' \ B=B'" (*<*) apply (erule rev_mp, erule rev_mp, erule rev_mp) apply (erule ns_public.induct, simp_all) (*Fake, NS1*) apply (blast intro: analz_insertI)+ done (*>*) section\Proving Secrecy Theorems \label{sec:secrecy}\ (*<*) (*Secrecy: Spy does not see the nonce sent in msg NS1 if A and B are secure The major premise "Says A B ..." makes it a dest-rule, so we use (erule rev_mp) rather than rule_format. *) theorem Spy_not_see_NA: "\Says A B (Crypt(pubK B) \Nonce NA, Agent A\) \ set evs; A \ bad; B \ bad; evs \ ns_public\ \ Nonce NA \ analz (knows Spy evs)" apply (erule rev_mp) apply (erule ns_public.induct, simp_all) apply spy_analz apply (blast dest: unique_NA intro: no_nonce_NS1_NS2)+ done (*Authentication for A: if she receives message 2 and has used NA to start a run, then B has sent message 2.*) lemma A_trusts_NS2_lemma [rule_format]: "\A \ bad; B \ bad; evs \ ns_public\ \ Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\ \ parts (knows Spy evs) \ Says A B (Crypt(pubK B) \Nonce NA, Agent A\) \ set evs \ Says B A (Crypt(pubK A) \Nonce NA, Nonce NB, Agent B\) \ set evs" apply (erule ns_public.induct, simp_all) (*Fake, NS1*) apply (blast dest: Spy_not_see_NA)+ done theorem A_trusts_NS2: "\Says A B (Crypt(pubK B) \Nonce NA, Agent A\) \ set evs; Says B' A (Crypt(pubK A) \Nonce NA, Nonce NB, Agent B\) \ set evs; A \ bad; B \ bad; evs \ ns_public\ \ Says B A (Crypt(pubK A) \Nonce NA, Nonce NB, Agent B\) \ set evs" by (blast intro: A_trusts_NS2_lemma) (*If the encrypted message appears then it originated with Alice in NS1*) lemma B_trusts_NS1 [rule_format]: "evs \ ns_public \ Crypt (pubK B) \Nonce NA, Agent A\ \ parts (knows Spy evs) \ Nonce NA \ analz (knows Spy evs) \ Says A B (Crypt (pubK B) \Nonce NA, Agent A\) \ set evs" apply (erule ns_public.induct, simp_all) (*Fake*) apply (blast intro!: analz_insertI) done (*** Authenticity properties obtained from NS2 ***) (*Unicity for NS2: nonce NB identifies nonce NA and agents A, B [unicity of B makes Lowe's fix work] [proof closely follows that for unique_NA] *) lemma unique_NB [dest]: "\Crypt(pubK A) \Nonce NA, Nonce NB, Agent B\ \ parts(knows Spy evs); Crypt(pubK A') \Nonce NA', Nonce NB, Agent B'\ \ parts(knows Spy evs); Nonce NB \ analz (knows Spy evs); evs \ ns_public\ \ A=A' \ NA=NA' \ B=B'" apply (erule rev_mp, erule rev_mp, erule rev_mp) apply (erule ns_public.induct, simp_all) (*Fake, NS2*) apply (blast intro: analz_insertI)+ done (*>*) text \ The secrecy theorems for Bob (the second participant) are especially important because they fail for the original protocol. The following theorem states that if Bob sends message~2 to Alice, and both agents are uncompromised, then Bob's nonce will never reach the spy. \ theorem Spy_not_see_NB [dest]: "\Says B A (Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\) \ set evs; A \ bad; B \ bad; evs \ ns_public\ \ Nonce NB \ analz (knows Spy evs)" txt \ To prove it, we must formulate the induction properly (one of the assumptions mentions~\evs\), apply induction, and simplify: \ apply (erule rev_mp, erule ns_public.induct, simp_all) (*<*) apply spy_analz defer apply (blast intro: no_nonce_NS1_NS2) apply (blast intro: no_nonce_NS1_NS2) (*>*) txt \ The proof states are too complicated to present in full. Let's examine the simplest subgoal, that for message~1. The following event has just occurred: \[ 1.\quad A'\to B' : \comp{Na',A'}\sb{Kb'} \] The variables above have been primed because this step belongs to a different run from that referred to in the theorem statement --- the theorem refers to a past instance of message~2, while this subgoal concerns message~1 being sent just now. In the Isabelle subgoal, instead of primed variables like $B'$ and $Na'$ we have \Ba\ and~\NAa\: @{subgoals[display,indent=0]} The simplifier has used a default simplification rule that does a case analysis for each encrypted message on whether or not the decryption key is compromised. @{named_thms [display,indent=0,margin=50] analz_Crypt_if [no_vars] (analz_Crypt_if)} The simplifier has also used \Spy_see_priK\, proved in {\S}\ref{sec:regularity} above, to yield \<^term>\Ba \ bad\. Recall that this subgoal concerns the case where the last message to be sent was \[ 1.\quad A'\to B' : \comp{Na',A'}\sb{Kb'}. \] This message can compromise $Nb$ only if $Nb=Na'$ and $B'$ is compromised, allowing the spy to decrypt the message. The Isabelle subgoal says precisely this, if we allow for its choice of variable names. Proving \<^term>\NB \ NAa\ is easy: \NB\ was sent earlier, while \NAa\ is fresh; formally, we have the assumption \<^term>\Nonce NAa \ used evs1\. Note that our reasoning concerned \B\'s participation in another run. Agents may engage in several runs concurrently, and some attacks work by interleaving the messages of two runs. With model checking, this possibility can cause a state-space explosion, and for us it certainly complicates proofs. The biggest subgoal concerns message~2. It splits into several cases, such as whether or not the message just sent is the very message mentioned in the theorem statement. Some of the cases are proved by unicity, others by the induction hypothesis. For all those complications, the proofs are automatic by \blast\ with the theorem \no_nonce_NS1_NS2\. The remaining theorems about the protocol are not hard to prove. The following one asserts a form of \emph{authenticity}: if \B\ has sent an instance of message~2 to~\A\ and has received the expected reply, then that reply really originated with~\A\. The proof is a simple induction. \ (*<*) by (blast intro: no_nonce_NS1_NS2) lemma B_trusts_NS3_lemma [rule_format]: "\A \ bad; B \ bad; evs \ ns_public\ \ Crypt (pubK B) (Nonce NB) \ parts (knows Spy evs) \ Says B A (Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\) \ set evs \ Says A B (Crypt (pubK B) (Nonce NB)) \ set evs" by (erule ns_public.induct, auto) (*>*) theorem B_trusts_NS3: "\Says B A (Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\) \ set evs; Says A' B (Crypt (pubK B) (Nonce NB)) \ set evs; A \ bad; B \ bad; evs \ ns_public\ \ Says A B (Crypt (pubK B) (Nonce NB)) \ set evs" (*<*) by (blast intro: B_trusts_NS3_lemma) (*** Overall guarantee for B ***) (*If NS3 has been sent and the nonce NB agrees with the nonce B joined with NA, then A initiated the run using NA.*) theorem B_trusts_protocol [rule_format]: "\A \ bad; B \ bad; evs \ ns_public\ \ Crypt (pubK B) (Nonce NB) \ parts (knows Spy evs) \ Says B A (Crypt (pubK A) \Nonce NA, Nonce NB, Agent B\) \ set evs \ Says A B (Crypt (pubK B) \Nonce NA, Agent A\) \ set evs" by (erule ns_public.induct, auto) (*>*) text \ From similar assumptions, we can prove that \A\ started the protocol run by sending an instance of message~1 involving the nonce~\NA\\@. For this theorem, the conclusion is @{thm [display] (concl) B_trusts_protocol [no_vars]} Analogous theorems can be proved for~\A\, stating that nonce~\NA\ remains secret and that message~2 really originates with~\B\. Even the flawed protocol establishes these properties for~\A\; the flaw only harms the second participant. \medskip Detailed information on this protocol verification technique can be found -elsewhere~@{cite "paulson-jcs"}, including proofs of an Internet -protocol~@{cite "paulson-tls"}. We must stress that the protocol discussed +elsewhere~\<^cite>\"paulson-jcs"\, including proofs of an Internet +protocol~\<^cite>\"paulson-tls"\. We must stress that the protocol discussed in this chapter is trivial. There are only three messages; no keys are exchanged; we merely have to prove that encrypted data remains secret. Real world protocols are much longer and distribute many secrets to their participants. To be realistic, the model has to include the possibility of keys being lost dynamically due to carelessness. If those keys have been used to encrypt other sensitive information, there may be cascading losses. We may still be able to establish a bound on the losses and to prove that other protocol runs function -correctly~@{cite "paulson-yahalom"}. Proofs of real-world protocols follow +correctly~\<^cite>\"paulson-yahalom"\. Proofs of real-world protocols follow the strategy illustrated above, but the subgoals can be much bigger and there are more of them. \index{protocols!security|)} \ (*<*)end(*>*) diff --git a/src/Doc/Tutorial/Types/Axioms.thy b/src/Doc/Tutorial/Types/Axioms.thy --- a/src/Doc/Tutorial/Types/Axioms.thy +++ b/src/Doc/Tutorial/Types/Axioms.thy @@ -1,258 +1,258 @@ (*<*)theory Axioms imports Overloading "../Setup" begin(*>*) subsection \Axioms\ text \Attaching axioms to our classes lets us reason on the level of classes. The results will be applicable to all types in a class, just as in axiomatic mathematics. \begin{warn} Proofs in this section use structured \emph{Isar} proofs, which are not -covered in this tutorial; but see @{cite "Nipkow-TYPES02"}.% +covered in this tutorial; but see \<^cite>\"Nipkow-TYPES02"\.% \end{warn}\ subsubsection \Semigroups\ text\We specify \emph{semigroups} as subclass of \<^class>\plus\:\ class semigroup = plus + assumes assoc: "(x \ y) \ z = x \ (y \ z)" text \\noindent This @{command class} specification requires that all instances of \<^class>\semigroup\ obey @{fact "assoc:"}~@{prop [source] "\x y z :: 'a::semigroup. (x \ y) \ z = x \ (y \ z)"}. We can use this class axiom to derive further abstract theorems relative to class \<^class>\semigroup\:\ lemma assoc_left: fixes x y z :: "'a::semigroup" shows "x \ (y \ z) = (x \ y) \ z" using assoc by (rule sym) text \\noindent The \<^class>\semigroup\ constraint on type \<^typ>\'a\ restricts instantiations of \<^typ>\'a\ to types of class \<^class>\semigroup\ and during the proof enables us to use the fact @{fact assoc} whose type parameter is itself constrained to class \<^class>\semigroup\. The main advantage of classes is that theorems can be proved in the abstract and freely reused for each instance. On instantiation, we have to give a proof that the given operations obey the class axioms:\ instantiation nat :: semigroup begin instance proof txt \\noindent The proof opens with a default proof step, which for instance judgements invokes method @{method intro_classes}.\ fix m n q :: nat show "(m \ n) \ q = m \ (n \ q)" by (induct m) simp_all qed end text \\noindent Again, the interesting things enter the stage with parametric types:\ instantiation prod :: (semigroup, semigroup) semigroup begin instance proof fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "'a::semigroup \ 'b::semigroup" show "p\<^sub>1 \ p\<^sub>2 \ p\<^sub>3 = p\<^sub>1 \ (p\<^sub>2 \ p\<^sub>3)" by (cases p\<^sub>1, cases p\<^sub>2, cases p\<^sub>3) (simp add: assoc) txt \\noindent Associativity of product semigroups is established using the hypothetical associativity @{fact assoc} of the type components, which holds due to the \<^class>\semigroup\ constraints imposed on the type components by the @{command instance} proposition. Indeed, this pattern often occurs with parametric types and type classes.\ qed end subsubsection \Monoids\ text \We define a subclass \monoidl\ (a semigroup with a left-hand neutral) by extending \<^class>\semigroup\ with one additional parameter \neutral\ together with its property:\ class monoidl = semigroup + fixes neutral :: "'a" ("\") assumes neutl: "\ \ x = x" text \\noindent Again, we prove some instances, by providing suitable parameter definitions and proofs for the additional specifications.\ instantiation nat :: monoidl begin definition neutral_nat_def: "\ = (0::nat)" instance proof fix n :: nat show "\ \ n = n" unfolding neutral_nat_def by simp qed end text \\noindent In contrast to the examples above, we here have both specification of class operations and a non-trivial instance proof. This covers products as well: \ instantiation prod :: (monoidl, monoidl) monoidl begin definition neutral_prod_def: "\ = (\, \)" instance proof fix p :: "'a::monoidl \ 'b::monoidl" show "\ \ p = p" by (cases p) (simp add: neutral_prod_def neutl) qed end text \\noindent Fully-fledged monoids are modelled by another subclass which does not add new parameters but tightens the specification:\ class monoid = monoidl + assumes neutr: "x \ \ = x" text \\noindent Corresponding instances for \<^typ>\nat\ and products are left as an exercise to the reader.\ subsubsection \Groups\ text \\noindent To finish our small algebra example, we add a \group\ class:\ class group = monoidl + fixes inv :: "'a \ 'a" ("\
_" [81] 80) assumes invl: "\
x \ x = \" text \\noindent We continue with a further example for abstract proofs relative to type classes:\ lemma left_cancel: fixes x y z :: "'a::group" shows "x \ y = x \ z \ y = z" proof assume "x \ y = x \ z" then have "\
x \ (x \ y) = \
x \ (x \ z)" by simp then have "(\
x \ x) \ y = (\
x \ x) \ z" by (simp add: assoc) then show "y = z" by (simp add: invl neutl) next assume "y = z" then show "x \ y = x \ z" by simp qed text \\noindent Any \group\ is also a \monoid\; this can be made explicit by claiming an additional subclass relation, together with a proof of the logical difference:\ instance group \ monoid proof fix x from invl have "\
x \ x = \" . then have "\
x \ (x \ \) = \
x \ x" by (simp add: neutl invl assoc [symmetric]) then show "x \ \ = x" by (simp add: left_cancel) qed text \\noindent The proof result is propagated to the type system, making \group\ an instance of \monoid\ by adding an additional edge to the graph of subclass relation; see also Figure~\ref{fig:subclass}. \begin{figure}[htbp] \begin{center} \small \unitlength 0.6mm \begin{picture}(40,60)(0,0) \put(20,60){\makebox(0,0){\semigroup\}} \put(20,40){\makebox(0,0){\monoidl\}} \put(00,20){\makebox(0,0){\monoid\}} \put(40,00){\makebox(0,0){\group\}} \put(20,55){\vector(0,-1){10}} \put(15,35){\vector(-1,-1){10}} \put(25,35){\vector(1,-3){10}} \end{picture} \hspace{8em} \begin{picture}(40,60)(0,0) \put(20,60){\makebox(0,0){\semigroup\}} \put(20,40){\makebox(0,0){\monoidl\}} \put(00,20){\makebox(0,0){\monoid\}} \put(40,00){\makebox(0,0){\group\}} \put(20,55){\vector(0,-1){10}} \put(15,35){\vector(-1,-1){10}} \put(05,15){\vector(3,-1){30}} \end{picture} \caption{Subclass relationship of monoids and groups: before and after establishing the relationship \group \ monoid\; transitive edges are left out.} \label{fig:subclass} \end{center} \end{figure} \ subsubsection \Inconsistencies\ text \The reader may be wondering what happens if we attach an inconsistent set of axioms to a class. So far we have always avoided to add new axioms to HOL for fear of inconsistencies and suddenly it seems that we are throwing all caution to the wind. So why is there no problem? The point is that by construction, all type variables in the axioms of a \isacommand{class} are automatically constrained with the class being defined (as shown for axiom @{thm[source]refl} above). These constraints are always carried around and Isabelle takes care that they are never lost, unless the type variable is instantiated with a type that has been shown to belong to that class. Thus you may be able to prove \<^prop>\False\ from your axioms, but Isabelle will remind you that this theorem has the hidden hypothesis that the class is non-empty. Even if each individual class is consistent, intersections of (unrelated) classes readily become inconsistent in practice. Now we know this need not worry us.\ subsubsection\Syntactic Classes and Predefined Overloading\ text \In our algebra example, we have started with a \emph{syntactic class} \<^class>\plus\ which only specifies operations but no axioms; it would have been also possible to start immediately with class \<^class>\semigroup\, specifying the \\\ operation and associativity at the same time. Which approach is more appropriate depends. Usually it is more convenient to introduce operations and axioms in the same class: then the type checker will automatically insert the corresponding class constraints whenever the operations occur, reducing the need of manual annotations. However, when operations are decorated with popular syntax, syntactic classes can be an option to re-use the syntax in different contexts; this is indeed the way most overloaded constants in HOL are introduced, of which the most important are listed in Table~\ref{tab:overloading} in the appendix. Section \ref{sec:numeric-classes} covers a range of corresponding classes \emph{with} axioms. Further note that classes may contain axioms but \emph{no} operations. An example is class \<^class>\finite\ from theory \<^theory>\HOL.Finite_Set\ which specifies a type to be finite: @{lemma [source] "finite (UNIV :: 'a::finite set)" by (fact finite_UNIV)}.\ (*<*)end(*>*) diff --git a/src/Doc/Tutorial/Types/Records.thy b/src/Doc/Tutorial/Types/Records.thy --- a/src/Doc/Tutorial/Types/Records.thy +++ b/src/Doc/Tutorial/Types/Records.thy @@ -1,394 +1,394 @@ section \Records \label{sec:records}\ (*<*) theory Records imports Main begin (*>*) text \ \index{records|(}% Records are familiar from programming languages. A record of $n$ fields is essentially an $n$-tuple, but the record's components have names, which can make expressions easier to read and reduces the risk of confusing one field for another. A record of Isabelle/HOL covers a collection of fields, with select and update operations. Each field has a specified type, which may be polymorphic. The field names are part of the record type, and the order of the fields is significant --- as it is in Pascal but not in Standard ML. If two different record types have field names in common, then the ambiguity is resolved in the usual way, by qualified names. Record types can also be defined by extending other record types. Extensible records make use of the reserved pseudo-field \cdx{more}, which is present in every record type. Generic record operations work on all possible extensions of a given type scheme; polymorphism takes care of structural sub-typing behind the scenes. There are also explicit coercion functions between fixed record types. \ subsection \Record Basics\ text \ Record types are not primitive in Isabelle and have a delicate - internal representation @{cite "NaraschewskiW-TPHOLs98"}, based on + internal representation \<^cite>\"NaraschewskiW-TPHOLs98"\, based on nested copies of the primitive product type. A \commdx{record} declaration introduces a new record type scheme by specifying its fields, which are packaged internally to hold up the perception of the record as a distinguished entity. Here is a simple example: \ record point = Xcoord :: int Ycoord :: int text \\noindent Records of type \<^typ>\point\ have two fields named \<^const>\Xcoord\ and \<^const>\Ycoord\, both of type~\<^typ>\int\. We now define a constant of type \<^typ>\point\: \ definition pt1 :: point where "pt1 \ (| Xcoord = 999, Ycoord = 23 |)" text \\noindent We see above the ASCII notation for record brackets. You can also use the symbolic brackets \\\ and \\\. Record type expressions can be also written directly with individual fields. The type name above is merely an abbreviation. \ definition pt2 :: "\Xcoord :: int, Ycoord :: int\" where "pt2 \ \Xcoord = -45, Ycoord = 97\" text \ For each field, there is a \emph{selector}\index{selector!record} function of the same name. For example, if \p\ has type \<^typ>\point\ then \Xcoord p\ denotes the value of the \Xcoord\ field of~\p\. Expressions involving field selection of explicit records are simplified automatically: \ lemma "Xcoord \Xcoord = a, Ycoord = b\ = a" by simp text \ The \emph{update}\index{update!record} operation is functional. For example, \<^term>\p\Xcoord := 0\\ is a record whose \<^const>\Xcoord\ value is zero and whose \<^const>\Ycoord\ value is copied from~\p\. Updates of explicit records are also simplified automatically: \ lemma "\Xcoord = a, Ycoord = b\\Xcoord := 0\ = \Xcoord = 0, Ycoord = b\" by simp text \ \begin{warn} Field names are declared as constants and can no longer be used as variables. It would be unwise, for example, to call the fields of type \<^typ>\point\ simply \x\ and~\y\. \end{warn} \ subsection \Extensible Records and Generic Operations\ text \ \index{records!extensible|(}% Now, let us define coloured points (type \cpoint\) to be points extended with a field \col\ of type \colour\: \ datatype colour = Red | Green | Blue record cpoint = point + col :: colour text \\noindent The fields of this new type are \<^const>\Xcoord\, \Ycoord\ and \col\, in that order. \ definition cpt1 :: cpoint where "cpt1 \ \Xcoord = 999, Ycoord = 23, col = Green\" text \ We can define generic operations that work on arbitrary instances of a record scheme, e.g.\ covering \<^typ>\point\, \<^typ>\cpoint\, and any further extensions. Every record structure has an implicit pseudo-field, \cdx{more}, that keeps the extension as an explicit value. Its type is declared as completely polymorphic:~\<^typ>\'a\. When a fixed record value is expressed using just its standard fields, the value of \more\ is implicitly set to \()\, the empty tuple, which has type \<^typ>\unit\. Within the record brackets, you can refer to the \more\ field by writing ``\\\'' (three dots): \ lemma "Xcoord \Xcoord = a, Ycoord = b, \ = p\ = a" by simp text \ This lemma applies to any record whose first two fields are \Xcoord\ and~\<^const>\Ycoord\. Note that \\Xcoord = a, Ycoord = b, \ = ()\\ is exactly the same as \\Xcoord = a, Ycoord = b\\. Selectors and updates are always polymorphic wrt.\ the \more\ part of a record scheme, its value is just ignored (for select) or copied (for update). The \more\ pseudo-field may be manipulated directly as well, but the identifier needs to be qualified: \ lemma "point.more cpt1 = \col = Green\" by (simp add: cpt1_def) text \\noindent We see that the colour part attached to this \<^typ>\point\ is a rudimentary record in its own right, namely \\col = Green\\. In order to select or update \col\, this fragment needs to be put back into the context of the parent type scheme, say as \more\ part of another \<^typ>\point\. To define generic operations, we need to know a bit more about records. Our definition of \<^typ>\point\ above has generated two type abbreviations: \medskip \begin{tabular}{l} \<^typ>\point\~\=\~\\Xcoord :: int, Ycoord :: int\\ \\ \<^typ>\'a point_scheme\~\=\~\\Xcoord :: int, Ycoord :: int, \ :: 'a\\ \\ \end{tabular} \medskip \noindent Type \<^typ>\point\ is for fixed records having exactly the two fields \<^const>\Xcoord\ and~\Ycoord\, while the polymorphic type \<^typ>\'a point_scheme\ comprises all possible extensions to those two fields. Note that \<^typ>\unit point_scheme\ coincides with \<^typ>\point\, and \<^typ>\\col :: colour\ point_scheme\ with \<^typ>\cpoint\. In the following example we define two operations --- methods, if we regard records as objects --- to get and set any point's \Xcoord\ field. \ definition getX :: "'a point_scheme \ int" where "getX r \ Xcoord r" definition setX :: "'a point_scheme \ int \ 'a point_scheme" where "setX r a \ r\Xcoord := a\" text \ Here is a generic method that modifies a point, incrementing its \<^const>\Xcoord\ field. The \Ycoord\ and \more\ fields are copied across. It works for any record type scheme derived from \<^typ>\point\ (including \<^typ>\cpoint\ etc.): \ definition incX :: "'a point_scheme \ 'a point_scheme" where "incX r \ \Xcoord = Xcoord r + 1, Ycoord = Ycoord r, \ = point.more r\" text \ Generic theorems can be proved about generic methods. This trivial lemma relates \<^const>\incX\ to \getX\ and \setX\: \ lemma "incX r = setX r (getX r + 1)" by (simp add: getX_def setX_def incX_def) text \ \begin{warn} If you use the symbolic record brackets \\\ and \\\, then you must also use the symbolic ellipsis, ``\\\'', rather than three consecutive periods, ``\...\''. Mixing the ASCII and symbolic versions causes a syntax error. (The two versions are more distinct on screen than they are on paper.) \end{warn}% \index{records!extensible|)} \ subsection \Record Equality\ text \ Two records are equal\index{equality!of records} if all pairs of corresponding fields are equal. Concrete record equalities are simplified automatically: \ lemma "(\Xcoord = a, Ycoord = b\ = \Xcoord = a', Ycoord = b'\) = (a = a' \ b = b')" by simp text \ The following equality is similar, but generic, in that \r\ can be any instance of \<^typ>\'a point_scheme\: \ lemma "r\Xcoord := a, Ycoord := b\ = r\Ycoord := b, Xcoord := a\" by simp text \\noindent We see above the syntax for iterated updates. We could equivalently have written the left-hand side as \r\Xcoord := a\\Ycoord := b\\. Record equality is \emph{extensional}: \index{extensionality!for records} a record is determined entirely by the values of its fields. \ lemma "r = \Xcoord = Xcoord r, Ycoord = Ycoord r\" by simp text \\noindent The generic version of this equality includes the pseudo-field \more\: \ lemma "r = \Xcoord = Xcoord r, Ycoord = Ycoord r, \ = point.more r\" by simp text \ The simplifier can prove many record equalities automatically, but general equality reasoning can be tricky. Consider proving this obvious fact: \ lemma "r\Xcoord := a\ = r\Xcoord := a'\ \ a = a'" apply simp? oops text \\noindent Here the simplifier can do nothing, since general record equality is not eliminated automatically. One way to proceed is by an explicit forward step that applies the selector \<^const>\Xcoord\ to both sides of the assumed record equality: \ lemma "r\Xcoord := a\ = r\Xcoord := a'\ \ a = a'" apply (drule_tac f = Xcoord in arg_cong) txt \@{subgoals [display, indent = 0, margin = 65]} Now, \simp\ will reduce the assumption to the desired conclusion.\ apply simp done text \ The \cases\ method is preferable to such a forward proof. We state the desired lemma again: \ lemma "r\Xcoord := a\ = r\Xcoord := a'\ \ a = a'" txt \The \methdx{cases} method adds an equality to replace the named record term by an explicit record expression, listing all fields. It even includes the pseudo-field \more\, since the record equality stated here is generic for all extensions.\ apply (cases r) txt \@{subgoals [display, indent = 0, margin = 65]} Again, \simp\ finishes the proof. Because \r\ is now represented as an explicit record construction, the updates can be applied and the record equality can be replaced by equality of the corresponding fields (due to injectivity).\ apply simp done text \ The generic cases method does not admit references to locally bound parameters of a goal. In longer proof scripts one might have to fall back on the primitive \rule_tac\ used together with the internal field representation rules of records. The above use of \(cases r)\ would become \(rule_tac r = r in point.cases_scheme)\. \ subsection \Extending and Truncating Records\ text \ Each record declaration introduces a number of derived operations to refer collectively to a record's fields and to convert between fixed record types. They can, for instance, convert between types \<^typ>\point\ and \<^typ>\cpoint\. We can add a colour to a point or convert a \<^typ>\cpoint\ to a \<^typ>\point\ by forgetting its colour. \begin{itemize} \item Function \cdx{make} takes as arguments all of the record's fields (including those inherited from ancestors). It returns the corresponding record. \item Function \cdx{fields} takes the record's very own fields and returns a record fragment consisting of just those fields. This may be filled into the \more\ part of the parent record scheme. \item Function \cdx{extend} takes two arguments: a record to be extended and a record containing the new fields. \item Function \cdx{truncate} takes a record (possibly an extension of the original record type) and returns a fixed record, removing any additional fields. \end{itemize} These functions provide useful abbreviations for standard record expressions involving constructors and selectors. The definitions, which are \emph{not} unfolded by default, are made available by the collective name of \defs\ (\point.defs\, \cpoint.defs\, etc.). For example, here are the versions of those functions generated for record \<^typ>\point\. We omit \point.fields\, which happens to be the same as \point.make\. @{thm [display, indent = 0, margin = 65] point.make_def [no_vars] point.extend_def [no_vars] point.truncate_def [no_vars]} Contrast those with the corresponding functions for record \<^typ>\cpoint\. Observe \cpoint.fields\ in particular. @{thm [display, indent = 0, margin = 65] cpoint.make_def [no_vars] cpoint.fields_def [no_vars] cpoint.extend_def [no_vars] cpoint.truncate_def [no_vars]} To demonstrate these functions, we declare a new coloured point by extending an ordinary point. Function \point.extend\ augments \pt1\ with a colour value, which is converted into an appropriate record fragment by \cpoint.fields\. \ definition cpt2 :: cpoint where "cpt2 \ point.extend pt1 (cpoint.fields Green)" text \ The coloured points \<^const>\cpt1\ and \cpt2\ are equal. The proof is trivial, by unfolding all the definitions. We deliberately omit the definition of~\pt1\ in order to reveal the underlying comparison on type \<^typ>\point\. \ lemma "cpt1 = cpt2" apply (simp add: cpt1_def cpt2_def point.defs cpoint.defs) txt \@{subgoals [display, indent = 0, margin = 65]}\ apply (simp add: pt1_def) done text \ In the example below, a coloured point is truncated to leave a point. We use the \truncate\ function of the target record. \ lemma "point.truncate cpt2 = pt1" by (simp add: pt1_def cpt2_def point.defs) text \ \begin{exercise} Extend record \<^typ>\cpoint\ to have a further field, \intensity\, of type~\<^typ>\nat\. Experiment with generic operations (using polymorphic selectors and updates) and explicit coercions (using \extend\, \truncate\ etc.) among the three record types. \end{exercise} \begin{exercise} (For Java programmers.) Model a small class hierarchy using records. \end{exercise} \index{records|)} \ (*<*) end (*>*) diff --git a/src/Doc/Tutorial/Types/Typedefs.thy b/src/Doc/Tutorial/Types/Typedefs.thy --- a/src/Doc/Tutorial/Types/Typedefs.thy +++ b/src/Doc/Tutorial/Types/Typedefs.thy @@ -1,234 +1,234 @@ (*<*)theory Typedefs imports Main begin(*>*) section\Introducing New Types\ text\\label{sec:adv-typedef} For most applications, a combination of predefined types like \<^typ>\bool\ and \\\ with recursive datatypes and records is quite sufficient. Very occasionally you may feel the need for a more advanced type. If you are certain that your type is not definable by any of the standard means, then read on. \begin{warn} Types in HOL must be non-empty; otherwise the quantifier rules would be unsound, because $\exists x.\ x=x$ is a theorem. \end{warn} \ subsection\Declaring New Types\ text\\label{sec:typedecl} \index{types!declaring|(}% \index{typedecl@\isacommand {typedecl} (command)}% The most trivial way of introducing a new type is by a \textbf{type declaration}:\ typedecl my_new_type text\\noindent This does not define \<^typ>\my_new_type\ at all but merely introduces its name. Thus we know nothing about this type, except that it is non-empty. Such declarations without definitions are useful if that type can be viewed as a parameter of the theory. A typical example is given in \S\ref{sec:VMC}, where we define a transition relation over an arbitrary type of states. In principle we can always get rid of such type declarations by making those types parameters of every other type, thus keeping the theory generic. In practice, however, the resulting clutter can make types hard to read. If you are looking for a quick and dirty way of introducing a new type together with its properties: declare the type and state its properties as axioms. Example: \ axiomatization where just_one: "\x::my_new_type. \y. x = y" text\\noindent However, we strongly discourage this approach, except at explorative stages of your development. It is extremely easy to write down contradictory sets of axioms, in which case you will be able to prove everything but it will mean nothing. In the example above, the axiomatic approach is unnecessary: a one-element type called \<^typ>\unit\ is already defined in HOL. \index{types!declaring|)} \ subsection\Defining New Types\ text\\label{sec:typedef} \index{types!defining|(}% \index{typedecl@\isacommand {typedef} (command)|(}% Now we come to the most general means of safely introducing a new type, the \textbf{type definition}. All other means, for example \isacommand{datatype}, are based on it. The principle is extremely simple: any non-empty subset of an existing type can be turned into a new type. More precisely, the new type is specified to be isomorphic to some non-empty subset of an existing type. Let us work a simple example, the definition of a three-element type. It is easily represented by the first three natural numbers: \ typedef three = "{0::nat, 1, 2}" txt\\noindent In order to enforce that the representing set on the right-hand side is non-empty, this definition actually starts a proof to that effect: @{subgoals[display,indent=0]} Fortunately, this is easy enough to show, even \isa{auto} could do it. In general, one has to provide a witness, in our case 0: \ apply(rule_tac x = 0 in exI) by simp text\ This type definition introduces the new type \<^typ>\three\ and asserts that it is a copy of the set \<^term>\{0::nat,1,2}\. This assertion is expressed via a bijection between the \emph{type} \<^typ>\three\ and the \emph{set} \<^term>\{0::nat,1,2}\. To this end, the command declares the following constants behind the scenes: \begin{center} \begin{tabular}{rcl} \<^term>\Rep_three\ &::& \<^typ>\three \ nat\\\ \<^term>\Abs_three\ &::& \<^typ>\nat \ three\ \end{tabular} \end{center} The situation is best summarized with the help of the following diagram, where squares denote types and the irregular region denotes a set: \begin{center} \includegraphics[scale=.8]{typedef} \end{center} Finally, \isacommand{typedef} asserts that \<^term>\Rep_three\ is surjective on the subset and \<^term>\Abs_three\ and \<^term>\Rep_three\ are inverses of each other: \begin{center} \begin{tabular}{@ {}r@ {\qquad\qquad}l@ {}} @{thm Rep_three[no_vars]} & (@{thm[source]Rep_three}) \\ @{thm Rep_three_inverse[no_vars]} & (@{thm[source]Rep_three_inverse}) \\ @{thm Abs_three_inverse[no_vars]} & (@{thm[source]Abs_three_inverse}) \end{tabular} \end{center} % From this example it should be clear what \isacommand{typedef} does in general given a name (here \three\) and a set (here \<^term>\{0::nat,1,2}\). Our next step is to define the basic functions expected on the new type. Although this depends on the type at hand, the following strategy works well: \begin{itemize} \item define a small kernel of basic functions that can express all other functions you anticipate. \item define the kernel in terms of corresponding functions on the representing type using \<^term>\Abs\ and \<^term>\Rep\ to convert between the two levels. \end{itemize} In our example it suffices to give the three elements of type \<^typ>\three\ names: \ definition A :: three where "A \ Abs_three 0" definition B :: three where "B \ Abs_three 1" definition C :: three where "C \ Abs_three 2" text\ So far, everything was easy. But it is clear that reasoning about \<^typ>\three\ will be hell if we have to go back to \<^typ>\nat\ every time. Thus our aim must be to raise our level of abstraction by deriving enough theorems about type \<^typ>\three\ to characterize it completely. And those theorems should be phrased in terms of \<^term>\A\, \<^term>\B\ and \<^term>\C\, not \<^term>\Abs_three\ and \<^term>\Rep_three\. Because of the simplicity of the example, we merely need to prove that \<^term>\A\, \<^term>\B\ and \<^term>\C\ are distinct and that they exhaust the type. In processing our \isacommand{typedef} declaration, Isabelle proves several helpful lemmas. The first two express injectivity of \<^term>\Rep_three\ and \<^term>\Abs_three\: \begin{center} \begin{tabular}{@ {}r@ {\qquad}l@ {}} @{thm Rep_three_inject[no_vars]} & (@{thm[source]Rep_three_inject}) \\ \begin{tabular}{@ {}l@ {}} \\x \ {0, 1, 2}; y \ {0, 1, 2} \\ \\ \\ (Abs_three x = Abs_three y) = (x = y)\ \end{tabular} & (@{thm[source]Abs_three_inject}) \\ \end{tabular} \end{center} The following ones allow to replace some \x::three\ by \Abs_three(y::nat)\, and conversely \<^term>\y\ by \<^term>\Rep_three x\: \begin{center} \begin{tabular}{@ {}r@ {\qquad}l@ {}} @{thm Rep_three_cases[no_vars]} & (@{thm[source]Rep_three_cases}) \\ @{thm Abs_three_cases[no_vars]} & (@{thm[source]Abs_three_cases}) \\ @{thm Rep_three_induct[no_vars]} & (@{thm[source]Rep_three_induct}) \\ @{thm Abs_three_induct[no_vars]} & (@{thm[source]Abs_three_induct}) \\ \end{tabular} \end{center} These theorems are proved for any type definition, with \three\ replaced by the name of the type in question. Distinctness of \<^term>\A\, \<^term>\B\ and \<^term>\C\ follows immediately if we expand their definitions and rewrite with the injectivity of \<^term>\Abs_three\: \ lemma "A \ B \ B \ A \ A \ C \ C \ A \ B \ C \ C \ B" by(simp add: Abs_three_inject A_def B_def C_def) text\\noindent Of course we rely on the simplifier to solve goals like \<^prop>\(0::nat) \ 1\. The fact that \<^term>\A\, \<^term>\B\ and \<^term>\C\ exhaust type \<^typ>\three\ is best phrased as a case distinction theorem: if you want to prove \<^prop>\P x\ (where \<^term>\x\ is of type \<^typ>\three\) it suffices to prove \<^prop>\P A\, \<^prop>\P B\ and \<^prop>\P C\:\ lemma three_cases: "\ P A; P B; P C \ \ P x" txt\\noindent Again this follows easily using the induction principle stemming from the type definition:\ apply(induct_tac x) txt\ @{subgoals[display,indent=0]} Simplification leads to the disjunction \<^prop>\y = 0 \ y = 1 \ y = (2::nat)\ which \isa{auto} separates into three subgoals, each of which is easily solved by simplification:\ apply(auto simp add: A_def B_def C_def) done text\\noindent This concludes the derivation of the characteristic theorems for type \<^typ>\three\. The attentive reader has realized long ago that the above lengthy definition can be collapsed into one line: \ datatype better_three = A | B | C text\\noindent In fact, the \isacommand{datatype} command performs internally more or less the same derivations as we did, which gives you some idea what life would be like without \isacommand{datatype}. Although \<^typ>\three\ could be defined in one line, we have chosen this example to demonstrate \isacommand{typedef} because its simplicity makes the key concepts particularly easy to grasp. If you would like to see a non-trivial example that cannot be defined more directly, we recommend the -definition of \emph{finite multisets} in the Library~@{cite "HOL-Library"}. +definition of \emph{finite multisets} in the Library~\<^cite>\"HOL-Library"\. Let us conclude by summarizing the above procedure for defining a new type. Given some abstract axiomatic description $P$ of a type $ty$ in terms of a set of functions $F$, this involves three steps: \begin{enumerate} \item Find an appropriate type $\tau$ and subset $A$ which has the desired properties $P$, and make a type definition based on this representation. \item Define the required functions $F$ on $ty$ by lifting analogous functions on the representation via $Abs_ty$ and $Rep_ty$. \item Prove that $P$ holds for $ty$ by lifting $P$ from the representation. \end{enumerate} You can now forget about the representation and work solely in terms of the abstract functions $F$ and properties $P$.% \index{typedecl@\isacommand {typedef} (command)|)}% \index{types!defining|)} \ (*<*)end(*>*) diff --git a/src/Doc/Typeclass_Hierarchy/Typeclass_Hierarchy.thy b/src/Doc/Typeclass_Hierarchy/Typeclass_Hierarchy.thy --- a/src/Doc/Typeclass_Hierarchy/Typeclass_Hierarchy.thy +++ b/src/Doc/Typeclass_Hierarchy/Typeclass_Hierarchy.thy @@ -1,368 +1,368 @@ theory Typeclass_Hierarchy imports Setup begin section \Introduction\ text \ The {Isabelle/HOL} type-class hierarchy entered the stage in a quite ancient era -- first related \<^emph>\NEWS\ entries date back to release {Isabelle99-1}. Since then, there have been ongoing modifications and additions, leading to ({Isabelle2016}) more than 180 type-classes. This sheer complexity makes access and understanding of that type-class hierarchy difficult and involved, let alone maintenance. The purpose of this primer is to shed some light on this, not only on the mere ingredients, but also on the design principles which have evolved and proven useful over time. \ section \Foundations\ subsection \Locales and type classes\ text \ Some familiarity with the Isabelle module system is assumed: defining locales and type-classes, interpreting locales and instantiating type-classes, adding relationships between locales (@{command sublocale}) and type-classes (@{command subclass}). Handy introductions are the - respective tutorials @{cite "isabelle-locale"} - @{cite "isabelle-classes"}. + respective tutorials \<^cite>\"isabelle-locale"\ + \<^cite>\"isabelle-classes"\. \ subsection \Strengths and restrictions of type classes\ text \ The primary motivation for using type classes in {Isabelle/HOL} always have been numerical types, which form an inclusion chain: \begin{center} \<^typ>\nat\ \\\ \<^typ>\int\ \\\ \<^typ>\rat\ \\\ \<^typ>\real\ \\\ \<^typ>\complex\ \end{center} \noindent The inclusion \\\ means that any value of the numerical type to the left hand side mathematically can be transferred to the numerical type on the right hand side. How to accomplish this given the quite restrictive type system - of {Isabelle/HOL}? Paulson @{cite "paulson-numerical"} explains + of {Isabelle/HOL}? Paulson \<^cite>\"paulson-numerical"\ explains that each numerical type has some characteristic properties which define an characteristic algebraic structure; \\\ then corresponds to specialization of the corresponding characteristic algebraic structures. These algebraic structures are expressed using algebraic type classes and embeddings of numerical types into them: \begin{center}\begin{tabular}{lccc} \<^term>\of_nat\ \::\ & \<^typ>\nat\ & \\\ & @{typ [source] "'a::semiring_1"} \\ & \\\ & & \\\ \\ \<^term>\of_int\ \::\ & \<^typ>\int\ & \\\ & @{typ [source] "'a::ring_1"} \\ & \\\ & & \\\ \\ \<^term>\of_rat\ \::\ & \<^typ>\rat\ & \\\ & @{typ [source] "'a::field_char_0"} \\ & \\\ & & \\\ \\ \<^term>\of_real\ \::\ & \<^typ>\real\ & \\\ & @{typ [source] "'a::real_algebra_1"} \\ & \\\ \\ & \<^typ>\complex\ \end{tabular}\end{center} \noindent \d \ c\ means that \c\ is subclass of \d\. Hence each characteristic embedding \of_num\ can transform any value of type \num\ to any numerical type further up in the inclusion chain. This canonical example exhibits key strengths of type classes: \<^item> Sharing of operations and facts among different types, hence also sharing of notation and names: there is only one plus operation using infix syntax \+\, only one zero written \0\, and neutrality (@{thm (frugal_sorts) add_0_left [all, no_vars]} and @{thm (frugal_sorts) add_0_right [all, no_vars]}) is referred to uniformly by names @{fact add_0_left} and @{fact add_0_right}. \<^item> Proof tool setups are shared implicitly: @{fact add_0} and @{fact add_0_right} are simplification rules by default. \<^item> Hence existing proofs about particular numerical types are often easy to generalize to algebraic structures, given that they do not depend on specific properties of those numerical types. Considerable restrictions include: \<^item> Type class operations are restricted to one type parameter; this is insufficient e.g. for expressing a unified power operation adequately (see \secref{sec:power}). \<^item> Parameters are fixed over the whole type class hierarchy and cannot be refined in specific situations: think of integral domains with a predicate \<^term>\is_unit\; for natural numbers, this degenerates to the much simpler @{term [source] "HOL.equal (1::nat)"} but facts refer to \<^term>\is_unit\ nonetheless. \<^item> Type classes are not apt for meta-theory. There is no practically usable way to express that the units of an integral domain form a multiplicative group using type classes. But see session \HOL-Algebra\ which provides locales with an explicit carrier. \ subsection \Navigating the hierarchy\ text \ An indispensable tool to inspect the class hierarchy is @{command class_deps} which displays the graph of classes, optionally showing the logical content for each class also. Optional parameters restrict the graph to a particular segment which is useful to get a graspable view. See - the Isar reference manual @{cite "isabelle-isar-ref"} for details. + the Isar reference manual \<^cite>\"isabelle-isar-ref"\ for details. \ section \The hierarchy\ subsection \Syntactic type classes \label{sec:syntactic-type-class}\ text \ At the top of the hierarchy there are a couple of syntactic type classes, ie. classes with operations but with no axioms, most notably: \<^item> @{command class} \<^class>\plus\ with @{term [source] "(a::'a::plus) + b"} \<^item> @{command class} \<^class>\zero\ with @{term [source] "0::'a::zero"} \<^item> @{command class} \<^class>\times\ with @{term [source] "(a::'a::times) * b"} \<^item> @{command class} \<^class>\one\ with @{term [source] "1::'a::one"} \noindent Before the introduction of the @{command class} statement in - Isabelle @{cite "Haftmann-Wenzel:2006:classes"} it was impossible + Isabelle \<^cite>\"Haftmann-Wenzel:2006:classes"\ it was impossible to define operations with associated axioms in the same class, hence there were always pairs of syntactic and logical type classes. This restriction is lifted nowadays, but there are still reasons to maintain syntactic type classes: \<^item> Syntactic type classes allow generic notation to be used regardless of a particular logical interpretation; e.g. although multiplication \*\ is usually associative, there are examples where it is not (e.g. octonions), and leaving \*\ without axioms allows to re-use this syntax by means of type class instantiation also for such exotic examples. \<^item> Type classes might share operations but not necessarily axioms on them, e.g. \<^term>\gcd\ (see \secref{sec:gcd}). Hence their common base is a syntactic type class. \noindent However syntactic type classes should only be used with striking cause. Otherwise there is risk for confusion if the notation suggests properties which do not hold without particular constraints. This can be illustrated using numerals (see \secref{sec:numerals}): @{lemma "2 + 2 = 4" by simp} is provable without further ado, and this also meets the typical expectation towards a numeral notation; in more ancient releases numerals were purely syntactic and \<^prop>\2 + 2 = 4\ was not provable without particular type constraints. \ subsection \Additive and multiplicative semigroups and monoids\ text \ In common literature, notation for semigroups and monoids is either multiplicative \(*, 1)\ or additive \(+, 0)\ with underlying properties isomorphic. In {Isabelle/HOL}, this is accomplished using the following abstract setup: \<^item> A \<^locale>\semigroup\ introduces an abstract binary associative operation. \<^item> A \<^locale>\monoid\ is an extension of \<^locale>\semigroup\ with a neutral element. \<^item> Both \<^locale>\semigroup\ and \<^locale>\monoid\ provide dedicated syntax for their operations \(\<^bold>*, \<^bold>1)\. This syntax is not visible on the global theory level but only for abstract reasoning inside the respective locale. \<^item> Concrete global syntax is added building on existing syntactic type classes \secref{sec:syntactic-type-class} using the following classes: \<^item> @{command class} \<^class>\semigroup_mult\ = \<^class>\times\ \<^item> @{command class} \<^class>\monoid_mult\ = \<^class>\one\ + \<^class>\semigroup_mult\ \<^item> @{command class} \<^class>\semigroup_add\ = \<^class>\plus\ \<^item> @{command class} \<^class>\monoid_add\ = \<^class>\zero\ + \<^class>\semigroup_add\ Locales \<^locale>\semigroup\ and \<^locale>\monoid\ are interpreted (using @{command sublocale}) into their corresponding type classes, with prefixes \add\ and \mult\; hence facts derived in \<^locale>\semigroup\ and \<^locale>\monoid\ are propagated simultaneously to \<^emph>\both\ using a consistent naming policy, ie. \<^item> @{fact semigroup.assoc}: @{thm (frugal_sorts) semigroup.assoc [all, no_vars]} \<^item> @{fact mult.assoc}: @{thm (frugal_sorts) mult.assoc [all, no_vars]} \<^item> @{fact add.assoc}: @{thm (frugal_sorts) add.assoc [all, no_vars]} \<^item> @{fact monoid.right_neutral}: @{thm (frugal_sorts) monoid.right_neutral [all, no_vars]} \<^item> @{fact mult.right_neutral}: @{thm (frugal_sorts) mult.right_neutral [all, no_vars]} \<^item> @{fact add.right_neutral}: @{thm (frugal_sorts) add.right_neutral [all, no_vars]} \<^item> Note that the syntax in \<^locale>\semigroup\ and \<^locale>\monoid\ is bold; this avoids clashes when writing properties inside one of these locales in presence of that global concrete type class syntax. \noindent That hierarchy extends in a straightforward manner to abelian semigroups and commutative monoids\footnote{The designation \<^emph>\abelian\ is quite standard concerning (semi)groups, but not for monoids}: \<^item> Locales \<^locale>\abel_semigroup\ and \<^locale>\comm_monoid\ add commutativity as property. \<^item> Concrete syntax emerges through \<^item> @{command class} \<^class>\ab_semigroup_add\ = \<^class>\semigroup_add\ \<^item> @{command class} \<^class>\ab_semigroup_mult\ = \<^class>\semigroup_mult\ \<^item> @{command class} \<^class>\comm_monoid_add\ = \<^class>\zero\ + \<^class>\ab_semigroup_add\ \<^item> @{command class} \<^class>\comm_monoid_mult\ = \<^class>\one\ + \<^class>\ab_semigroup_mult\ and corresponding interpretation of the locales above, yielding \<^item> @{fact abel_semigroup.commute}: @{thm (frugal_sorts) abel_semigroup.commute [all, no_vars]} \<^item> @{fact mult.commute}: @{thm (frugal_sorts) mult.commute [all, no_vars]} \<^item> @{fact add.commute}: @{thm (frugal_sorts) add.commute [all, no_vars]} \ paragraph \Named collection of theorems\ text \ Locale interpretation interacts smoothly with named collections of theorems as introduced by command @{command named_theorems}. In our example, rules concerning associativity and commutativity are no simplification rules by default since they desired orientation may vary depending on the situation. However, there is a collection @{fact ac_simps} where facts @{fact abel_semigroup.assoc}, @{fact abel_semigroup.commute} and @{fact abel_semigroup.left_commute} are declared as members. Due to interpretation, also @{fact mult.assoc}, @{fact mult.commute} and @{fact mult.left_commute} are also members of @{fact ac_simps}, as any corresponding facts stemming from interpretation of \<^locale>\abel_semigroup\. Hence adding @{fact ac_simps} to the simplification rules for a single method call uses all associativity and commutativity rules known by means of interpretation. \ subsection \Additive and multiplicative groups\ text \ The hierarchy for inverse group operations takes into account that there are weaker algebraic structures with only a partially inverse operation. E. g. the natural numbers have bounded subtraction \<^term>\m - (n::nat)\ which is only an inverse operation if \<^term>\m \ (n::nat)\; unary minus \-\ is pointless on the natural numbers. Hence for both additive and multiplicative notation there are syntactic classes for inverse operations, both unary and binary: \<^item> @{command class} \<^class>\minus\ with @{term [source] "(a::'a::minus) - b"} \<^item> @{command class} \<^class>\uminus\ with @{term [source] "- a::'a::uminus"} \<^item> @{command class} \<^class>\divide\ with @{term [source] "(a::'a::divide) div b"} \<^item> @{command class} \<^class>\inverse\ = \<^class>\divide\ with @{term [source] "inverse a::'a::inverse"} \\ and @{term [source] "(a::'a::inverse) / b"} \noindent Here \<^class>\inverse\ specializes the ``partial'' syntax @{term [source] "a div b"} to the more specific @{term [source] "a / b"}. Semantic properties are added by \<^item> @{command class} \<^class>\cancel_ab_semigroup_add\ = \<^class>\ab_semigroup_add\ + \<^class>\minus\ \<^item> @{command class} \<^class>\cancel_comm_monoid_add\ = \<^class>\cancel_ab_semigroup_add\ + \<^class>\comm_monoid_add\ \noindent which specify a minimal binary partially inverse operation as \<^item> @{fact add_diff_cancel_left'}: @{thm (frugal_sorts) add_diff_cancel_left' [all, no_vars]} \<^item> @{fact diff_diff_add}: @{thm (frugal_sorts) diff_diff_add [all, no_vars]} \noindent which in turn allow to derive facts like \<^item> @{fact add_left_imp_eq}: @{thm (frugal_sorts) add_left_imp_eq [all, no_vars]} \noindent The total inverse operation is established as follows: \<^item> Locale \<^locale>\group\ extends the abstract hierarchy with the inverse operation. \<^item> The concrete additive inverse operation emerges through \<^item> @{command class} \<^class>\group_add\ = \<^class>\minus\ + \<^class>\uminus\ + \<^class>\monoid_add\ (in \<^theory>\HOL.Groups\) \\ \<^item> @{command class} \<^class>\ab_group_add\ = \<^class>\minus\ + \<^class>\uminus\ + \<^class>\comm_monoid_add\ (in \<^theory>\HOL.Groups\) and corresponding interpretation of locale \<^locale>\group\, yielding e.g. \<^item> @{fact group.left_inverse}: @{thm (frugal_sorts) group.left_inverse [all, no_vars]} \<^item> @{fact add.left_inverse}: @{thm (frugal_sorts) add.left_inverse [all, no_vars]} \noindent There is no multiplicative counterpart. Why? In rings, the multiplicative group excludes the zero element, hence the inverse operation is not total. See further \secref{sec:rings}. \ paragraph \Mitigating against redundancy by default simplification rules\ text \ Inverse operations imposes some redundancy on the type class hierarchy: in a group with a total inverse operation, the unary operation is simpler and more primitive than the binary one; but we cannot eliminate the binary one in favour of a mere syntactic abbreviation since the binary one is vital to express a partial inverse operation. This is mitigated by providing suitable default simplification rules: expression involving the unary inverse operation are simplified to binary inverse operation whenever appropriate. The rationale is that simplification is a central device in explorative proving, where proof obligation remaining after certain default proof steps including simplification are inspected to get an idea what is missing to finish a proof. When preferable normal forms are encoded into default simplification rules, proof obligations after simplification are normalized and hence more proof-friendly. \ end diff --git a/src/HOL/Algebra/Group.thy b/src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy +++ b/src/HOL/Algebra/Group.thy @@ -1,1746 +1,1746 @@ (* Title: HOL/Algebra/Group.thy Author: Clemens Ballarin, started 4 February 2003 Based on work by Florian Kammueller, L C Paulson and Markus Wenzel. With additional contributions from Martin Baillon and Paulo Emílio de Vilhena. *) theory Group imports Complete_Lattice "HOL-Library.FuncSet" begin section \Monoids and Groups\ subsection \Definitions\ text \ - Definitions follow @{cite "Jacobson:1985"}. + Definitions follow \<^cite>\"Jacobson:1985"\. \ record 'a monoid = "'a partial_object" + mult :: "['a, 'a] \ 'a" (infixl "\\" 70) one :: 'a ("\\") definition m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" ("inv\ _" [81] 80) where "inv\<^bsub>G\<^esub> x = (THE y. y \ carrier G \ x \\<^bsub>G\<^esub> y = \\<^bsub>G\<^esub> \ y \\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub>)" definition Units :: "_ => 'a set" \ \The set of invertible elements\ where "Units G = {y. y \ carrier G \ (\x \ carrier G. x \\<^bsub>G\<^esub> y = \\<^bsub>G\<^esub> \ y \\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub>)}" locale monoid = fixes G (structure) assumes m_closed [intro, simp]: "\x \ carrier G; y \ carrier G\ \ x \ y \ carrier G" and m_assoc: "\x \ carrier G; y \ carrier G; z \ carrier G\ \ (x \ y) \ z = x \ (y \ z)" and one_closed [intro, simp]: "\ \ carrier G" and l_one [simp]: "x \ carrier G \ \ \ x = x" and r_one [simp]: "x \ carrier G \ x \ \ = x" lemma monoidI: fixes G (structure) assumes m_closed: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \ y \ carrier G" and one_closed: "\ \ carrier G" and m_assoc: "!!x y z. [| x \ carrier G; y \ carrier G; z \ carrier G |] ==> (x \ y) \ z = x \ (y \ z)" and l_one: "!!x. x \ carrier G ==> \ \ x = x" and r_one: "!!x. x \ carrier G ==> x \ \ = x" shows "monoid G" by (fast intro!: monoid.intro intro: assms) lemma (in monoid) Units_closed [dest]: "x \ Units G ==> x \ carrier G" by (unfold Units_def) fast lemma (in monoid) one_unique: assumes "u \ carrier G" and "\x. x \ carrier G \ u \ x = x" shows "u = \" using assms(2)[OF one_closed] r_one[OF assms(1)] by simp lemma (in monoid) inv_unique: assumes eq: "y \ x = \" "x \ y' = \" and G: "x \ carrier G" "y \ carrier G" "y' \ carrier G" shows "y = y'" proof - from G eq have "y = y \ (x \ y')" by simp also from G have "... = (y \ x) \ y'" by (simp add: m_assoc) also from G eq have "... = y'" by simp finally show ?thesis . qed lemma (in monoid) Units_m_closed [simp, intro]: assumes x: "x \ Units G" and y: "y \ Units G" shows "x \ y \ Units G" proof - from x obtain x' where x: "x \ carrier G" "x' \ carrier G" and xinv: "x \ x' = \" "x' \ x = \" unfolding Units_def by fast from y obtain y' where y: "y \ carrier G" "y' \ carrier G" and yinv: "y \ y' = \" "y' \ y = \" unfolding Units_def by fast from x y xinv yinv have "y' \ (x' \ x) \ y = \" by simp moreover from x y xinv yinv have "x \ (y \ y') \ x' = \" by simp moreover note x y ultimately show ?thesis unfolding Units_def by simp (metis m_assoc m_closed) qed lemma (in monoid) Units_one_closed [intro, simp]: "\ \ Units G" by (unfold Units_def) auto lemma (in monoid) Units_inv_closed [intro, simp]: "x \ Units G ==> inv x \ carrier G" apply (simp add: Units_def m_inv_def) by (metis (mono_tags, lifting) inv_unique the_equality) lemma (in monoid) Units_l_inv_ex: "x \ Units G ==> \y \ carrier G. y \ x = \" by (unfold Units_def) auto lemma (in monoid) Units_r_inv_ex: "x \ Units G ==> \y \ carrier G. x \ y = \" by (unfold Units_def) auto lemma (in monoid) Units_l_inv [simp]: "x \ Units G ==> inv x \ x = \" apply (unfold Units_def m_inv_def, simp) by (metis (mono_tags, lifting) inv_unique the_equality) lemma (in monoid) Units_r_inv [simp]: "x \ Units G ==> x \ inv x = \" by (metis (full_types) Units_closed Units_inv_closed Units_l_inv Units_r_inv_ex inv_unique) lemma (in monoid) inv_one [simp]: "inv \ = \" by (metis Units_one_closed Units_r_inv l_one monoid.Units_inv_closed monoid_axioms) lemma (in monoid) Units_inv_Units [intro, simp]: "x \ Units G ==> inv x \ Units G" proof - assume x: "x \ Units G" show "inv x \ Units G" by (auto simp add: Units_def intro: Units_l_inv Units_r_inv x Units_closed [OF x]) qed lemma (in monoid) Units_l_cancel [simp]: "[| x \ Units G; y \ carrier G; z \ carrier G |] ==> (x \ y = x \ z) = (y = z)" proof assume eq: "x \ y = x \ z" and G: "x \ Units G" "y \ carrier G" "z \ carrier G" then have "(inv x \ x) \ y = (inv x \ x) \ z" by (simp add: m_assoc Units_closed del: Units_l_inv) with G show "y = z" by simp next assume eq: "y = z" and G: "x \ Units G" "y \ carrier G" "z \ carrier G" then show "x \ y = x \ z" by simp qed lemma (in monoid) Units_inv_inv [simp]: "x \ Units G ==> inv (inv x) = x" proof - assume x: "x \ Units G" then have "inv x \ inv (inv x) = inv x \ x" by simp with x show ?thesis by (simp add: Units_closed del: Units_l_inv Units_r_inv) qed lemma (in monoid) inv_inj_on_Units: "inj_on (m_inv G) (Units G)" proof (rule inj_onI) fix x y assume G: "x \ Units G" "y \ Units G" and eq: "inv x = inv y" then have "inv (inv x) = inv (inv y)" by simp with G show "x = y" by simp qed lemma (in monoid) Units_inv_comm: assumes inv: "x \ y = \" and G: "x \ Units G" "y \ Units G" shows "y \ x = \" proof - from G have "x \ y \ x = x \ \" by (auto simp add: inv Units_closed) with G show ?thesis by (simp del: r_one add: m_assoc Units_closed) qed lemma (in monoid) carrier_not_empty: "carrier G \ {}" by auto (* Jacobson defines submonoid here. *) (* Jacobson defines the order of a monoid here. *) subsection \Groups\ text \ A group is a monoid all of whose elements are invertible. \ locale group = monoid + assumes Units: "carrier G <= Units G" lemma (in group) is_group [iff]: "group G" by (rule group_axioms) lemma (in group) is_monoid [iff]: "monoid G" by (rule monoid_axioms) theorem groupI: fixes G (structure) assumes m_closed [simp]: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \ y \ carrier G" and one_closed [simp]: "\ \ carrier G" and m_assoc: "!!x y z. [| x \ carrier G; y \ carrier G; z \ carrier G |] ==> (x \ y) \ z = x \ (y \ z)" and l_one [simp]: "!!x. x \ carrier G ==> \ \ x = x" and l_inv_ex: "!!x. x \ carrier G ==> \y \ carrier G. y \ x = \" shows "group G" proof - have l_cancel [simp]: "!!x y z. [| x \ carrier G; y \ carrier G; z \ carrier G |] ==> (x \ y = x \ z) = (y = z)" proof fix x y z assume eq: "x \ y = x \ z" and G: "x \ carrier G" "y \ carrier G" "z \ carrier G" with l_inv_ex obtain x_inv where xG: "x_inv \ carrier G" and l_inv: "x_inv \ x = \" by fast from G eq xG have "(x_inv \ x) \ y = (x_inv \ x) \ z" by (simp add: m_assoc) with G show "y = z" by (simp add: l_inv) next fix x y z assume eq: "y = z" and G: "x \ carrier G" "y \ carrier G" "z \ carrier G" then show "x \ y = x \ z" by simp qed have r_one: "!!x. x \ carrier G ==> x \ \ = x" proof - fix x assume x: "x \ carrier G" with l_inv_ex obtain x_inv where xG: "x_inv \ carrier G" and l_inv: "x_inv \ x = \" by fast from x xG have "x_inv \ (x \ \) = x_inv \ x" by (simp add: m_assoc [symmetric] l_inv) with x xG show "x \ \ = x" by simp qed have inv_ex: "\x. x \ carrier G \ \y \ carrier G. y \ x = \ \ x \ y = \" proof - fix x assume x: "x \ carrier G" with l_inv_ex obtain y where y: "y \ carrier G" and l_inv: "y \ x = \" by fast from x y have "y \ (x \ y) = y \ \" by (simp add: m_assoc [symmetric] l_inv r_one) with x y have r_inv: "x \ y = \" by simp from x y show "\y \ carrier G. y \ x = \ \ x \ y = \" by (fast intro: l_inv r_inv) qed then have carrier_subset_Units: "carrier G \ Units G" by (unfold Units_def) fast show ?thesis by standard (auto simp: r_one m_assoc carrier_subset_Units) qed lemma (in monoid) group_l_invI: assumes l_inv_ex: "!!x. x \ carrier G ==> \y \ carrier G. y \ x = \" shows "group G" by (rule groupI) (auto intro: m_assoc l_inv_ex) lemma (in group) Units_eq [simp]: "Units G = carrier G" proof show "Units G \ carrier G" by fast next show "carrier G \ Units G" by (rule Units) qed lemma (in group) inv_closed [intro, simp]: "x \ carrier G ==> inv x \ carrier G" using Units_inv_closed by simp lemma (in group) l_inv_ex [simp]: "x \ carrier G ==> \y \ carrier G. y \ x = \" using Units_l_inv_ex by simp lemma (in group) r_inv_ex [simp]: "x \ carrier G ==> \y \ carrier G. x \ y = \" using Units_r_inv_ex by simp lemma (in group) l_inv [simp]: "x \ carrier G ==> inv x \ x = \" by simp subsection \Cancellation Laws and Basic Properties\ lemma (in group) inv_eq_1_iff [simp]: assumes "x \ carrier G" shows "inv\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub> \ x = \\<^bsub>G\<^esub>" proof - have "x = \" if "inv x = \" proof - have "inv x \ x = \" using assms l_inv by blast then show "x = \" using that assms by simp qed then show ?thesis by auto qed lemma (in group) r_inv [simp]: "x \ carrier G ==> x \ inv x = \" by simp lemma (in group) right_cancel [simp]: "[| x \ carrier G; y \ carrier G; z \ carrier G |] ==> (y \ x = z \ x) = (y = z)" by (metis inv_closed m_assoc r_inv r_one) lemma (in group) inv_inv [simp]: "x \ carrier G ==> inv (inv x) = x" using Units_inv_inv by simp lemma (in group) inv_inj: "inj_on (m_inv G) (carrier G)" using inv_inj_on_Units by simp lemma (in group) inv_mult_group: "[| x \ carrier G; y \ carrier G |] ==> inv (x \ y) = inv y \ inv x" proof - assume G: "x \ carrier G" "y \ carrier G" then have "inv (x \ y) \ (x \ y) = (inv y \ inv x) \ (x \ y)" by (simp add: m_assoc) (simp add: m_assoc [symmetric]) with G show ?thesis by (simp del: l_inv Units_l_inv) qed lemma (in group) inv_comm: "[| x \ y = \; x \ carrier G; y \ carrier G |] ==> y \ x = \" by (rule Units_inv_comm) auto lemma (in group) inv_equality: "[|y \ x = \; x \ carrier G; y \ carrier G|] ==> inv x = y" using inv_unique r_inv by blast lemma (in group) inv_solve_left: "\ a \ carrier G; b \ carrier G; c \ carrier G \ \ a = inv b \ c \ c = b \ a" by (metis inv_equality l_inv_ex l_one m_assoc r_inv) lemma (in group) inv_solve_left': "\ a \ carrier G; b \ carrier G; c \ carrier G \ \ inv b \ c = a \ c = b \ a" by (metis inv_equality l_inv_ex l_one m_assoc r_inv) lemma (in group) inv_solve_right: "\ a \ carrier G; b \ carrier G; c \ carrier G \ \ a = b \ inv c \ b = a \ c" by (metis inv_equality l_inv_ex l_one m_assoc r_inv) lemma (in group) inv_solve_right': "\a \ carrier G; b \ carrier G; c \ carrier G\ \ b \ inv c = a \ b = a \ c" by (auto simp: m_assoc) subsection \Power\ consts pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr "[^]\" 75) overloading nat_pow == "pow :: [_, 'a, nat] => 'a" begin definition "nat_pow G a n = rec_nat \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) n" end lemma (in monoid) nat_pow_closed [intro, simp]: "x \ carrier G ==> x [^] (n::nat) \ carrier G" by (induct n) (simp_all add: nat_pow_def) lemma (in monoid) nat_pow_0 [simp]: "x [^] (0::nat) = \" by (simp add: nat_pow_def) lemma (in monoid) nat_pow_Suc [simp]: "x [^] (Suc n) = x [^] n \ x" by (simp add: nat_pow_def) lemma (in monoid) nat_pow_one [simp]: "\ [^] (n::nat) = \" by (induct n) simp_all lemma (in monoid) nat_pow_mult: "x \ carrier G ==> x [^] (n::nat) \ x [^] m = x [^] (n + m)" by (induct m) (simp_all add: m_assoc [THEN sym]) lemma (in monoid) nat_pow_comm: "x \ carrier G \ (x [^] (n::nat)) \ (x [^] (m :: nat)) = (x [^] m) \ (x [^] n)" using nat_pow_mult[of x n m] nat_pow_mult[of x m n] by (simp add: add.commute) lemma (in monoid) nat_pow_Suc2: "x \ carrier G \ x [^] (Suc n) = x \ (x [^] n)" using nat_pow_mult[of x 1 n] Suc_eq_plus1[of n] by (metis One_nat_def Suc_eq_plus1_left l_one nat.rec(1) nat_pow_Suc nat_pow_def) lemma (in monoid) nat_pow_pow: "x \ carrier G ==> (x [^] n) [^] m = x [^] (n * m::nat)" by (induct m) (simp, simp add: nat_pow_mult add.commute) lemma (in monoid) nat_pow_consistent: "x [^] (n :: nat) = x [^]\<^bsub>(G \ carrier := H \)\<^esub> n" unfolding nat_pow_def by simp lemma nat_pow_0 [simp]: "x [^]\<^bsub>G\<^esub> (0::nat) = \\<^bsub>G\<^esub>" by (simp add: nat_pow_def) lemma nat_pow_Suc [simp]: "x [^]\<^bsub>G\<^esub> (Suc n) = (x [^]\<^bsub>G\<^esub> n)\\<^bsub>G\<^esub> x" by (simp add: nat_pow_def) lemma (in group) nat_pow_inv: assumes "x \ carrier G" shows "(inv x) [^] (i :: nat) = inv (x [^] i)" proof (induction i) case 0 thus ?case by simp next case (Suc i) have "(inv x) [^] Suc i = ((inv x) [^] i) \ inv x" by simp also have " ... = (inv (x [^] i)) \ inv x" by (simp add: Suc.IH Suc.prems) also have " ... = inv (x \ (x [^] i))" by (simp add: assms inv_mult_group) also have " ... = inv (x [^] (Suc i))" using assms nat_pow_Suc2 by auto finally show ?case . qed overloading int_pow == "pow :: [_, 'a, int] => 'a" begin definition "int_pow G a z = (let p = rec_nat \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) in if z < 0 then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z))" end lemma int_pow_int: "x [^]\<^bsub>G\<^esub> (int n) = x [^]\<^bsub>G\<^esub> n" by(simp add: int_pow_def nat_pow_def) lemma pow_nat: assumes "i\0" shows "x [^]\<^bsub>G\<^esub> nat i = x [^]\<^bsub>G\<^esub> i" proof (cases i rule: int_cases) case (nonneg n) then show ?thesis by (simp add: int_pow_int) next case (neg n) then show ?thesis using assms by linarith qed lemma int_pow_0 [simp]: "x [^]\<^bsub>G\<^esub> (0::int) = \\<^bsub>G\<^esub>" by (simp add: int_pow_def) lemma int_pow_def2: "a [^]\<^bsub>G\<^esub> z = (if z < 0 then inv\<^bsub>G\<^esub> (a [^]\<^bsub>G\<^esub> (nat (-z))) else a [^]\<^bsub>G\<^esub> (nat z))" by (simp add: int_pow_def nat_pow_def) lemma (in group) int_pow_one [simp]: "\ [^] (z::int) = \" by (simp add: int_pow_def2) lemma (in group) int_pow_closed [intro, simp]: "x \ carrier G ==> x [^] (i::int) \ carrier G" by (simp add: int_pow_def2) lemma (in group) int_pow_1 [simp]: "x \ carrier G \ x [^] (1::int) = x" by (simp add: int_pow_def2) lemma (in group) int_pow_neg: "x \ carrier G \ x [^] (-i::int) = inv (x [^] i)" by (simp add: int_pow_def2) lemma (in group) int_pow_neg_int: "x \ carrier G \ x [^] -(int n) = inv (x [^] n)" by (simp add: int_pow_neg int_pow_int) lemma (in group) int_pow_mult: assumes "x \ carrier G" shows "x [^] (i + j::int) = x [^] i \ x [^] j" proof - have [simp]: "-i - j = -j - i" by simp show ?thesis by (auto simp: assms int_pow_def2 inv_solve_left inv_solve_right nat_add_distrib [symmetric] nat_pow_mult) qed lemma (in group) int_pow_inv: "x \ carrier G \ (inv x) [^] (i :: int) = inv (x [^] i)" by (metis int_pow_def2 nat_pow_inv) lemma (in group) int_pow_pow: assumes "x \ carrier G" shows "(x [^] (n :: int)) [^] (m :: int) = x [^] (n * m :: int)" proof (cases) assume n_ge: "n \ 0" thus ?thesis proof (cases) assume m_ge: "m \ 0" thus ?thesis using n_ge nat_pow_pow[OF assms, of "nat n" "nat m"] int_pow_def2 [where G=G] by (simp add: mult_less_0_iff nat_mult_distrib) next assume m_lt: "\ m \ 0" with n_ge show ?thesis apply (simp add: int_pow_def2 mult_less_0_iff) by (metis assms mult_minus_right n_ge nat_mult_distrib nat_pow_pow) qed next assume n_lt: "\ n \ 0" thus ?thesis proof (cases) assume m_ge: "m \ 0" have "inv x [^] (nat m * nat (- n)) = inv x [^] nat (- (m * n))" by (metis (full_types) m_ge mult_minus_right nat_mult_distrib) with m_ge n_lt show ?thesis by (simp add: int_pow_def2 mult_less_0_iff assms mult.commute nat_pow_inv nat_pow_pow) next assume m_lt: "\ m \ 0" thus ?thesis using n_lt by (auto simp: int_pow_def2 mult_less_0_iff assms nat_mult_distrib_neg nat_pow_inv nat_pow_pow) qed qed lemma (in group) int_pow_diff: "x \ carrier G \ x [^] (n - m :: int) = x [^] n \ inv (x [^] m)" by(simp only: diff_conv_add_uminus int_pow_mult int_pow_neg) lemma (in group) inj_on_multc: "c \ carrier G \ inj_on (\x. x \ c) (carrier G)" by(simp add: inj_on_def) lemma (in group) inj_on_cmult: "c \ carrier G \ inj_on (\x. c \ x) (carrier G)" by(simp add: inj_on_def) lemma (in monoid) group_commutes_pow: fixes n::nat shows "\x \ y = y \ x; x \ carrier G; y \ carrier G\ \ x [^] n \ y = y \ x [^] n" apply (induction n, auto) by (metis m_assoc nat_pow_closed) lemma (in monoid) pow_mult_distrib: assumes eq: "x \ y = y \ x" and xy: "x \ carrier G" "y \ carrier G" shows "(x \ y) [^] (n::nat) = x [^] n \ y [^] n" proof (induct n) case (Suc n) have "x \ (y [^] n \ y) = y [^] n \ x \ y" by (simp add: eq group_commutes_pow m_assoc xy) then show ?case using assms Suc.hyps m_assoc by auto qed auto lemma (in group) int_pow_mult_distrib: assumes eq: "x \ y = y \ x" and xy: "x \ carrier G" "y \ carrier G" shows "(x \ y) [^] (i::int) = x [^] i \ y [^] i" proof (cases i rule: int_cases) case (nonneg n) then show ?thesis by (metis eq int_pow_int pow_mult_distrib xy) next case (neg n) then show ?thesis unfolding neg apply (simp add: xy int_pow_neg_int del: of_nat_Suc) by (metis eq inv_mult_group local.nat_pow_Suc nat_pow_closed pow_mult_distrib xy) qed lemma (in group) pow_eq_div2: fixes m n :: nat assumes x_car: "x \ carrier G" assumes pow_eq: "x [^] m = x [^] n" shows "x [^] (m - n) = \" proof (cases "m < n") case False have "\ \ x [^] m = x [^] m" by (simp add: x_car) also have "\ = x [^] (m - n) \ x [^] n" using False by (simp add: nat_pow_mult x_car) also have "\ = x [^] (m - n) \ x [^] m" by (simp add: pow_eq) finally show ?thesis by (metis nat_pow_closed one_closed right_cancel x_car) qed simp subsection \Submonoids\ locale submonoid = \<^marker>\contributor \Martin Baillon\\ fixes H and G (structure) assumes subset: "H \ carrier G" and m_closed [intro, simp]: "\x \ H; y \ H\ \ x \ y \ H" and one_closed [simp]: "\ \ H" lemma (in submonoid) is_submonoid: \<^marker>\contributor \Martin Baillon\\ "submonoid H G" by (rule submonoid_axioms) lemma (in submonoid) mem_carrier [simp]: \<^marker>\contributor \Martin Baillon\\ "x \ H \ x \ carrier G" using subset by blast lemma (in submonoid) submonoid_is_monoid [intro]: \<^marker>\contributor \Martin Baillon\\ assumes "monoid G" shows "monoid (G\carrier := H\)" proof - interpret monoid G by fact show ?thesis by (simp add: monoid_def m_assoc) qed lemma submonoid_nonempty: \<^marker>\contributor \Martin Baillon\\ "~ submonoid {} G" by (blast dest: submonoid.one_closed) lemma (in submonoid) finite_monoid_imp_card_positive: \<^marker>\contributor \Martin Baillon\\ "finite (carrier G) ==> 0 < card H" proof (rule classical) assume "finite (carrier G)" and a: "~ 0 < card H" then have "finite H" by (blast intro: finite_subset [OF subset]) with is_submonoid a have "submonoid {} G" by simp with submonoid_nonempty show ?thesis by contradiction qed lemma (in monoid) monoid_incl_imp_submonoid : \<^marker>\contributor \Martin Baillon\\ assumes "H \ carrier G" and "monoid (G\carrier := H\)" shows "submonoid H G" proof (intro submonoid.intro[OF assms(1)]) have ab_eq : "\ a b. a \ H \ b \ H \ a \\<^bsub>G\carrier := H\\<^esub> b = a \ b" using assms by simp have "\a b. a \ H \ b \ H \ a \ b \ carrier (G\carrier := H\) " using assms ab_eq unfolding group_def using monoid.m_closed by fastforce thus "\a b. a \ H \ b \ H \ a \ b \ H" by simp show "\ \ H " using monoid.one_closed[OF assms(2)] assms by simp qed lemma (in monoid) inv_unique': \<^marker>\contributor \Martin Baillon\\ assumes "x \ carrier G" "y \ carrier G" shows "\ x \ y = \; y \ x = \ \ \ y = inv x" proof - assume "x \ y = \" and l_inv: "y \ x = \" hence unit: "x \ Units G" using assms unfolding Units_def by auto show "y = inv x" using inv_unique[OF l_inv Units_r_inv[OF unit] assms Units_inv_closed[OF unit]] . qed lemma (in monoid) m_inv_monoid_consistent: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "x \ Units (G \ carrier := H \)" and "submonoid H G" shows "inv\<^bsub>(G \ carrier := H \)\<^esub> x = inv x" proof - have monoid: "monoid (G \ carrier := H \)" using submonoid.submonoid_is_monoid[OF assms(2) monoid_axioms] . obtain y where y: "y \ H" "x \ y = \" "y \ x = \" using assms(1) unfolding Units_def by auto have x: "x \ H" and in_carrier: "x \ carrier G" "y \ carrier G" using y(1) submonoid.subset[OF assms(2)] assms(1) unfolding Units_def by auto show ?thesis using monoid.inv_unique'[OF monoid, of x y] x y using inv_unique'[OF in_carrier y(2-3)] by auto qed subsection \Subgroups\ locale subgroup = fixes H and G (structure) assumes subset: "H \ carrier G" and m_closed [intro, simp]: "\x \ H; y \ H\ \ x \ y \ H" and one_closed [simp]: "\ \ H" and m_inv_closed [intro,simp]: "x \ H \ inv x \ H" lemma (in subgroup) is_subgroup: "subgroup H G" by (rule subgroup_axioms) declare (in subgroup) group.intro [intro] lemma (in subgroup) mem_carrier [simp]: "x \ H \ x \ carrier G" using subset by blast lemma (in subgroup) subgroup_is_group [intro]: assumes "group G" shows "group (G\carrier := H\)" proof - interpret group G by fact have "Group.monoid (G\carrier := H\)" by (simp add: monoid_axioms submonoid.intro submonoid.submonoid_is_monoid subset) then show ?thesis by (rule monoid.group_l_invI) (auto intro: l_inv mem_carrier) qed lemma subgroup_is_submonoid: assumes "subgroup H G" shows "submonoid H G" using assms by (auto intro: submonoid.intro simp add: subgroup_def) lemma (in group) subgroup_Units: assumes "subgroup H G" shows "H \ Units (G \ carrier := H \)" using group.Units[OF subgroup.subgroup_is_group[OF assms group_axioms]] by simp lemma (in group) m_inv_consistent [simp]: assumes "subgroup H G" "x \ H" shows "inv\<^bsub>(G \ carrier := H \)\<^esub> x = inv x" using assms m_inv_monoid_consistent[OF _ subgroup_is_submonoid] subgroup_Units[of H] by auto lemma (in group) int_pow_consistent: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "subgroup H G" "x \ H" shows "x [^] (n :: int) = x [^]\<^bsub>(G \ carrier := H \)\<^esub> n" proof (cases) assume ge: "n \ 0" hence "x [^] n = x [^] (nat n)" using int_pow_def2 [of G] by auto also have " ... = x [^]\<^bsub>(G \ carrier := H \)\<^esub> (nat n)" using nat_pow_consistent by simp also have " ... = x [^]\<^bsub>(G \ carrier := H \)\<^esub> n" by (metis ge int_nat_eq int_pow_int) finally show ?thesis . next assume "\ n \ 0" hence lt: "n < 0" by simp hence "x [^] n = inv (x [^] (nat (- n)))" using int_pow_def2 [of G] by auto also have " ... = (inv x) [^] (nat (- n))" by (metis assms nat_pow_inv subgroup.mem_carrier) also have " ... = (inv\<^bsub>(G \ carrier := H \)\<^esub> x) [^]\<^bsub>(G \ carrier := H \)\<^esub> (nat (- n))" using m_inv_consistent[OF assms] nat_pow_consistent by auto also have " ... = inv\<^bsub>(G \ carrier := H \)\<^esub> (x [^]\<^bsub>(G \ carrier := H \)\<^esub> (nat (- n)))" using group.nat_pow_inv[OF subgroup.subgroup_is_group[OF assms(1) is_group]] assms(2) by auto also have " ... = x [^]\<^bsub>(G \ carrier := H \)\<^esub> n" by (simp add: int_pow_def2 lt) finally show ?thesis . qed text \ Since \<^term>\H\ is nonempty, it contains some element \<^term>\x\. Since it is closed under inverse, it contains \inv x\. Since it is closed under product, it contains \x \ inv x = \\. \ lemma (in group) one_in_subset: "[| H \ carrier G; H \ {}; \a \ H. inv a \ H; \a\H. \b\H. a \ b \ H |] ==> \ \ H" by force text \A characterization of subgroups: closed, non-empty subset.\ lemma (in group) subgroupI: assumes subset: "H \ carrier G" and non_empty: "H \ {}" and inv: "!!a. a \ H \ inv a \ H" and mult: "!!a b. \a \ H; b \ H\ \ a \ b \ H" shows "subgroup H G" proof (simp add: subgroup_def assms) show "\ \ H" by (rule one_in_subset) (auto simp only: assms) qed lemma (in group) subgroupE: assumes "subgroup H G" shows "H \ carrier G" and "H \ {}" and "\a. a \ H \ inv a \ H" and "\a b. \ a \ H; b \ H \ \ a \ b \ H" using assms unfolding subgroup_def[of H G] by auto declare monoid.one_closed [iff] group.inv_closed [simp] monoid.l_one [simp] monoid.r_one [simp] group.inv_inv [simp] lemma subgroup_nonempty: "\ subgroup {} G" by (blast dest: subgroup.one_closed) lemma (in subgroup) finite_imp_card_positive: "finite (carrier G) \ 0 < card H" using subset one_closed card_gt_0_iff finite_subset by blast lemma (in subgroup) subgroup_is_submonoid : \<^marker>\contributor \Martin Baillon\\ "submonoid H G" by (simp add: submonoid.intro subset) lemma (in group) submonoid_subgroupI : \<^marker>\contributor \Martin Baillon\\ assumes "submonoid H G" and "\a. a \ H \ inv a \ H" shows "subgroup H G" by (metis assms subgroup_def submonoid_def) lemma (in group) group_incl_imp_subgroup: \<^marker>\contributor \Martin Baillon\\ assumes "H \ carrier G" and "group (G\carrier := H\)" shows "subgroup H G" proof (intro submonoid_subgroupI[OF monoid_incl_imp_submonoid[OF assms(1)]]) show "monoid (G\carrier := H\)" using group_def assms by blast have ab_eq : "\ a b. a \ H \ b \ H \ a \\<^bsub>G\carrier := H\\<^esub> b = a \ b" using assms by simp fix a assume aH : "a \ H" have " inv\<^bsub>G\carrier := H\\<^esub> a \ carrier G" using assms aH group.inv_closed[OF assms(2)] by auto moreover have "\\<^bsub>G\carrier := H\\<^esub> = \" using assms monoid.one_closed ab_eq one_def by simp hence "a \\<^bsub>G\carrier := H\\<^esub> inv\<^bsub>G\carrier := H\\<^esub> a= \" using assms ab_eq aH group.r_inv[OF assms(2)] by simp hence "a \ inv\<^bsub>G\carrier := H\\<^esub> a= \" using aH assms group.inv_closed[OF assms(2)] ab_eq by simp ultimately have "inv\<^bsub>G\carrier := H\\<^esub> a = inv a" by (metis aH assms(1) contra_subsetD group.inv_inv is_group local.inv_equality) moreover have "inv\<^bsub>G\carrier := H\\<^esub> a \ H" using aH group.inv_closed[OF assms(2)] by auto ultimately show "inv a \ H" by auto qed subsection \Direct Products\ definition DirProd :: "_ \ _ \ ('a \ 'b) monoid" (infixr "\\" 80) where "G \\ H = \carrier = carrier G \ carrier H, mult = (\(g, h) (g', h'). (g \\<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h')), one = (\\<^bsub>G\<^esub>, \\<^bsub>H\<^esub>)\" lemma DirProd_monoid: assumes "monoid G" and "monoid H" shows "monoid (G \\ H)" proof - interpret G: monoid G by fact interpret H: monoid H by fact from assms show ?thesis by (unfold monoid_def DirProd_def, auto) qed text\Does not use the previous result because it's easier just to use auto.\ lemma DirProd_group: assumes "group G" and "group H" shows "group (G \\ H)" proof - interpret G: group G by fact interpret H: group H by fact show ?thesis by (rule groupI) (auto intro: G.m_assoc H.m_assoc G.l_inv H.l_inv simp add: DirProd_def) qed lemma carrier_DirProd [simp]: "carrier (G \\ H) = carrier G \ carrier H" by (simp add: DirProd_def) lemma one_DirProd [simp]: "\\<^bsub>G \\ H\<^esub> = (\\<^bsub>G\<^esub>, \\<^bsub>H\<^esub>)" by (simp add: DirProd_def) lemma mult_DirProd [simp]: "(g, h) \\<^bsub>(G \\ H)\<^esub> (g', h') = (g \\<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h')" by (simp add: DirProd_def) lemma mult_DirProd': "x \\<^bsub>(G \\ H)\<^esub> y = (fst x \\<^bsub>G\<^esub> fst y, snd x \\<^bsub>H\<^esub> snd y)" by (subst mult_DirProd [symmetric]) simp lemma DirProd_assoc: "(G \\ H \\ I) = (G \\ (H \\ I))" by auto lemma inv_DirProd [simp]: assumes "group G" and "group H" assumes g: "g \ carrier G" and h: "h \ carrier H" shows "m_inv (G \\ H) (g, h) = (inv\<^bsub>G\<^esub> g, inv\<^bsub>H\<^esub> h)" proof - interpret G: group G by fact interpret H: group H by fact interpret Prod: group "G \\ H" by (auto intro: DirProd_group group.intro group.axioms assms) show ?thesis by (simp add: Prod.inv_equality g h) qed lemma DirProd_subgroups : assumes "group G" and "subgroup H G" and "group K" and "subgroup I K" shows "subgroup (H \ I) (G \\ K)" proof (intro group.group_incl_imp_subgroup[OF DirProd_group[OF assms(1)assms(3)]]) have "H \ carrier G" "I \ carrier K" using subgroup.subset assms by blast+ thus "(H \ I) \ carrier (G \\ K)" unfolding DirProd_def by auto have "Group.group ((G\carrier := H\) \\ (K\carrier := I\))" using DirProd_group[OF subgroup.subgroup_is_group[OF assms(2)assms(1)] subgroup.subgroup_is_group[OF assms(4)assms(3)]]. moreover have "((G\carrier := H\) \\ (K\carrier := I\)) = ((G \\ K)\carrier := H \ I\)" unfolding DirProd_def using assms by simp ultimately show "Group.group ((G \\ K)\carrier := H \ I\)" by simp qed subsection \Homomorphisms (mono and epi) and Isomorphisms\ definition hom :: "_ => _ => ('a => 'b) set" where "hom G H = {h. h \ carrier G \ carrier H \ (\x \ carrier G. \y \ carrier G. h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y)}" lemma homI: "\\x. x \ carrier G \ h x \ carrier H; \x y. \x \ carrier G; y \ carrier G\ \ h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y\ \ h \ hom G H" by (auto simp: hom_def) lemma hom_carrier: "h \ hom G H \ h ` carrier G \ carrier H" by (auto simp: hom_def) lemma hom_in_carrier: "\h \ hom G H; x \ carrier G\ \ h x \ carrier H" by (auto simp: hom_def) lemma hom_compose: "\ f \ hom G H; g \ hom H I \ \ g \ f \ hom G I" unfolding hom_def by (auto simp add: Pi_iff) lemma (in group) hom_restrict: assumes "h \ hom G H" and "\g. g \ carrier G \ h g = t g" shows "t \ hom G H" using assms unfolding hom_def by (auto simp add: Pi_iff) lemma (in group) hom_compose: "[|h \ hom G H; i \ hom H I|] ==> compose (carrier G) i h \ hom G I" by (fastforce simp add: hom_def compose_def) lemma (in group) restrict_hom_iff [simp]: "(\x. if x \ carrier G then f x else g x) \ hom G H \ f \ hom G H" by (simp add: hom_def Pi_iff) definition iso :: "_ => _ => ('a => 'b) set" where "iso G H = {h. h \ hom G H \ bij_betw h (carrier G) (carrier H)}" definition is_iso :: "_ \ _ \ bool" (infixr "\" 60) where "G \ H = (iso G H \ {})" definition mon where "mon G H = {f \ hom G H. inj_on f (carrier G)}" definition epi where "epi G H = {f \ hom G H. f ` (carrier G) = carrier H}" lemma isoI: "\h \ hom G H; bij_betw h (carrier G) (carrier H)\ \ h \ iso G H" by (auto simp: iso_def) lemma is_isoI: "h \ iso G H \ G \ H" using is_iso_def by auto lemma epi_iff_subset: "f \ epi G G' \ f \ hom G G' \ carrier G' \ f ` carrier G" by (auto simp: epi_def hom_def) lemma iso_iff_mon_epi: "f \ iso G H \ f \ mon G H \ f \ epi G H" by (auto simp: iso_def mon_def epi_def bij_betw_def) lemma iso_set_refl: "(\x. x) \ iso G G" by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) lemma id_iso: "id \ iso G G" by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) corollary iso_refl [simp]: "G \ G" using iso_set_refl unfolding is_iso_def by auto lemma iso_iff: "h \ iso G H \ h \ hom G H \ h ` (carrier G) = carrier H \ inj_on h (carrier G)" by (auto simp: iso_def hom_def bij_betw_def) lemma iso_imp_homomorphism: "h \ iso G H \ h \ hom G H" by (simp add: iso_iff) lemma trivial_hom: "group H \ (\x. one H) \ hom G H" by (auto simp: hom_def Group.group_def) lemma (in group) hom_eq: assumes "f \ hom G H" "\x. x \ carrier G \ f' x = f x" shows "f' \ hom G H" using assms by (auto simp: hom_def) lemma (in group) iso_eq: assumes "f \ iso G H" "\x. x \ carrier G \ f' x = f x" shows "f' \ iso G H" using assms by (fastforce simp: iso_def inj_on_def bij_betw_def hom_eq image_iff) lemma (in group) iso_set_sym: assumes "h \ iso G H" shows "inv_into (carrier G) h \ iso H G" proof - have h: "h \ hom G H" "bij_betw h (carrier G) (carrier H)" using assms by (auto simp add: iso_def bij_betw_inv_into) then have HG: "bij_betw (inv_into (carrier G) h) (carrier H) (carrier G)" by (simp add: bij_betw_inv_into) have "inv_into (carrier G) h \ hom H G" unfolding hom_def proof safe show *: "\x. x \ carrier H \ inv_into (carrier G) h x \ carrier G" by (meson HG bij_betwE) show "inv_into (carrier G) h (x \\<^bsub>H\<^esub> y) = inv_into (carrier G) h x \ inv_into (carrier G) h y" if "x \ carrier H" "y \ carrier H" for x y proof (rule inv_into_f_eq) show "inj_on h (carrier G)" using bij_betw_def h(2) by blast show "inv_into (carrier G) h x \ inv_into (carrier G) h y \ carrier G" by (simp add: * that) show "h (inv_into (carrier G) h x \ inv_into (carrier G) h y) = x \\<^bsub>H\<^esub> y" using h bij_betw_inv_into_right [of h] unfolding hom_def by (simp add: "*" that) qed qed then show ?thesis by (simp add: Group.iso_def bij_betw_inv_into h) qed corollary (in group) iso_sym: "G \ H \ H \ G" using iso_set_sym unfolding is_iso_def by auto lemma iso_set_trans: "\h \ Group.iso G H; i \ Group.iso H I\ \ i \ h \ Group.iso G I" by (force simp: iso_def hom_compose intro: bij_betw_trans) corollary iso_trans [trans]: "\G \ H ; H \ I\ \ G \ I" using iso_set_trans unfolding is_iso_def by blast lemma iso_same_card: "G \ H \ card (carrier G) = card (carrier H)" using bij_betw_same_card unfolding is_iso_def iso_def by auto lemma iso_finite: "G \ H \ finite(carrier G) \ finite(carrier H)" by (auto simp: is_iso_def iso_def bij_betw_finite) lemma mon_compose: "\f \ mon G H; g \ mon H K\ \ (g \ f) \ mon G K" by (auto simp: mon_def intro: hom_compose comp_inj_on inj_on_subset [OF _ hom_carrier]) lemma mon_compose_rev: "\f \ hom G H; g \ hom H K; (g \ f) \ mon G K\ \ f \ mon G H" using inj_on_imageI2 by (auto simp: mon_def) lemma epi_compose: "\f \ epi G H; g \ epi H K\ \ (g \ f) \ epi G K" using hom_compose by (force simp: epi_def hom_compose simp flip: image_image) lemma epi_compose_rev: "\f \ hom G H; g \ hom H K; (g \ f) \ epi G K\ \ g \ epi H K" by (fastforce simp: epi_def hom_def Pi_iff image_def set_eq_iff) lemma iso_compose_rev: "\f \ hom G H; g \ hom H K; (g \ f) \ iso G K\ \ f \ mon G H \ g \ epi H K" unfolding iso_iff_mon_epi using mon_compose_rev epi_compose_rev by blast lemma epi_iso_compose_rev: assumes "f \ epi G H" "g \ hom H K" "(g \ f) \ iso G K" shows "f \ iso G H \ g \ iso H K" proof show "f \ iso G H" by (metis (no_types, lifting) assms epi_def iso_compose_rev iso_iff_mon_epi mem_Collect_eq) then have "f \ hom G H \ bij_betw f (carrier G) (carrier H)" using Group.iso_def \f \ Group.iso G H\ by blast then have "bij_betw g (carrier H) (carrier K)" using Group.iso_def assms(3) bij_betw_comp_iff by blast then show "g \ iso H K" using Group.iso_def assms(2) by blast qed lemma mon_left_invertible: "\f \ hom G H; \x. x \ carrier G \ g(f x) = x\ \ f \ mon G H" by (simp add: mon_def inj_on_def) metis lemma epi_right_invertible: "\g \ hom H G; f \ carrier G \ carrier H; \x. x \ carrier G \ g(f x) = x\ \ g \ epi H G" by (force simp: Pi_iff epi_iff_subset image_subset_iff_funcset subset_iff) lemma (in monoid) hom_imp_img_monoid: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "h \ hom G H" shows "monoid (H \ carrier := h ` (carrier G), one := h \\<^bsub>G\<^esub> \)" (is "monoid ?h_img") proof (rule monoidI) show "\\<^bsub>?h_img\<^esub> \ carrier ?h_img" by auto next fix x y z assume "x \ carrier ?h_img" "y \ carrier ?h_img" "z \ carrier ?h_img" then obtain g1 g2 g3 where g1: "g1 \ carrier G" "x = h g1" and g2: "g2 \ carrier G" "y = h g2" and g3: "g3 \ carrier G" "z = h g3" using image_iff[where ?f = h and ?A = "carrier G"] by auto have aux_lemma: "\a b. \ a \ carrier G; b \ carrier G \ \ h a \\<^bsub>(?h_img)\<^esub> h b = h (a \ b)" using assms unfolding hom_def by auto show "x \\<^bsub>(?h_img)\<^esub> \\<^bsub>(?h_img)\<^esub> = x" using aux_lemma[OF g1(1) one_closed] g1(2) r_one[OF g1(1)] by simp show "\\<^bsub>(?h_img)\<^esub> \\<^bsub>(?h_img)\<^esub> x = x" using aux_lemma[OF one_closed g1(1)] g1(2) l_one[OF g1(1)] by simp have "x \\<^bsub>(?h_img)\<^esub> y = h (g1 \ g2)" using aux_lemma g1 g2 by auto thus "x \\<^bsub>(?h_img)\<^esub> y \ carrier ?h_img" using g1(1) g2(1) by simp have "(x \\<^bsub>(?h_img)\<^esub> y) \\<^bsub>(?h_img)\<^esub> z = h ((g1 \ g2) \ g3)" using aux_lemma g1 g2 g3 by auto also have " ... = h (g1 \ (g2 \ g3))" using m_assoc[OF g1(1) g2(1) g3(1)] by simp also have " ... = x \\<^bsub>(?h_img)\<^esub> (y \\<^bsub>(?h_img)\<^esub> z)" using aux_lemma g1 g2 g3 by auto finally show "(x \\<^bsub>(?h_img)\<^esub> y) \\<^bsub>(?h_img)\<^esub> z = x \\<^bsub>(?h_img)\<^esub> (y \\<^bsub>(?h_img)\<^esub> z)" . qed lemma (in group) hom_imp_img_group: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "h \ hom G H" shows "group (H \ carrier := h ` (carrier G), one := h \\<^bsub>G\<^esub> \)" (is "group ?h_img") proof - interpret monoid ?h_img using hom_imp_img_monoid[OF assms] . show ?thesis proof (unfold_locales) show "carrier ?h_img \ Units ?h_img" proof (auto simp add: Units_def) have aux_lemma: "\g1 g2. \ g1 \ carrier G; g2 \ carrier G \ \ h g1 \\<^bsub>H\<^esub> h g2 = h (g1 \ g2)" using assms unfolding hom_def by auto fix g1 assume g1: "g1 \ carrier G" thus "\g2 \ carrier G. (h g2) \\<^bsub>H\<^esub> (h g1) = h \ \ (h g1) \\<^bsub>H\<^esub> (h g2) = h \" using aux_lemma[OF g1 inv_closed[OF g1]] aux_lemma[OF inv_closed[OF g1] g1] inv_closed by auto qed qed qed lemma (in group) iso_imp_group: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "G \ H" and "monoid H" shows "group H" proof - obtain \ where phi: "\ \ iso G H" "inv_into (carrier G) \ \ iso H G" using iso_set_sym assms unfolding is_iso_def by blast define \ where psi_def: "\ = inv_into (carrier G) \" have surj: "\ ` (carrier G) = (carrier H)" "\ ` (carrier H) = (carrier G)" and inj: "inj_on \ (carrier G)" "inj_on \ (carrier H)" and phi_hom: "\g1 g2. \ g1 \ carrier G; g2 \ carrier G \ \ \ (g1 \ g2) = (\ g1) \\<^bsub>H\<^esub> (\ g2)" and psi_hom: "\h1 h2. \ h1 \ carrier H; h2 \ carrier H \ \ \ (h1 \\<^bsub>H\<^esub> h2) = (\ h1) \ (\ h2)" using phi psi_def unfolding iso_def bij_betw_def hom_def by auto have phi_one: "\ \ = \\<^bsub>H\<^esub>" proof - have "(\ \) \\<^bsub>H\<^esub> \\<^bsub>H\<^esub> = (\ \) \\<^bsub>H\<^esub> (\ \)" by (metis assms(2) image_eqI monoid.r_one one_closed phi_hom r_one surj(1)) thus ?thesis by (metis (no_types, opaque_lifting) Units_eq Units_one_closed assms(2) f_inv_into_f imageI monoid.l_one monoid.one_closed phi_hom psi_def r_one surj) qed have "carrier H \ Units H" proof fix h assume h: "h \ carrier H" let ?inv_h = "\ (inv (\ h))" have "h \\<^bsub>H\<^esub> ?inv_h = \ (\ h) \\<^bsub>H\<^esub> ?inv_h" by (simp add: f_inv_into_f h psi_def surj(1)) also have " ... = \ ((\ h) \ inv (\ h))" by (metis h imageI inv_closed phi_hom surj(2)) also have " ... = \ \" by (simp add: h inv_into_into psi_def surj(1)) finally have 1: "h \\<^bsub>H\<^esub> ?inv_h = \\<^bsub>H\<^esub>" using phi_one by simp have "?inv_h \\<^bsub>H\<^esub> h = ?inv_h \\<^bsub>H\<^esub> \ (\ h)" by (simp add: f_inv_into_f h psi_def surj(1)) also have " ... = \ (inv (\ h) \ (\ h))" by (metis h imageI inv_closed phi_hom surj(2)) also have " ... = \ \" by (simp add: h inv_into_into psi_def surj(1)) finally have 2: "?inv_h \\<^bsub>H\<^esub> h = \\<^bsub>H\<^esub>" using phi_one by simp thus "h \ Units H" unfolding Units_def using 1 2 h surj by fastforce qed thus ?thesis unfolding group_def group_axioms_def using assms(2) by simp qed corollary (in group) iso_imp_img_group: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "h \ iso G H" shows "group (H \ one := h \ \)" proof - let ?h_img = "H \ carrier := h ` (carrier G), one := h \ \" have "h \ iso G ?h_img" using assms unfolding iso_def hom_def bij_betw_def by auto hence "G \ ?h_img" unfolding is_iso_def by auto hence "group ?h_img" using iso_imp_group[of ?h_img] hom_imp_img_monoid[of h H] assms unfolding iso_def by simp moreover have "carrier H = carrier ?h_img" using assms unfolding iso_def bij_betw_def by simp hence "H \ one := h \ \ = ?h_img" by simp ultimately show ?thesis by simp qed subsubsection \HOL Light's concept of an isomorphism pair\ definition group_isomorphisms where "group_isomorphisms G H f g \ f \ hom G H \ g \ hom H G \ (\x \ carrier G. g(f x) = x) \ (\y \ carrier H. f(g y) = y)" lemma group_isomorphisms_sym: "group_isomorphisms G H f g \ group_isomorphisms H G g f" by (auto simp: group_isomorphisms_def) lemma group_isomorphisms_imp_iso: "group_isomorphisms G H f g \ f \ iso G H" by (auto simp: iso_def inj_on_def image_def group_isomorphisms_def hom_def bij_betw_def Pi_iff, metis+) lemma (in group) iso_iff_group_isomorphisms: "f \ iso G H \ (\g. group_isomorphisms G H f g)" proof safe show "\g. group_isomorphisms G H f g" if "f \ Group.iso G H" unfolding group_isomorphisms_def proof (intro exI conjI) let ?g = "inv_into (carrier G) f" show "\x\carrier G. ?g (f x) = x" by (metis (no_types, lifting) Group.iso_def bij_betw_inv_into_left mem_Collect_eq that) show "\y\carrier H. f (?g y) = y" by (metis (no_types, lifting) Group.iso_def bij_betw_inv_into_right mem_Collect_eq that) qed (use Group.iso_def iso_set_sym that in \blast+\) next fix g assume "group_isomorphisms G H f g" then show "f \ Group.iso G H" by (auto simp: iso_def group_isomorphisms_def hom_in_carrier intro: bij_betw_byWitness) qed subsubsection \Involving direct products\ lemma DirProd_commute_iso_set: shows "(\(x,y). (y,x)) \ iso (G \\ H) (H \\ G)" by (auto simp add: iso_def hom_def inj_on_def bij_betw_def) corollary DirProd_commute_iso : "(G \\ H) \ (H \\ G)" using DirProd_commute_iso_set unfolding is_iso_def by blast lemma DirProd_assoc_iso_set: shows "(\(x,y,z). (x,(y,z))) \ iso (G \\ H \\ I) (G \\ (H \\ I))" by (auto simp add: iso_def hom_def inj_on_def bij_betw_def) lemma (in group) DirProd_iso_set_trans: assumes "g \ iso G G2" and "h \ iso H I" shows "(\(x,y). (g x, h y)) \ iso (G \\ H) (G2 \\ I)" proof- have "(\(x,y). (g x, h y)) \ hom (G \\ H) (G2 \\ I)" using assms unfolding iso_def hom_def by auto moreover have " inj_on (\(x,y). (g x, h y)) (carrier (G \\ H))" using assms unfolding iso_def DirProd_def bij_betw_def inj_on_def by auto moreover have "(\(x, y). (g x, h y)) ` carrier (G \\ H) = carrier (G2 \\ I)" using assms unfolding iso_def bij_betw_def image_def DirProd_def by fastforce ultimately show "(\(x,y). (g x, h y)) \ iso (G \\ H) (G2 \\ I)" unfolding iso_def bij_betw_def by auto qed corollary (in group) DirProd_iso_trans : assumes "G \ G2" and "H \ I" shows "G \\ H \ G2 \\ I" using DirProd_iso_set_trans assms unfolding is_iso_def by blast lemma hom_pairwise: "f \ hom G (DirProd H K) \ (fst \ f) \ hom G H \ (snd \ f) \ hom G K" apply (auto simp: hom_def mult_DirProd' dest: Pi_mem) apply (metis Product_Type.mem_Times_iff comp_eq_dest_lhs funcset_mem) by (metis mult_DirProd prod.collapse) lemma hom_paired: "(\x. (f x,g x)) \ hom G (DirProd H K) \ f \ hom G H \ g \ hom G K" by (simp add: hom_pairwise o_def) lemma hom_paired2: assumes "group G" "group H" shows "(\(x,y). (f x,g y)) \ hom (DirProd G H) (DirProd G' H') \ f \ hom G G' \ g \ hom H H'" using assms by (fastforce simp: hom_def Pi_def dest!: group.is_monoid) lemma iso_paired2: assumes "group G" "group H" shows "(\(x,y). (f x,g y)) \ iso (DirProd G H) (DirProd G' H') \ f \ iso G G' \ g \ iso H H'" using assms by (fastforce simp add: iso_def inj_on_def bij_betw_def hom_paired2 image_paired_Times times_eq_iff group_def monoid.carrier_not_empty) lemma hom_of_fst: assumes "group H" shows "(f \ fst) \ hom (DirProd G H) K \ f \ hom G K" proof - interpret group H by (rule assms) show ?thesis using one_closed by (auto simp: hom_def Pi_def) qed lemma hom_of_snd: assumes "group G" shows "(f \ snd) \ hom (DirProd G H) K \ f \ hom H K" proof - interpret group G by (rule assms) show ?thesis using one_closed by (auto simp: hom_def Pi_def) qed subsection\The locale for a homomorphism between two groups\ text\Basis for homomorphism proofs: we assume two groups \<^term>\G\ and \<^term>\H\, with a homomorphism \<^term>\h\ between them\ locale group_hom = G?: group G + H?: group H for G (structure) and H (structure) + fixes h assumes homh [simp]: "h \ hom G H" declare group_hom.homh [simp] lemma (in group_hom) hom_mult [simp]: "[| x \ carrier G; y \ carrier G |] ==> h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y" proof - assume "x \ carrier G" "y \ carrier G" with homh [unfolded hom_def] show ?thesis by simp qed lemma (in group_hom) hom_closed [simp]: "x \ carrier G ==> h x \ carrier H" proof - assume "x \ carrier G" with homh [unfolded hom_def] show ?thesis by auto qed lemma (in group_hom) one_closed: "h \ \ carrier H" by simp lemma (in group_hom) hom_one [simp]: "h \ = \\<^bsub>H\<^esub>" proof - have "h \ \\<^bsub>H\<^esub> \\<^bsub>H\<^esub> = h \ \\<^bsub>H\<^esub> h \" by (simp add: hom_mult [symmetric] del: hom_mult) then show ?thesis by (metis H.Units_eq H.Units_l_cancel H.one_closed local.one_closed) qed lemma hom_one: assumes "h \ hom G H" "group G" "group H" shows "h (one G) = one H" apply (rule group_hom.hom_one) by (simp add: assms group_hom_axioms_def group_hom_def) lemma hom_mult: "\h \ hom G H; x \ carrier G; y \ carrier G\ \ h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y" by (auto simp: hom_def) lemma (in group_hom) inv_closed [simp]: "x \ carrier G ==> h (inv x) \ carrier H" by simp lemma (in group_hom) hom_inv [simp]: assumes "x \ carrier G" shows "h (inv x) = inv\<^bsub>H\<^esub> (h x)" proof - have "h x \\<^bsub>H\<^esub> h (inv x) = h x \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h x)" using assms by (simp flip: hom_mult) with assms show ?thesis by (simp del: H.r_inv H.Units_r_inv) qed lemma (in group) int_pow_is_hom: \<^marker>\contributor \Joachim Breitner\\ "x \ carrier G \ (([^]) x) \ hom \ carrier = UNIV, mult = (+), one = 0::int \ G " unfolding hom_def by (simp add: int_pow_mult) lemma (in group_hom) img_is_subgroup: "subgroup (h ` (carrier G)) H" \<^marker>\contributor \Paulo Emílio de Vilhena\\ apply (rule subgroupI) apply (auto simp add: image_subsetI) apply (metis G.inv_closed hom_inv image_iff) by (metis G.monoid_axioms hom_mult image_eqI monoid.m_closed) lemma (in group_hom) subgroup_img_is_subgroup: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "subgroup I G" shows "subgroup (h ` I) H" proof - have "h \ hom (G \ carrier := I \) H" using G.subgroupE[OF assms] subgroup.mem_carrier[OF assms] homh unfolding hom_def by auto hence "group_hom (G \ carrier := I \) H h" using subgroup.subgroup_is_group[OF assms G.is_group] is_group unfolding group_hom_def group_hom_axioms_def by simp thus ?thesis using group_hom.img_is_subgroup[of "G \ carrier := I \" H h] by simp qed lemma (in group_hom) induced_group_hom: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "subgroup I G" shows "group_hom (G \ carrier := I \) (H \ carrier := h ` I \) h" proof - have "h \ hom (G \ carrier := I \) (H \ carrier := h ` I \)" using homh subgroup.mem_carrier[OF assms] unfolding hom_def by auto thus ?thesis unfolding group_hom_def group_hom_axioms_def using subgroup.subgroup_is_group[OF assms G.is_group] subgroup.subgroup_is_group[OF subgroup_img_is_subgroup[OF assms] is_group] by simp qed lemma (in group) canonical_inj_is_hom: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "subgroup H G" shows "group_hom (G \ carrier := H \) G id" unfolding group_hom_def group_hom_axioms_def hom_def using subgroup.subgroup_is_group[OF assms is_group] is_group subgroup.subset[OF assms] by auto lemma (in group_hom) hom_nat_pow: \<^marker>\contributor \Paulo Emílio de Vilhena\\ "x \ carrier G \ h (x [^] (n :: nat)) = (h x) [^]\<^bsub>H\<^esub> n" by (induction n) auto lemma (in group_hom) hom_int_pow: \<^marker>\contributor \Paulo Emílio de Vilhena\\ "x \ carrier G \ h (x [^] (n :: int)) = (h x) [^]\<^bsub>H\<^esub> n" using hom_nat_pow by (simp add: int_pow_def2) lemma hom_nat_pow: "\h \ hom G H; x \ carrier G; group G; group H\ \ h (x [^]\<^bsub>G\<^esub> (n :: nat)) = (h x) [^]\<^bsub>H\<^esub> n" by (simp add: group_hom.hom_nat_pow group_hom_axioms_def group_hom_def) lemma hom_int_pow: "\h \ hom G H; x \ carrier G; group G; group H\ \ h (x [^]\<^bsub>G\<^esub> (n :: int)) = (h x) [^]\<^bsub>H\<^esub> n" by (simp add: group_hom.hom_int_pow group_hom_axioms.intro group_hom_def) subsection \Commutative Structures\ text \ Naming convention: multiplicative structures that are commutative are called \emph{commutative}, additive structures are called \emph{Abelian}. \ locale comm_monoid = monoid + assumes m_comm: "\x \ carrier G; y \ carrier G\ \ x \ y = y \ x" lemma (in comm_monoid) m_lcomm: "\x \ carrier G; y \ carrier G; z \ carrier G\ \ x \ (y \ z) = y \ (x \ z)" proof - assume xyz: "x \ carrier G" "y \ carrier G" "z \ carrier G" from xyz have "x \ (y \ z) = (x \ y) \ z" by (simp add: m_assoc) also from xyz have "... = (y \ x) \ z" by (simp add: m_comm) also from xyz have "... = y \ (x \ z)" by (simp add: m_assoc) finally show ?thesis . qed lemmas (in comm_monoid) m_ac = m_assoc m_comm m_lcomm lemma comm_monoidI: fixes G (structure) assumes m_closed: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \ y \ carrier G" and one_closed: "\ \ carrier G" and m_assoc: "!!x y z. [| x \ carrier G; y \ carrier G; z \ carrier G |] ==> (x \ y) \ z = x \ (y \ z)" and l_one: "!!x. x \ carrier G ==> \ \ x = x" and m_comm: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \ y = y \ x" shows "comm_monoid G" using l_one by (auto intro!: comm_monoid.intro comm_monoid_axioms.intro monoid.intro intro: assms simp: m_closed one_closed m_comm) lemma (in monoid) monoid_comm_monoidI: assumes m_comm: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \ y = y \ x" shows "comm_monoid G" by (rule comm_monoidI) (auto intro: m_assoc m_comm) lemma (in comm_monoid) submonoid_is_comm_monoid : assumes "submonoid H G" shows "comm_monoid (G\carrier := H\)" proof (intro monoid.monoid_comm_monoidI) show "monoid (G\carrier := H\)" using submonoid.submonoid_is_monoid assms comm_monoid_axioms comm_monoid_def by blast show "\x y. x \ carrier (G\carrier := H\) \ y \ carrier (G\carrier := H\) \ x \\<^bsub>G\carrier := H\\<^esub> y = y \\<^bsub>G\carrier := H\\<^esub> x" by simp (meson assms m_comm submonoid.mem_carrier) qed locale comm_group = comm_monoid + group lemma (in group) group_comm_groupI: assumes m_comm: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \ y = y \ x" shows "comm_group G" by standard (simp_all add: m_comm) lemma comm_groupI: fixes G (structure) assumes m_closed: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \ y \ carrier G" and one_closed: "\ \ carrier G" and m_assoc: "!!x y z. [| x \ carrier G; y \ carrier G; z \ carrier G |] ==> (x \ y) \ z = x \ (y \ z)" and m_comm: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \ y = y \ x" and l_one: "!!x. x \ carrier G ==> \ \ x = x" and l_inv_ex: "!!x. x \ carrier G ==> \y \ carrier G. y \ x = \" shows "comm_group G" by (fast intro: group.group_comm_groupI groupI assms) lemma comm_groupE: fixes G (structure) assumes "comm_group G" shows "\x y. \ x \ carrier G; y \ carrier G \ \ x \ y \ carrier G" and "\ \ carrier G" and "\x y z. \ x \ carrier G; y \ carrier G; z \ carrier G \ \ (x \ y) \ z = x \ (y \ z)" and "\x y. \ x \ carrier G; y \ carrier G \ \ x \ y = y \ x" and "\x. x \ carrier G \ \ \ x = x" and "\x. x \ carrier G \ \y \ carrier G. y \ x = \" apply (simp_all add: group.axioms assms comm_group.axioms comm_monoid.m_comm comm_monoid.m_ac(1)) by (simp_all add: Group.group.axioms(1) assms comm_group.axioms(2) monoid.m_closed group.r_inv_ex) lemma (in comm_group) inv_mult: "[| x \ carrier G; y \ carrier G |] ==> inv (x \ y) = inv x \ inv y" by (simp add: m_ac inv_mult_group) lemma (in comm_monoid) nat_pow_distrib: fixes n::nat assumes "x \ carrier G" "y \ carrier G" shows "(x \ y) [^] n = x [^] n \ y [^] n" by (simp add: assms pow_mult_distrib m_comm) lemma (in comm_group) int_pow_distrib: assumes "x \ carrier G" "y \ carrier G" shows "(x \ y) [^] (i::int) = x [^] i \ y [^] i" by (simp add: assms int_pow_mult_distrib m_comm) lemma (in comm_monoid) hom_imp_img_comm_monoid: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "h \ hom G H" shows "comm_monoid (H \ carrier := h ` (carrier G), one := h \\<^bsub>G\<^esub> \)" (is "comm_monoid ?h_img") proof (rule monoid.monoid_comm_monoidI) show "monoid ?h_img" using hom_imp_img_monoid[OF assms] . next fix x y assume "x \ carrier ?h_img" "y \ carrier ?h_img" then obtain g1 g2 where g1: "g1 \ carrier G" "x = h g1" and g2: "g2 \ carrier G" "y = h g2" by auto have "x \\<^bsub>(?h_img)\<^esub> y = h (g1 \ g2)" using g1 g2 assms unfolding hom_def by auto also have " ... = h (g2 \ g1)" using m_comm[OF g1(1) g2(1)] by simp also have " ... = y \\<^bsub>(?h_img)\<^esub> x" using g1 g2 assms unfolding hom_def by auto finally show "x \\<^bsub>(?h_img)\<^esub> y = y \\<^bsub>(?h_img)\<^esub> x" . qed lemma (in comm_group) hom_group_mult: assumes "f \ hom H G" "g \ hom H G" shows "(\x. f x \\<^bsub>G\<^esub> g x) \ hom H G" using assms by (auto simp: hom_def Pi_def m_ac) lemma (in comm_group) hom_imp_img_comm_group: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "h \ hom G H" shows "comm_group (H \ carrier := h ` (carrier G), one := h \\<^bsub>G\<^esub> \)" unfolding comm_group_def using hom_imp_img_group[OF assms] hom_imp_img_comm_monoid[OF assms] by simp lemma (in comm_group) iso_imp_img_comm_group: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "h \ iso G H" shows "comm_group (H \ one := h \\<^bsub>G\<^esub> \)" proof - let ?h_img = "H \ carrier := h ` (carrier G), one := h \ \" have "comm_group ?h_img" using hom_imp_img_comm_group[of h H] assms unfolding iso_def by auto moreover have "carrier H = carrier ?h_img" using assms unfolding iso_def bij_betw_def by simp hence "H \ one := h \ \ = ?h_img" by simp ultimately show ?thesis by simp qed lemma (in comm_group) iso_imp_comm_group: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "G \ H" "monoid H" shows "comm_group H" proof - obtain h where h: "h \ iso G H" using assms(1) unfolding is_iso_def by auto hence comm_gr: "comm_group (H \ one := h \ \)" using iso_imp_img_comm_group[of h H] by simp hence "\x. x \ carrier H \ h \ \\<^bsub>H\<^esub> x = x" using monoid.l_one[of "H \ one := h \ \"] unfolding comm_group_def comm_monoid_def by simp moreover have "h \ \ carrier H" using h one_closed unfolding iso_def hom_def by auto ultimately have "h \ = \\<^bsub>H\<^esub>" using monoid.one_unique[OF assms(2), of "h \"] by simp hence "H = H \ one := h \ \" by simp thus ?thesis using comm_gr by simp qed (*A subgroup of a subgroup is a subgroup of the group*) lemma (in group) incl_subgroup: assumes "subgroup J G" and "subgroup I (G\carrier:=J\)" shows "subgroup I G" unfolding subgroup_def proof have H1: "I \ carrier (G\carrier:=J\)" using assms(2) subgroup.subset by blast also have H2: "...\J" by simp also have "...\(carrier G)" by (simp add: assms(1) subgroup.subset) finally have H: "I \ carrier G" by simp have "(\x y. \x \ I ; y \ I\ \ x \ y \ I)" using assms(2) by (auto simp add: subgroup_def) thus "I \ carrier G \ (\x y. x \ I \ y \ I \ x \ y \ I)" using H by blast have K: "\ \ I" using assms(2) by (auto simp add: subgroup_def) have "(\x. x \ I \ inv x \ I)" using assms subgroup.m_inv_closed H by (metis H1 H2 m_inv_consistent subsetCE) thus "\ \ I \ (\x. x \ I \ inv x \ I)" using K by blast qed (*A subgroup included in another subgroup is a subgroup of the subgroup*) lemma (in group) subgroup_incl: assumes "subgroup I G" and "subgroup J G" and "I \ J" shows "subgroup I (G \ carrier := J \)" using group.group_incl_imp_subgroup[of "G \ carrier := J \" I] assms(1-2)[THEN subgroup.subgroup_is_group[OF _ group_axioms]] assms(3) by auto subsection \The Lattice of Subgroups of a Group\ text_raw \\label{sec:subgroup-lattice}\ theorem (in group) subgroups_partial_order: "partial_order \carrier = {H. subgroup H G}, eq = (=), le = (\)\" by standard simp_all lemma (in group) subgroup_self: "subgroup (carrier G) G" by (rule subgroupI) auto lemma (in group) subgroup_imp_group: "subgroup H G ==> group (G\carrier := H\)" by (erule subgroup.subgroup_is_group) (rule group_axioms) lemma (in group) subgroup_mult_equality: "\ subgroup H G; h1 \ H; h2 \ H \ \ h1 \\<^bsub>G \ carrier := H \\<^esub> h2 = h1 \ h2" unfolding subgroup_def by simp theorem (in group) subgroups_Inter: assumes subgr: "(\H. H \ A \ subgroup H G)" and not_empty: "A \ {}" shows "subgroup (\A) G" proof (rule subgroupI) from subgr [THEN subgroup.subset] and not_empty show "\A \ carrier G" by blast next from subgr [THEN subgroup.one_closed] show "\A \ {}" by blast next fix x assume "x \ \A" with subgr [THEN subgroup.m_inv_closed] show "inv x \ \A" by blast next fix x y assume "x \ \A" "y \ \A" with subgr [THEN subgroup.m_closed] show "x \ y \ \A" by blast qed lemma (in group) subgroups_Inter_pair : assumes "subgroup I G" and "subgroup J G" shows "subgroup (I\J) G" using subgroups_Inter[ where ?A = "{I,J}"] assms by auto theorem (in group) subgroups_complete_lattice: "complete_lattice \carrier = {H. subgroup H G}, eq = (=), le = (\)\" (is "complete_lattice ?L") proof (rule partial_order.complete_lattice_criterion1) show "partial_order ?L" by (rule subgroups_partial_order) next have "greatest ?L (carrier G) (carrier ?L)" by (unfold greatest_def) (simp add: subgroup.subset subgroup_self) then show "\G. greatest ?L G (carrier ?L)" .. next fix A assume L: "A \ carrier ?L" and non_empty: "A \ {}" then have Int_subgroup: "subgroup (\A) G" by (fastforce intro: subgroups_Inter) have "greatest ?L (\A) (Lower ?L A)" (is "greatest _ ?Int _") proof (rule greatest_LowerI) fix H assume H: "H \ A" with L have subgroupH: "subgroup H G" by auto from subgroupH have groupH: "group (G \carrier := H\)" (is "group ?H") by (rule subgroup_imp_group) from groupH have monoidH: "monoid ?H" by (rule group.is_monoid) from H have Int_subset: "?Int \ H" by fastforce then show "le ?L ?Int H" by simp next fix H assume H: "H \ Lower ?L A" with L Int_subgroup show "le ?L H ?Int" by (fastforce simp: Lower_def intro: Inter_greatest) next show "A \ carrier ?L" by (rule L) next show "?Int \ carrier ?L" by simp (rule Int_subgroup) qed then show "\I. greatest ?L I (Lower ?L A)" .. qed subsection\The units in any monoid give rise to a group\ text \Thanks to Jeremy Avigad. The file Residues.thy provides some infrastructure to use facts about the unit group within the ring locale. \ definition units_of :: "('a, 'b) monoid_scheme \ 'a monoid" where "units_of G = \carrier = Units G, Group.monoid.mult = Group.monoid.mult G, one = one G\" lemma (in monoid) units_group: "group (units_of G)" proof - have "\x y z. \x \ Units G; y \ Units G; z \ Units G\ \ x \ y \ z = x \ (y \ z)" by (simp add: Units_closed m_assoc) moreover have "\x. x \ Units G \ \y\Units G. y \ x = \" using Units_l_inv by blast ultimately show ?thesis unfolding units_of_def by (force intro!: groupI) qed lemma (in comm_monoid) units_comm_group: "comm_group (units_of G)" proof - have "\x y. \x \ carrier (units_of G); y \ carrier (units_of G)\ \ x \\<^bsub>units_of G\<^esub> y = y \\<^bsub>units_of G\<^esub> x" by (simp add: Units_closed m_comm units_of_def) then show ?thesis by (rule group.group_comm_groupI [OF units_group]) auto qed lemma units_of_carrier: "carrier (units_of G) = Units G" by (auto simp: units_of_def) lemma units_of_mult: "mult (units_of G) = mult G" by (auto simp: units_of_def) lemma units_of_one: "one (units_of G) = one G" by (auto simp: units_of_def) lemma (in monoid) units_of_inv: assumes "x \ Units G" shows "m_inv (units_of G) x = m_inv G x" by (simp add: assms group.inv_equality units_group units_of_carrier units_of_mult units_of_one) lemma units_of_units [simp] : "Units (units_of G) = Units G" unfolding units_of_def Units_def by force lemma (in group) surj_const_mult: "a \ carrier G \ (\x. a \ x) ` carrier G = carrier G" apply (auto simp add: image_def) by (metis inv_closed inv_solve_left m_closed) lemma (in group) l_cancel_one [simp]: "x \ carrier G \ a \ carrier G \ x \ a = x \ a = one G" by (metis Units_eq Units_l_cancel monoid.r_one monoid_axioms one_closed) lemma (in group) r_cancel_one [simp]: "x \ carrier G \ a \ carrier G \ a \ x = x \ a = one G" by (metis monoid.l_one monoid_axioms one_closed right_cancel) lemma (in group) l_cancel_one' [simp]: "x \ carrier G \ a \ carrier G \ x = x \ a \ a = one G" using l_cancel_one by fastforce lemma (in group) r_cancel_one' [simp]: "x \ carrier G \ a \ carrier G \ x = a \ x \ a = one G" using r_cancel_one by fastforce declare pow_nat [simp] (*causes looping if added above, especially with int_pow_def2*) end diff --git a/src/HOL/Algebra/Multiplicative_Group.thy b/src/HOL/Algebra/Multiplicative_Group.thy --- a/src/HOL/Algebra/Multiplicative_Group.thy +++ b/src/HOL/Algebra/Multiplicative_Group.thy @@ -1,1061 +1,1061 @@ (* Title: HOL/Algebra/Multiplicative_Group.thy Author: Simon Wimmer Author: Lars Noschinski *) theory Multiplicative_Group imports Complex_Main Group Coset UnivPoly Generated_Groups Elementary_Groups begin section \Simplification Rules for Polynomials\ text_raw \\label{sec:simp-rules}\ lemma (in ring_hom_cring) hom_sub[simp]: assumes "x \ carrier R" "y \ carrier R" shows "h (x \ y) = h x \\<^bsub>S\<^esub> h y" using assms by (simp add: R.minus_eq S.minus_eq) context UP_ring begin lemma deg_nzero_nzero: assumes deg_p_nzero: "deg R p \ 0" shows "p \ \\<^bsub>P\<^esub>" using deg_zero deg_p_nzero by auto lemma deg_add_eq: assumes c: "p \ carrier P" "q \ carrier P" assumes "deg R q \ deg R p" shows "deg R (p \\<^bsub>P\<^esub> q) = max (deg R p) (deg R q)" proof - let ?m = "max (deg R p) (deg R q)" from assms have "coeff P p ?m = \ \ coeff P q ?m \ \" by (metis deg_belowI lcoeff_nonzero[OF deg_nzero_nzero] linear max.absorb_iff2 max.absorb1) then have "coeff P (p \\<^bsub>P\<^esub> q) ?m \ \" using assms by auto then have "deg R (p \\<^bsub>P\<^esub> q) \ ?m" using assms by (blast intro: deg_belowI) with deg_add[OF c] show ?thesis by arith qed lemma deg_minus_eq: assumes "p \ carrier P" "q \ carrier P" "deg R q \ deg R p" shows "deg R (p \\<^bsub>P\<^esub> q) = max (deg R p) (deg R q)" using assms by (simp add: deg_add_eq a_minus_def) end context UP_cring begin lemma evalRR_add: assumes "p \ carrier P" "q \ carrier P" assumes x: "x \ carrier R" shows "eval R R id x (p \\<^bsub>P\<^esub> q) = eval R R id x p \ eval R R id x q" proof - interpret UP_pre_univ_prop R R id by unfold_locales simp interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) show ?thesis using assms by simp qed lemma evalRR_sub: assumes "p \ carrier P" "q \ carrier P" assumes x: "x \ carrier R" shows "eval R R id x (p \\<^bsub>P\<^esub> q) = eval R R id x p \ eval R R id x q" proof - interpret UP_pre_univ_prop R R id by unfold_locales simp interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) show ?thesis using assms by simp qed lemma evalRR_mult: assumes "p \ carrier P" "q \ carrier P" assumes x: "x \ carrier R" shows "eval R R id x (p \\<^bsub>P\<^esub> q) = eval R R id x p \ eval R R id x q" proof - interpret UP_pre_univ_prop R R id by unfold_locales simp interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) show ?thesis using assms by simp qed lemma evalRR_monom: assumes a: "a \ carrier R" and x: "x \ carrier R" shows "eval R R id x (monom P a d) = a \ x [^] d" proof - interpret UP_pre_univ_prop R R id by unfold_locales simp show ?thesis using assms by (simp add: eval_monom) qed lemma evalRR_one: assumes x: "x \ carrier R" shows "eval R R id x \\<^bsub>P\<^esub> = \" proof - interpret UP_pre_univ_prop R R id by unfold_locales simp interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) show ?thesis using assms by simp qed lemma carrier_evalRR: assumes x: "x \ carrier R" and "p \ carrier P" shows "eval R R id x p \ carrier R" proof - interpret UP_pre_univ_prop R R id by unfold_locales simp interpret ring_hom_cring P R "eval R R id x" by unfold_locales (rule eval_ring_hom[OF x]) show ?thesis using assms by simp qed lemmas evalRR_simps = evalRR_add evalRR_sub evalRR_mult evalRR_monom evalRR_one carrier_evalRR end section \Properties of the Euler \\\-function\ text_raw \\label{sec:euler-phi}\ text\ In this section we prove that for every positive natural number the equation $\sum_{d | n}^n \varphi(d) = n$ holds. \ lemma dvd_div_ge_1: fixes a b :: nat assumes "a \ 1" "b dvd a" shows "a div b \ 1" proof - from \b dvd a\ obtain c where "a = b * c" .. with \a \ 1\ show ?thesis by simp qed lemma dvd_nat_bounds: fixes n p :: nat assumes "p > 0" "n dvd p" shows "n > 0 \ n \ p" using assms by (simp add: dvd_pos_nat dvd_imp_le) (* TODO FIXME: This is the "totient" function from HOL-Number_Theory, but since part of HOL-Number_Theory depends on HOL-Algebra.Multiplicative_Group, there would be a cyclic dependency. *) definition phi' :: "nat => nat" where "phi' m = card {x. 1 \ x \ x \ m \ coprime x m}" notation (latex output) phi' ("\ _") lemma phi'_nonzero: assumes "m > 0" shows "phi' m > 0" proof - have "1 \ {x. 1 \ x \ x \ m \ coprime x m}" using assms by simp hence "card {x. 1 \ x \ x \ m \ coprime x m} > 0" by (auto simp: card_gt_0_iff) thus ?thesis unfolding phi'_def by simp qed lemma dvd_div_eq_1: fixes a b c :: nat assumes "c dvd a" "c dvd b" "a div c = b div c" shows "a = b" using assms dvd_mult_div_cancel[OF \c dvd a\] dvd_mult_div_cancel[OF \c dvd b\] by presburger lemma dvd_div_eq_2: fixes a b c :: nat assumes "c>0" "a dvd c" "b dvd c" "c div a = c div b" shows "a = b" proof - have "a > 0" "a \ c" using dvd_nat_bounds[OF assms(1-2)] by auto have "a*(c div a) = c" using assms dvd_mult_div_cancel by fastforce also have "\ = b*(c div a)" using assms dvd_mult_div_cancel by fastforce finally show "a = b" using \c>0\ dvd_div_ge_1[OF _ \a dvd c\] by fastforce qed lemma div_mult_mono: fixes a b c :: nat assumes "a > 0" "a\d" shows "a * b div d \ b" proof - have "a*b div d \ b*a div a" using assms div_le_mono2 mult.commute[of a b] by presburger thus ?thesis using assms by force qed text\ We arrive at the main result of this section: For every positive natural number the equation $\sum_{d | n}^n \varphi(d) = n$ holds. The outline of the proof for this lemma is as follows: We count the $n$ fractions $1/n$, $\ldots$, $(n-1)/n$, $n/n$. We analyze the reduced form $a/d = m/n$ for any of those fractions. We want to know how many fractions $m/n$ have the reduced form denominator $d$. The condition $1 \leq m \leq n$ is equivalent to the condition $1 \leq a \leq d$. Therefore we want to know how many $a$ with $1 \leq a \leq d$ exist, s.t. \<^term>\gcd a d = 1\. This number is exactly \<^term>\phi' d\. Finally, by counting the fractions $m/n$ according to their reduced form denominator, we get: @{term [display] "(\d | d dvd n . phi' d) = n"}. To formalize this proof in Isabelle, we analyze for an arbitrary divisor $d$ of $n$ \begin{itemize} \item the set of reduced form numerators \<^term>\{a. (1::nat) \ a \ a \ d \ coprime a d}\ \item the set of numerators $m$, for which $m/n$ has the reduced form denominator $d$, i.e. the set \<^term>\{m \ {1::nat .. n}. n div gcd m n = d}\ \end{itemize} We show that \<^term>\\a. a*n div d\ with the inverse \<^term>\\a. a div gcd a n\ is a bijection between theses sets, thus yielding the equality @{term [display] "phi' d = card {m \ {1 .. n}. n div gcd m n = d}"} This gives us @{term [display] "(\d | d dvd n . phi' d) = card (\d \ {d. d dvd n}. {m \ {1 .. n}. n div gcd m n = d})"} and by showing \<^term>\(\d \ {d. d dvd n}. {m \ {1::nat .. n}. n div gcd m n = d}) \ {1 .. n}\ (this is our counting argument) the thesis follows. \ lemma sum_phi'_factors: fixes n :: nat assumes "n > 0" shows "(\d | d dvd n. phi' d) = n" proof - { fix d assume "d dvd n" then obtain q where q: "n = d * q" .. have "card {a. 1 \ a \ a \ d \ coprime a d} = card {m \ {1 .. n}. n div gcd m n = d}" (is "card ?RF = card ?F") proof (rule card_bij_eq) { fix a b assume "a * n div d = b * n div d" hence "a * (n div d) = b * (n div d)" using dvd_div_mult[OF \d dvd n\] by (fastforce simp add: mult.commute) hence "a = b" using dvd_div_ge_1[OF _ \d dvd n\] \n>0\ by (simp add: mult.commute nat_mult_eq_cancel1) } thus "inj_on (\a. a*n div d) ?RF" unfolding inj_on_def by blast { fix a assume a: "a\?RF" hence "a * (n div d) \ 1" using \n>0\ dvd_div_ge_1[OF _ \d dvd n\] by simp hence ge_1: "a * n div d \ 1" by (simp add: \d dvd n\ div_mult_swap) have le_n: "a * n div d \ n" using div_mult_mono a by simp have "gcd (a * n div d) n = n div d * gcd a d" by (simp add: gcd_mult_distrib_nat q ac_simps) hence "n div gcd (a * n div d) n = d*n div (d*(n div d))" using a by simp hence "a * n div d \ ?F" using ge_1 le_n by (fastforce simp add: \d dvd n\) } thus "(\a. a*n div d) ` ?RF \ ?F" by blast { fix m l assume A: "m \ ?F" "l \ ?F" "m div gcd m n = l div gcd l n" hence "gcd m n = gcd l n" using dvd_div_eq_2[OF assms] by fastforce hence "m = l" using dvd_div_eq_1[of "gcd m n" m l] A(3) by fastforce } thus "inj_on (\a. a div gcd a n) ?F" unfolding inj_on_def by blast { fix m assume "m \ ?F" hence "m div gcd m n \ ?RF" using dvd_div_ge_1 by (fastforce simp add: div_le_mono div_gcd_coprime) } thus "(\a. a div gcd a n) ` ?F \ ?RF" by blast qed force+ } hence phi'_eq: "\d. d dvd n \ phi' d = card {m \ {1 .. n}. n div gcd m n = d}" unfolding phi'_def by presburger have fin: "finite {d. d dvd n}" using dvd_nat_bounds[OF \n>0\] by force have "(\d | d dvd n. phi' d) = card (\d \ {d. d dvd n}. {m \ {1 .. n}. n div gcd m n = d})" using card_UN_disjoint[OF fin, of "(\d. {m \ {1 .. n}. n div gcd m n = d})"] phi'_eq by fastforce also have "(\d \ {d. d dvd n}. {m \ {1 .. n}. n div gcd m n = d}) = {1 .. n}" (is "?L = ?R") proof show "?L \ ?R" proof fix m assume m: "m \ ?R" thus "m \ ?L" using dvd_triv_right[of "n div gcd m n" "gcd m n"] by simp qed qed fastforce finally show ?thesis by force qed section \Order of an Element of a Group\ text_raw \\label{sec:order-elem}\ context group begin definition (in group) ord :: "'a \ nat" where "ord x \ (@d. \n::nat. x [^] n = \ \ d dvd n)" lemma (in group) pow_eq_id: assumes "x \ carrier G" shows "x [^] n = \ \ (ord x) dvd n" proof (cases "\n::nat. pow G x n = one G \ n = 0") case True show ?thesis unfolding ord_def by (rule someI2 [where a=0]) (auto simp: True) next case False define N where "N \ LEAST n::nat. x [^] n = \ \ n > 0" have N: "x [^] N = \ \ N > 0" using False apply (simp add: N_def) by (metis (mono_tags, lifting) LeastI) have eq0: "n = 0" if "x [^] n = \" "n < N" for n using N_def not_less_Least that by fastforce show ?thesis unfolding ord_def proof (rule someI2 [where a = N], rule allI) fix n :: "nat" show "(x [^] n = \) \ (N dvd n)" proof (cases "n = 0") case False show ?thesis unfolding dvd_def proof safe assume 1: "x [^] n = \" have "x [^] n = x [^] (n mod N + N * (n div N))" by simp also have "\ = x [^] (n mod N) \ x [^] (N * (n div N))" by (simp add: assms nat_pow_mult) also have "\ = x [^] (n mod N)" by (metis N assms l_cancel_one nat_pow_closed nat_pow_one nat_pow_pow) finally have "x [^] (n mod N) = \" by (simp add: "1") then have "n mod N = 0" using N eq0 mod_less_divisor by blast then show "\k. n = N * k" by blast next fix k :: "nat" assume "n = N * k" with N show "x [^] (N * k) = \" by (metis assms nat_pow_one nat_pow_pow) qed qed simp qed blast qed lemma (in group) pow_ord_eq_1 [simp]: "x \ carrier G \ x [^] ord x = \" by (simp add: pow_eq_id) lemma (in group) int_pow_eq_id: assumes "x \ carrier G" shows "(pow G x i = one G \ int (ord x) dvd i)" proof (cases i rule: int_cases2) case (nonneg n) then show ?thesis by (simp add: int_pow_int pow_eq_id assms) next case (nonpos n) then have "x [^] i = inv (x [^] n)" by (simp add: assms int_pow_int int_pow_neg) then show ?thesis by (simp add: assms pow_eq_id nonpos) qed lemma (in group) int_pow_eq: "x \ carrier G \ (x [^] m = x [^] n) \ int (ord x) dvd (n - m)" apply (simp flip: int_pow_eq_id) by (metis int_pow_closed int_pow_diff inv_closed r_inv right_cancel) lemma (in group) ord_eq_0: "x \ carrier G \ (ord x = 0 \ (\n::nat. n \ 0 \ x [^] n \ \))" by (auto simp: pow_eq_id) lemma (in group) ord_unique: "x \ carrier G \ ord x = d \ (\n. pow G x n = one G \ d dvd n)" by (meson dvd_antisym dvd_refl pow_eq_id) lemma (in group) ord_eq_1: "x \ carrier G \ (ord x = 1 \ x = \)" by (metis pow_eq_id nat_dvd_1_iff_1 nat_pow_eone) lemma (in group) ord_id [simp]: "ord (one G) = 1" using ord_eq_1 by blast lemma (in group) ord_inv [simp]: "x \ carrier G \ ord (m_inv G x) = ord x" by (simp add: ord_unique pow_eq_id nat_pow_inv) lemma (in group) ord_pow: assumes "x \ carrier G" "k dvd ord x" "k \ 0" shows "ord (pow G x k) = ord x div k" proof - have "(x [^] k) [^] (ord x div k) = \" using assms by (simp add: nat_pow_pow) moreover have "ord x dvd k * ord (x [^] k)" by (metis assms(1) pow_ord_eq_1 pow_eq_id nat_pow_closed nat_pow_pow) ultimately show ?thesis by (metis assms div_dvd_div dvd_antisym dvd_triv_left pow_eq_id nat_pow_closed nonzero_mult_div_cancel_left) qed lemma (in group) ord_mul_divides: assumes eq: "x \ y = y \ x" and xy: "x \ carrier G" "y \ carrier G" shows "ord (x \ y) dvd (ord x * ord y)" apply (simp add: xy flip: pow_eq_id eq) by (metis dvd_triv_left dvd_triv_right eq pow_eq_id one_closed pow_mult_distrib r_one xy) lemma (in comm_group) abelian_ord_mul_divides: "\x \ carrier G; y \ carrier G\ \ ord (x \ y) dvd (ord x * ord y)" by (simp add: ord_mul_divides m_comm) lemma ord_inj: assumes a: "a \ carrier G" shows "inj_on (\ x . a [^] x) {0 .. ord a - 1}" proof - let ?M = "Max (ord ` carrier G)" have "finite {d \ {..?M}. a [^] d = \}" by auto have *: False if A: "x < y" "x \ {0 .. ord a - 1}" "y \ {0 .. ord a - 1}" "a [^] x = a [^] y" for x y proof - have "y - x < ord a" using that by auto moreover have "a [^] (y-x) = \" using a A by (simp add: pow_eq_div2) ultimately have "min (y - x) (ord a) = ord a" using A(1) a pow_eq_id by auto with \y - x < ord a\ show False by linarith qed show ?thesis unfolding inj_on_def by (metis nat_neq_iff *) qed lemma ord_inj': assumes a: "a \ carrier G" shows "inj_on (\ x . a [^] x) {1 .. ord a}" proof (rule inj_onI, rule ccontr) fix x y :: nat assume A: "x \ {1 .. ord a}" "y \ {1 .. ord a}" "a [^] x = a [^] y" "x\y" { assume "x < ord a" "y < ord a" hence False using ord_inj[OF assms] A unfolding inj_on_def by fastforce } moreover { assume "x = ord a" "y < ord a" hence "a [^] y = a [^] (0::nat)" using pow_ord_eq_1 A by (simp add: a) hence "y=0" using ord_inj[OF assms] \y < ord a\ unfolding inj_on_def by force hence False using A by fastforce } moreover { assume "y = ord a" "x < ord a" hence "a [^] x = a [^] (0::nat)" using pow_ord_eq_1 A by (simp add: a) hence "x=0" using ord_inj[OF assms] \x < ord a\ unfolding inj_on_def by force hence False using A by fastforce } ultimately show False using A by force qed lemma (in group) ord_ge_1: assumes finite: "finite (carrier G)" and a: "a \ carrier G" shows "ord a \ 1" proof - have "((\n::nat. a [^] n) ` {0<..}) \ carrier G" using a by blast then have "finite ((\n::nat. a [^] n) ` {0<..})" using finite_subset finite by auto then have "\ inj_on (\n::nat. a [^] n) {0<..}" using finite_imageD infinite_Ioi by blast then obtain i j::nat where "i \ j" "a [^] i = a [^] j" by (auto simp: inj_on_def) then have "\n::nat. n>0 \ a [^] n = \" by (metis a diffs0_imp_equal pow_eq_div2 neq0_conv) then have "ord a \ 0" by (simp add: ord_eq_0 [OF a]) then show ?thesis by simp qed lemma ord_elems: assumes "finite (carrier G)" "a \ carrier G" shows "{a[^]x | x. x \ (UNIV :: nat set)} = {a[^]x | x. x \ {0 .. ord a - 1}}" (is "?L = ?R") proof show "?R \ ?L" by blast { fix y assume "y \ ?L" then obtain x::nat where x: "y = a[^]x" by auto define r q where "r = x mod ord a" and "q = x div ord a" then have "x = q * ord a + r" by (simp add: div_mult_mod_eq) hence "y = (a[^]ord a)[^]q \ a[^]r" using x assms by (metis mult.commute nat_pow_mult nat_pow_pow) hence "y = a[^]r" using assms by (simp add: pow_ord_eq_1) have "r < ord a" using ord_ge_1[OF assms] by (simp add: r_def) hence "r \ {0 .. ord a - 1}" by (force simp: r_def) hence "y \ {a[^]x | x. x \ {0 .. ord a - 1}}" using \y=a[^]r\ by blast } thus "?L \ ?R" by auto qed lemma (in group) assumes "x \ carrier G" shows finite_cyclic_subgroup: "finite(carrier(subgroup_generated G {x})) \ (\n::nat. n \ 0 \ x [^] n = \)" (is "?fin \ ?nat1") and infinite_cyclic_subgroup: "infinite(carrier(subgroup_generated G {x})) \ (\m n::nat. x [^] m = x [^] n \ m = n)" (is "\ ?fin \ ?nateq") and finite_cyclic_subgroup_int: "finite(carrier(subgroup_generated G {x})) \ (\i::int. i \ 0 \ x [^] i = \)" (is "?fin \ ?int1") and infinite_cyclic_subgroup_int: "infinite(carrier(subgroup_generated G {x})) \ (\i j::int. x [^] i = x [^] j \ i = j)" (is "\ ?fin \ ?inteq") proof - have 1: "\ ?fin" if ?nateq proof - have "infinite (range (\n::nat. x [^] n))" using that range_inj_infinite [of "(\n::nat. x [^] n)"] by (auto simp: inj_on_def) moreover have "range (\n::nat. x [^] n) \ range (\i::int. x [^] i)" apply clarify by (metis assms group.int_pow_neg int_pow_closed int_pow_neg_int is_group local.inv_equality nat_pow_closed r_inv rangeI) ultimately show ?thesis using carrier_subgroup_generated_by_singleton [OF assms] finite_subset by auto qed have 2: "m = n" if mn: "x [^] m = x [^] n" and eq [rule_format]: "?inteq" for m n::nat using eq [of "int m" "int n"] by (simp add: int_pow_int mn) have 3: ?nat1 if non: "\ ?inteq" proof - obtain i j::int where eq: "x [^] i = x [^] j" and "i \ j" using non by auto show ?thesis proof (cases i j rule: linorder_cases) case less then have [simp]: "x [^] (j - i) = \" by (simp add: eq assms int_pow_diff) show ?thesis using less by (rule_tac x="nat (j-i)" in exI) auto next case greater then have [simp]: "x [^] (i - j) = \" by (simp add: eq assms int_pow_diff) then show ?thesis using greater by (rule_tac x="nat (i-j)" in exI) auto qed (use \i \ j\ in auto) qed have 4: "\i::int. (i \ 0) \ x [^] i = \" if "n \ 0" "x [^] n = \" for n::nat apply (rule_tac x="int n" in exI) by (simp add: int_pow_int that) have 5: "finite (carrier (subgroup_generated G {x}))" if "i \ 0" and 1: "x [^] i = \" for i::int proof - obtain n::nat where n: "n > 0" "x [^] n = \" using "1" "3" \i \ 0\ by fastforce have "x [^] a \ ([^]) x ` {0.. {0.. ([^]) x ` {0.. ?nat1" "\ ?fin \ ?nateq" "?fin \ ?int1" "\ ?fin \ ?inteq" using 1 2 3 4 5 by meson+ qed lemma (in group) finite_cyclic_subgroup_order: "x \ carrier G \ finite(carrier(subgroup_generated G {x})) \ ord x \ 0" by (simp add: finite_cyclic_subgroup ord_eq_0) lemma (in group) infinite_cyclic_subgroup_order: "x \ carrier G \ infinite (carrier(subgroup_generated G {x})) \ ord x = 0" by (simp add: finite_cyclic_subgroup_order) lemma generate_pow_on_finite_carrier: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "finite (carrier G)" and a: "a \ carrier G" shows "generate G { a } = { a [^] k | k. k \ (UNIV :: nat set) }" proof show "{ a [^] k | k. k \ (UNIV :: nat set) } \ generate G { a }" proof fix b assume "b \ { a [^] k | k. k \ (UNIV :: nat set) }" then obtain k :: nat where "b = a [^] k" by blast hence "b = a [^] (int k)" by (simp add: int_pow_int) thus "b \ generate G { a }" unfolding generate_pow[OF a] by blast qed next show "generate G { a } \ { a [^] k | k. k \ (UNIV :: nat set) }" proof fix b assume "b \ generate G { a }" then obtain k :: int where k: "b = a [^] k" unfolding generate_pow[OF a] by blast show "b \ { a [^] k | k. k \ (UNIV :: nat set) }" proof (cases "k < 0") assume "\ k < 0" hence "b = a [^] (nat k)" by (simp add: k) thus ?thesis by blast next assume "k < 0" hence b: "b = inv (a [^] (nat (- k)))" using k a by (auto simp: int_pow_neg) obtain m where m: "ord a * m \ nat (- k)" by (metis assms mult.left_neutral mult_le_mono1 ord_ge_1) hence "a [^] (ord a * m) = \" by (metis a nat_pow_one nat_pow_pow pow_ord_eq_1) then obtain k' :: nat where "(a [^] (nat (- k))) \ (a [^] k') = \" using m a nat_le_iff_add nat_pow_mult by auto hence "b = a [^] k'" using b a by (metis inv_unique' nat_pow_closed nat_pow_comm) thus "b \ { a [^] k | k. k \ (UNIV :: nat set) }" by blast qed qed qed lemma ord_elems_inf_carrier: assumes "a \ carrier G" "ord a \ 0" shows "{a[^]x | x. x \ (UNIV :: nat set)} = {a[^]x | x. x \ {0 .. ord a - 1}}" (is "?L = ?R") proof show "?R \ ?L" by blast { fix y assume "y \ ?L" then obtain x::nat where x: "y = a[^]x" by auto define r q where "r = x mod ord a" and "q = x div ord a" then have "x = q * ord a + r" by (simp add: div_mult_mod_eq) hence "y = (a[^]ord a)[^]q \ a[^]r" using x assms by (metis mult.commute nat_pow_mult nat_pow_pow) hence "y = a[^]r" using assms by simp have "r < ord a" using assms by (simp add: r_def) hence "r \ {0 .. ord a - 1}" by (force simp: r_def) hence "y \ {a[^]x | x. x \ {0 .. ord a - 1}}" using \y=a[^]r\ by blast } thus "?L \ ?R" by auto qed lemma generate_pow_nat: assumes a: "a \ carrier G" and "ord a \ 0" shows "generate G { a } = { a [^] k | k. k \ (UNIV :: nat set) }" proof show "{ a [^] k | k. k \ (UNIV :: nat set) } \ generate G { a }" proof fix b assume "b \ { a [^] k | k. k \ (UNIV :: nat set) }" then obtain k :: nat where "b = a [^] k" by blast hence "b = a [^] (int k)" by (simp add: int_pow_int) thus "b \ generate G { a }" unfolding generate_pow[OF a] by blast qed next show "generate G { a } \ { a [^] k | k. k \ (UNIV :: nat set) }" proof fix b assume "b \ generate G { a }" then obtain k :: int where k: "b = a [^] k" unfolding generate_pow[OF a] by blast show "b \ { a [^] k | k. k \ (UNIV :: nat set) }" proof (cases "k < 0") assume "\ k < 0" hence "b = a [^] (nat k)" by (simp add: k) thus ?thesis by blast next assume "k < 0" hence b: "b = inv (a [^] (nat (- k)))" using k a by (auto simp: int_pow_neg) obtain m where m: "ord a * m \ nat (- k)" by (metis assms(2) dvd_imp_le dvd_triv_right le_zero_eq mult_eq_0_iff not_gr_zero) hence "a [^] (ord a * m) = \" by (metis a nat_pow_one nat_pow_pow pow_ord_eq_1) then obtain k' :: nat where "(a [^] (nat (- k))) \ (a [^] k') = \" using m a nat_le_iff_add nat_pow_mult by auto hence "b = a [^] k'" using b a by (metis inv_unique' nat_pow_closed nat_pow_comm) thus "b \ { a [^] k | k. k \ (UNIV :: nat set) }" by blast qed qed qed lemma generate_pow_card: assumes a: "a \ carrier G" shows "ord a = card (generate G { a })" proof (cases "ord a = 0") case True then have "infinite (carrier (subgroup_generated G {a}))" using infinite_cyclic_subgroup_order[OF a] by auto then have "infinite (generate G {a})" unfolding subgroup_generated_def using a by simp then show ?thesis using \ord a = 0\ by auto next case False note finite_subgroup = this then have "generate G { a } = (([^]) a) ` {0..ord a - 1}" using generate_pow_nat ord_elems_inf_carrier a by auto hence "card (generate G {a}) = card {0..ord a - 1}" using ord_inj[OF a] card_image by metis also have "... = ord a" using finite_subgroup by auto finally show ?thesis.. qed lemma (in group) cyclic_order_is_ord: assumes "g \ carrier G" shows "ord g = order (subgroup_generated G {g})" unfolding order_def subgroup_generated_def using assms generate_pow_card by simp lemma ord_dvd_group_order: assumes "a \ carrier G" shows "(ord a) dvd (order G)" using lagrange[OF generate_is_subgroup[of "{a}"]] assms unfolding generate_pow_card[OF assms] by (metis dvd_triv_right empty_subsetI insert_subset) lemma (in group) pow_order_eq_1: assumes "a \ carrier G" shows "a [^] order G = \" using assms by (metis nat_pow_pow ord_dvd_group_order pow_ord_eq_1 dvdE nat_pow_one) lemma dvd_gcd: fixes a b :: nat obtains q where "a * (b div gcd a b) = b*q" proof have "a * (b div gcd a b) = (a div gcd a b) * b" by (simp add: div_mult_swap dvd_div_mult) also have "\ = b * (a div gcd a b)" by simp finally show "a * (b div gcd a b) = b * (a div gcd a b) " . qed lemma (in group) ord_le_group_order: assumes finite: "finite (carrier G)" and a: "a \ carrier G" shows "ord a \ order G" by (simp add: a dvd_imp_le local.finite ord_dvd_group_order order_gt_0_iff_finite) lemma (in group) ord_pow_gen: assumes "x \ carrier G" shows "ord (pow G x k) = (if k = 0 then 1 else ord x div gcd (ord x) k)" proof - have "ord (x [^] k) = ord x div gcd (ord x) k" if "0 < k" proof - have "(d dvd k * n) = (d div gcd (d) k dvd n)" for d n using that by (simp add: div_dvd_iff_mult gcd_mult_distrib_nat mult.commute) then show ?thesis using that by (auto simp add: assms ord_unique nat_pow_pow pow_eq_id) qed then show ?thesis by auto qed lemma (in group) assumes finite': "finite (carrier G)" "a \ carrier G" shows pow_ord_eq_ord_iff: "group.ord G (a [^] k) = ord a \ coprime k (ord a)" (is "?L \ ?R") using assms ord_ge_1 [OF assms] by (auto simp: div_eq_dividend_iff ord_pow_gen coprime_iff_gcd_eq_1 gcd.commute split: if_split_asm) lemma element_generates_subgroup: assumes finite[simp]: "finite (carrier G)" assumes a[simp]: "a \ carrier G" shows "subgroup {a [^] i | i. i \ {0 .. ord a - 1}} G" using generate_is_subgroup[of "{ a }"] assms(2) generate_pow_on_finite_carrier[OF assms] unfolding ord_elems[OF assms] by auto end section \Number of Roots of a Polynomial\ text_raw \\label{sec:number-roots}\ definition mult_of :: "('a, 'b) ring_scheme \ 'a monoid" where "mult_of R \ \ carrier = carrier R - {\\<^bsub>R\<^esub>}, mult = mult R, one = \\<^bsub>R\<^esub>\" lemma carrier_mult_of [simp]: "carrier (mult_of R) = carrier R - {\\<^bsub>R\<^esub>}" by (simp add: mult_of_def) lemma mult_mult_of [simp]: "mult (mult_of R) = mult R" by (simp add: mult_of_def) lemma nat_pow_mult_of: "([^]\<^bsub>mult_of R\<^esub>) = (([^]\<^bsub>R\<^esub>) :: _ \ nat \ _)" by (simp add: mult_of_def fun_eq_iff nat_pow_def) lemma one_mult_of [simp]: "\\<^bsub>mult_of R\<^esub> = \\<^bsub>R\<^esub>" by (simp add: mult_of_def) lemmas mult_of_simps = carrier_mult_of mult_mult_of nat_pow_mult_of one_mult_of context field begin lemma mult_of_is_Units: "mult_of R = units_of R" unfolding mult_of_def units_of_def using field_Units by auto lemma m_inv_mult_of: "\x. x \ carrier (mult_of R) \ m_inv (mult_of R) x = m_inv R x" using mult_of_is_Units units_of_inv unfolding units_of_def by simp lemma (in field) field_mult_group: "group (mult_of R)" proof (rule groupI) show "\y\carrier (mult_of R). y \\<^bsub>mult_of R\<^esub> x = \\<^bsub>mult_of R\<^esub>" if "x \ carrier (mult_of R)" for x using group.l_inv_ex mult_of_is_Units that units_group by fastforce qed (auto simp: m_assoc dest: integral) lemma finite_mult_of: "finite (carrier R) \ finite (carrier (mult_of R))" by simp lemma order_mult_of: "finite (carrier R) \ order (mult_of R) = order R - 1" unfolding order_def carrier_mult_of by (simp add: card.remove) end lemma (in monoid) Units_pow_closed : fixes d :: nat assumes "x \ Units G" shows "x [^] d \ Units G" by (metis assms group.is_monoid monoid.nat_pow_closed units_group units_of_carrier units_of_pow) lemma (in ring) r_right_minus_eq[simp]: assumes "a \ carrier R" "b \ carrier R" shows "a \ b = \ \ a = b" using assms by (metis a_minus_def add.inv_closed minus_equality r_neg) context UP_cring begin lemma is_UP_cring: "UP_cring R" by (unfold_locales) lemma is_UP_ring: shows "UP_ring R" by (unfold_locales) end context UP_domain begin lemma roots_bound: assumes f [simp]: "f \ carrier P" assumes f_not_zero: "f \ \\<^bsub>P\<^esub>" assumes finite: "finite (carrier R)" shows "finite {a \ carrier R . eval R R id a f = \} \ card {a \ carrier R . eval R R id a f = \} \ deg R f" using f f_not_zero proof (induction "deg R f" arbitrary: f) case 0 have "\x. eval R R id x f \ \" proof - fix x have "(\i\{..deg R f}. id (coeff P f i) \ x [^] i) \ \" using 0 lcoeff_nonzero_nonzero[where p = f] by simp thus "eval R R id x f \ \" using 0 unfolding eval_def P_def by simp qed then have *: "{a \ carrier R. eval R R (\a. a) a f = \} = {}" by (auto simp: id_def) show ?case by (simp add: *) next case (Suc x) show ?case proof (cases "\ a \ carrier R . eval R R id a f = \") case True then obtain a where a_carrier[simp]: "a \ carrier R" and a_root: "eval R R id a f = \" by blast have R_not_triv: "carrier R \ {\}" by (metis R.one_zeroI R.zero_not_one) obtain q where q: "(q \ carrier P)" and f: "f = (monom P \\<^bsub>R\<^esub> 1 \\<^bsub> P\<^esub> monom P a 0) \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> monom P (eval R R id a f) 0" using remainder_theorem[OF Suc.prems(1) a_carrier R_not_triv] by auto hence lin_fac: "f = (monom P \\<^bsub>R\<^esub> 1 \\<^bsub> P\<^esub> monom P a 0) \\<^bsub>P\<^esub> q" using q by (simp add: a_root) have deg: "deg R (monom P \\<^bsub>R\<^esub> 1 \\<^bsub> P\<^esub> monom P a 0) = 1" using a_carrier by (simp add: deg_minus_eq) hence mon_not_zero: "(monom P \\<^bsub>R\<^esub> 1 \\<^bsub> P\<^esub> monom P a 0) \ \\<^bsub>P\<^esub>" by (fastforce simp del: r_right_minus_eq) have q_not_zero: "q \ \\<^bsub>P\<^esub>" using Suc by (auto simp add : lin_fac) hence "deg R q = x" using Suc deg deg_mult[OF mon_not_zero q_not_zero _ q] by (simp add : lin_fac) hence q_IH: "finite {a \ carrier R . eval R R id a q = \} \ card {a \ carrier R . eval R R id a q = \} \ x" using Suc q q_not_zero by blast have subs: "{a \ carrier R . eval R R id a f = \} \ {a \ carrier R . eval R R id a q = \} \ {a}" (is "?L \ ?R \ {a}") using a_carrier \q \ _\ by (auto simp: evalRR_simps lin_fac R.integral_iff) have "{a \ carrier R . eval R R id a f = \} \ insert a {a \ carrier R . eval R R id a q = \}" using subs by auto hence "card {a \ carrier R . eval R R id a f = \} \ card (insert a {a \ carrier R . eval R R id a q = \})" using q_IH by (blast intro: card_mono) also have "\ \ deg R f" using q_IH \Suc x = _\ by (simp add: card_insert_if) finally show ?thesis using q_IH \Suc x = _\ using finite by force next case False hence "card {a \ carrier R. eval R R id a f = \} = 0" using finite by auto also have "\ \ deg R f" by simp finally show ?thesis using finite by auto qed qed end lemma (in domain) num_roots_le_deg : fixes p d :: nat assumes finite: "finite (carrier R)" assumes d_neq_zero: "d \ 0" shows "card {x \ carrier R. x [^] d = \} \ d" proof - let ?f = "monom (UP R) \\<^bsub>R\<^esub> d \\<^bsub> (UP R)\<^esub> monom (UP R) \\<^bsub>R\<^esub> 0" have one_in_carrier: "\ \ carrier R" by simp interpret R: UP_domain R "UP R" by (unfold_locales) have "deg R ?f = d" using d_neq_zero by (simp add: R.deg_minus_eq) hence f_not_zero: "?f \ \\<^bsub>UP R\<^esub>" using d_neq_zero by (auto simp add : R.deg_nzero_nzero) have roots_bound: "finite {a \ carrier R . eval R R id a ?f = \} \ card {a \ carrier R . eval R R id a ?f = \} \ deg R ?f" using finite by (intro R.roots_bound[OF _ f_not_zero]) simp have subs: "{x \ carrier R. x [^] d = \} \ {a \ carrier R . eval R R id a ?f = \}" by (auto simp: R.evalRR_simps) then have "card {x \ carrier R. x [^] d = \} \ card {a \ carrier R. eval R R id a ?f = \}" using finite by (simp add : card_mono) thus ?thesis using \deg R ?f = d\ roots_bound by linarith qed section \The Multiplicative Group of a Field\ text_raw \\label{sec:mult-group}\ text \ In this section we show that the multiplicative group of a finite field is generated by a single element, i.e. it is cyclic. The proof is inspired - by the first proof given in the survey~@{cite "conrad-cyclicity"}. + by the first proof given in the survey~\<^cite>\"conrad-cyclicity"\. \ context field begin lemma num_elems_of_ord_eq_phi': assumes finite: "finite (carrier R)" and dvd: "d dvd order (mult_of R)" and exists: "\a\carrier (mult_of R). group.ord (mult_of R) a = d" shows "card {a \ carrier (mult_of R). group.ord (mult_of R) a = d} = phi' d" proof - note mult_of_simps[simp] have finite': "finite (carrier (mult_of R))" using finite by (rule finite_mult_of) interpret G:group "mult_of R" rewrites "([^]\<^bsub>mult_of R\<^esub>) = (([^]) :: _ \ nat \ _)" and "\\<^bsub>mult_of R\<^esub> = \" by (rule field_mult_group) simp_all from exists obtain a where a: "a \ carrier (mult_of R)" and ord_a: "group.ord (mult_of R) a = d" by (auto simp add: card_gt_0_iff) have set_eq1: "{a[^]n| n. n \ {1 .. d}} = {x \ carrier (mult_of R). x [^] d = \}" proof (rule card_seteq) show "finite {x \ carrier (mult_of R). x [^] d = \}" using finite by auto show "{a[^]n| n. n \ {1 ..d}} \ {x \ carrier (mult_of R). x[^]d = \}" proof fix x assume "x \ {a[^]n | n. n \ {1 .. d}}" then obtain n where n: "x = a[^]n \ n \ {1 .. d}" by auto have "x[^]d =(a[^]d)[^]n" using n a ord_a by (simp add:nat_pow_pow mult.commute) hence "x[^]d = \" using ord_a G.pow_ord_eq_1[OF a] by fastforce thus "x \ {x \ carrier (mult_of R). x[^]d = \}" using G.nat_pow_closed[OF a] n by blast qed show "card {x \ carrier (mult_of R). x [^] d = \} \ card {a[^]n | n. n \ {1 .. d}}" proof - have *: "{a[^]n | n. n \ {1 .. d }} = ((\ n. a[^]n) ` {1 .. d})" by auto have "0 < order (mult_of R)" unfolding order_mult_of[OF finite] using card_mono[OF finite, of "{\, \}"] by (simp add: order_def) have "card {x \ carrier (mult_of R). x [^] d = \} \ card {x \ carrier R. x [^] d = \}" using finite by (auto intro: card_mono) also have "\ \ d" using \0 < order (mult_of R)\ num_roots_le_deg[OF finite, of d] by (simp add : dvd_pos_nat[OF _ \d dvd order (mult_of R)\]) finally show ?thesis using G.ord_inj'[OF a] ord_a * by (simp add: card_image) qed qed have set_eq2: "{x \ carrier (mult_of R) . group.ord (mult_of R) x = d} = (\ n . a[^]n) ` {n \ {1 .. d}. group.ord (mult_of R) (a[^]n) = d}" (is "?L = ?R") proof { fix x assume x: "x \ (carrier (mult_of R)) \ group.ord (mult_of R) x = d" hence "x \ {x \ carrier (mult_of R). x [^] d = \}" by (simp add: G.pow_ord_eq_1[of x, symmetric]) then obtain n where n: "x = a[^]n \ n \ {1 .. d}" using set_eq1 by blast hence "x \ ?R" using x by fast } thus "?L \ ?R" by blast show "?R \ ?L" using a by (auto simp add: carrier_mult_of[symmetric] simp del: carrier_mult_of) qed have "inj_on (\ n . a[^]n) {n \ {1 .. d}. group.ord (mult_of R) (a[^]n) = d}" using G.ord_inj'[OF a, unfolded ord_a] unfolding inj_on_def by fast hence "card ((\n. a[^]n) ` {n \ {1 .. d}. group.ord (mult_of R) (a[^]n) = d}) = card {k \ {1 .. d}. group.ord (mult_of R) (a[^]k) = d}" using card_image by blast thus ?thesis using set_eq2 G.pow_ord_eq_ord_iff[OF finite' \a \ _\, unfolded ord_a] by (simp add: phi'_def) qed end theorem (in field) finite_field_mult_group_has_gen : assumes finite: "finite (carrier R)" shows "\ a \ carrier (mult_of R) . carrier (mult_of R) = {a[^]i | i::nat . i \ UNIV}" proof - note mult_of_simps[simp] have finite': "finite (carrier (mult_of R))" using finite by (rule finite_mult_of) interpret G: group "mult_of R" rewrites "([^]\<^bsub>mult_of R\<^esub>) = (([^]) :: _ \ nat \ _)" and "\\<^bsub>mult_of R\<^esub> = \" by (rule field_mult_group) (simp_all add: fun_eq_iff nat_pow_def) let ?N = "\ x . card {a \ carrier (mult_of R). group.ord (mult_of R) a = x}" have "0 < order R - 1" unfolding order_def using card_mono[OF finite, of "{\, \}"] by simp then have *: "0 < order (mult_of R)" using assms by (simp add: order_mult_of) have fin: "finite {d. d dvd order (mult_of R) }" using dvd_nat_bounds[OF *] by force have "(\d | d dvd order (mult_of R). ?N d) = card (UN d:{d . d dvd order (mult_of R) }. {a \ carrier (mult_of R). group.ord (mult_of R) a = d})" (is "_ = card ?U") using fin finite by (subst card_UN_disjoint) auto also have "?U = carrier (mult_of R)" proof { fix x assume x: "x \ carrier (mult_of R)" hence x': "x\carrier (mult_of R)" by simp then have "group.ord (mult_of R) x dvd order (mult_of R)" using G.ord_dvd_group_order by blast hence "x \ ?U" using dvd_nat_bounds[of "order (mult_of R)" "group.ord (mult_of R) x"] x by blast } thus "carrier (mult_of R) \ ?U" by blast qed auto also have "card ... = order (mult_of R)" using order_mult_of finite' by (simp add: order_def) finally have sum_Ns_eq: "(\d | d dvd order (mult_of R). ?N d) = order (mult_of R)" . { fix d assume d: "d dvd order (mult_of R)" have "card {a \ carrier (mult_of R). group.ord (mult_of R) a = d} \ phi' d" proof cases assume "card {a \ carrier (mult_of R). group.ord (mult_of R) a = d} = 0" thus ?thesis by presburger next assume "card {a \ carrier (mult_of R). group.ord (mult_of R) a = d} \ 0" hence "\a \ carrier (mult_of R). group.ord (mult_of R) a = d" by (auto simp: card_eq_0_iff) thus ?thesis using num_elems_of_ord_eq_phi'[OF finite d] by auto qed } hence all_le: "\i. i \ {d. d dvd order (mult_of R) } \ (\i. card {a \ carrier (mult_of R). group.ord (mult_of R) a = i}) i \ (\i. phi' i) i" by fast hence le: "(\i | i dvd order (mult_of R). ?N i) \ (\i | i dvd order (mult_of R). phi' i)" using sum_mono[of "{d . d dvd order (mult_of R)}" "\i. card {a \ carrier (mult_of R). group.ord (mult_of R) a = i}"] by presburger have "order (mult_of R) = (\d | d dvd order (mult_of R). phi' d)" using * by (simp add: sum_phi'_factors) hence eq: "(\i | i dvd order (mult_of R). ?N i) = (\i | i dvd order (mult_of R). phi' i)" using le sum_Ns_eq by presburger have "\i. i \ {d. d dvd order (mult_of R) } \ ?N i = (\i. phi' i) i" proof (rule ccontr) fix i assume i1: "i \ {d. d dvd order (mult_of R)}" and "?N i \ phi' i" hence "?N i = 0" using num_elems_of_ord_eq_phi'[OF finite, of i] by (auto simp: card_eq_0_iff) moreover have "0 < i" using * i1 by (simp add: dvd_nat_bounds[of "order (mult_of R)" i]) ultimately have "?N i < phi' i" using phi'_nonzero by presburger hence "(\i | i dvd order (mult_of R). ?N i) < (\i | i dvd order (mult_of R). phi' i)" using sum_strict_mono_ex1[OF fin, of "?N" "\ i . phi' i"] i1 all_le by auto thus False using eq by force qed hence "?N (order (mult_of R)) > 0" using * by (simp add: phi'_nonzero) then obtain a where a: "a \ carrier (mult_of R)" and a_ord: "group.ord (mult_of R) a = order (mult_of R)" by (auto simp add: card_gt_0_iff) hence set_eq: "{a[^]i | i::nat. i \ UNIV} = (\x. a[^]x) ` {0 .. group.ord (mult_of R) a - 1}" using G.ord_elems[OF finite'] by auto have card_eq: "card ((\x. a[^]x) ` {0 .. group.ord (mult_of R) a - 1}) = card {0 .. group.ord (mult_of R) a - 1}" by (intro card_image G.ord_inj finite' a) hence "card ((\ x . a[^]x) ` {0 .. group.ord (mult_of R) a - 1}) = card {0 ..order (mult_of R) - 1}" using assms by (simp add: card_eq a_ord) hence card_R_minus_1: "card {a[^]i | i::nat. i \ UNIV} = order (mult_of R)" using * by (subst set_eq) auto have **: "{a[^]i | i::nat. i \ UNIV} \ carrier (mult_of R)" using G.nat_pow_closed[OF a] by auto with _ have "carrier (mult_of R) = {a[^]i|i::nat. i \ UNIV}" by (rule card_seteq[symmetric]) (simp_all add: card_R_minus_1 finite order_def del: UNIV_I) thus ?thesis using a by blast qed end diff --git a/src/HOL/Algebra/Sylow.thy b/src/HOL/Algebra/Sylow.thy --- a/src/HOL/Algebra/Sylow.thy +++ b/src/HOL/Algebra/Sylow.thy @@ -1,355 +1,355 @@ (* Title: HOL/Algebra/Sylow.thy Author: Florian Kammueller, with new proofs by L C Paulson *) theory Sylow imports Coset Exponent begin -text \See also @{cite "Kammueller-Paulson:1999"}.\ +text \See also \<^cite>\"Kammueller-Paulson:1999"\.\ text \The combinatorial argument is in theory \Exponent\.\ lemma le_extend_mult: "\0 < c; a \ b\ \ a \ b * c" for c :: nat by (metis divisors_zero dvd_triv_left leI less_le_trans nat_dvd_not_less zero_less_iff_neq_zero) locale sylow = group + fixes p and a and m and calM and RelM assumes prime_p: "prime p" and order_G: "order G = (p^a) * m" and finite_G[iff]: "finite (carrier G)" defines "calM \ {s. s \ carrier G \ card s = p^a}" and "RelM \ {(N1, N2). N1 \ calM \ N2 \ calM \ (\g \ carrier G. N1 = N2 #> g)}" begin lemma RelM_refl_on: "refl_on calM RelM" by (auto simp: refl_on_def RelM_def calM_def) (blast intro!: coset_mult_one [symmetric]) lemma RelM_sym: "sym RelM" proof (unfold sym_def RelM_def, clarify) fix y g assume "y \ calM" and g: "g \ carrier G" then have "y = y #> g #> (inv g)" by (simp add: coset_mult_assoc calM_def) then show "\g'\carrier G. y = y #> g #> g'" by (blast intro: g) qed lemma RelM_trans: "trans RelM" by (auto simp add: trans_def RelM_def calM_def coset_mult_assoc) lemma RelM_equiv: "equiv calM RelM" unfolding equiv_def by (blast intro: RelM_refl_on RelM_sym RelM_trans) lemma M_subset_calM_prep: "M' \ calM // RelM \ M' \ calM" unfolding RelM_def by (blast elim!: quotientE) end subsection \Main Part of the Proof\ locale sylow_central = sylow + fixes H and M1 and M assumes M_in_quot: "M \ calM // RelM" and not_dvd_M: "\ (p ^ Suc (multiplicity p m) dvd card M)" and M1_in_M: "M1 \ M" defines "H \ {g. g \ carrier G \ M1 #> g = M1}" begin lemma M_subset_calM: "M \ calM" by (rule M_in_quot [THEN M_subset_calM_prep]) lemma card_M1: "card M1 = p^a" using M1_in_M M_subset_calM calM_def by blast lemma exists_x_in_M1: "\x. x \ M1" using prime_p [THEN prime_gt_Suc_0_nat] card_M1 by (metis Suc_lessD card_eq_0_iff empty_subsetI equalityI gr_implies_not0 nat_zero_less_power_iff subsetI) lemma M1_subset_G [simp]: "M1 \ carrier G" using M1_in_M M_subset_calM calM_def mem_Collect_eq subsetCE by blast lemma M1_inj_H: "\f \ H\M1. inj_on f H" proof - from exists_x_in_M1 obtain m1 where m1M: "m1 \ M1".. have m1: "m1 \ carrier G" by (simp add: m1M M1_subset_G [THEN subsetD]) show ?thesis proof show "inj_on (\z\H. m1 \ z) H" by (simp add: H_def inj_on_def m1) show "restrict ((\) m1) H \ H \ M1" proof (rule restrictI) fix z assume zH: "z \ H" show "m1 \ z \ M1" proof - from zH have zG: "z \ carrier G" and M1zeq: "M1 #> z = M1" by (auto simp add: H_def) show ?thesis by (rule subst [OF M1zeq]) (simp add: m1M zG rcosI) qed qed qed qed end subsection \Discharging the Assumptions of \sylow_central\\ context sylow begin lemma EmptyNotInEquivSet: "{} \ calM // RelM" by (blast elim!: quotientE dest: RelM_equiv [THEN equiv_class_self]) lemma existsM1inM: "M \ calM // RelM \ \M1. M1 \ M" using RelM_equiv equiv_Eps_in by blast lemma zero_less_o_G: "0 < order G" by (simp add: order_def card_gt_0_iff carrier_not_empty) lemma zero_less_m: "m > 0" using zero_less_o_G by (simp add: order_G) lemma card_calM: "card calM = (p^a) * m choose p^a" by (simp add: calM_def n_subsets order_G [symmetric] order_def) lemma zero_less_card_calM: "card calM > 0" by (simp add: card_calM zero_less_binomial le_extend_mult zero_less_m) lemma max_p_div_calM: "\ (p ^ Suc (multiplicity p m) dvd card calM)" proof assume "p ^ Suc (multiplicity p m) dvd card calM" with zero_less_card_calM prime_p have "Suc (multiplicity p m) \ multiplicity p (card calM)" by (intro multiplicity_geI) auto then have "multiplicity p m < multiplicity p (card calM)" by simp also have "multiplicity p m = multiplicity p (card calM)" by (simp add: const_p_fac prime_p zero_less_m card_calM) finally show False by simp qed lemma finite_calM: "finite calM" unfolding calM_def by (rule finite_subset [where B = "Pow (carrier G)"]) auto lemma lemma_A1: "\M \ calM // RelM. \ (p ^ Suc (multiplicity p m) dvd card M)" using RelM_equiv equiv_imp_dvd_card finite_calM max_p_div_calM by blast end subsubsection \Introduction and Destruct Rules for \H\\ context sylow_central begin lemma H_I: "\g \ carrier G; M1 #> g = M1\ \ g \ H" by (simp add: H_def) lemma H_into_carrier_G: "x \ H \ x \ carrier G" by (simp add: H_def) lemma in_H_imp_eq: "g \ H \ M1 #> g = M1" by (simp add: H_def) lemma H_m_closed: "\x \ H; y \ H\ \ x \ y \ H" by (simp add: H_def coset_mult_assoc [symmetric]) lemma H_not_empty: "H \ {}" by (force simp add: H_def intro: exI [of _ \]) lemma H_is_subgroup: "subgroup H G" proof (rule subgroupI) show "H \ carrier G" using H_into_carrier_G by blast show "\a. a \ H \ inv a \ H" by (metis H_I H_into_carrier_G H_m_closed M1_subset_G Units_eq Units_inv_closed Units_inv_inv coset_mult_inv1 in_H_imp_eq) show "\a b. \a \ H; b \ H\ \ a \ b \ H" by (blast intro: H_m_closed) qed (use H_not_empty in auto) lemma rcosetGM1g_subset_G: "\g \ carrier G; x \ M1 #> g\ \ x \ carrier G" by (blast intro: M1_subset_G [THEN r_coset_subset_G, THEN subsetD]) lemma finite_M1: "finite M1" by (rule finite_subset [OF M1_subset_G finite_G]) lemma finite_rcosetGM1g: "g \ carrier G \ finite (M1 #> g)" using rcosetGM1g_subset_G finite_G M1_subset_G cosets_finite rcosetsI by blast lemma M1_cardeq_rcosetGM1g: "g \ carrier G \ card (M1 #> g) = card M1" by (metis M1_subset_G card_rcosets_equal rcosetsI) lemma M1_RelM_rcosetGM1g: assumes "g \ carrier G" shows "(M1, M1 #> g) \ RelM" proof - have "M1 #> g \ carrier G" by (simp add: assms r_coset_subset_G) moreover have "card (M1 #> g) = p ^ a" using assms by (simp add: card_M1 M1_cardeq_rcosetGM1g) moreover have "\h\carrier G. M1 = M1 #> g #> h" by (metis assms M1_subset_G coset_mult_assoc coset_mult_one r_inv_ex) ultimately show ?thesis by (simp add: RelM_def calM_def card_M1) qed end subsection \Equal Cardinalities of \M\ and the Set of Cosets\ text \Injections between \<^term>\M\ and \<^term>\rcosets\<^bsub>G\<^esub> H\ show that their cardinalities are equal.\ lemma ElemClassEquiv: "\equiv A r; C \ A // r\ \ \x \ C. \y \ C. (x, y) \ r" unfolding equiv_def quotient_def sym_def trans_def by blast context sylow_central begin lemma M_elem_map: "M2 \ M \ \g. g \ carrier G \ M1 #> g = M2" using M1_in_M M_in_quot [THEN RelM_equiv [THEN ElemClassEquiv]] by (simp add: RelM_def) (blast dest!: bspec) lemmas M_elem_map_carrier = M_elem_map [THEN someI_ex, THEN conjunct1] lemmas M_elem_map_eq = M_elem_map [THEN someI_ex, THEN conjunct2] lemma M_funcset_rcosets_H: "(\x\M. H #> (SOME g. g \ carrier G \ M1 #> g = x)) \ M \ rcosets H" by (metis (lifting) H_is_subgroup M_elem_map_carrier rcosetsI restrictI subgroup.subset) lemma inj_M_GmodH: "\f \ M \ rcosets H. inj_on f M" proof let ?inv = "\x. SOME g. g \ carrier G \ M1 #> g = x" show "inj_on (\x\M. H #> ?inv x) M" proof (rule inj_onI, simp) fix x y assume eq: "H #> ?inv x = H #> ?inv y" and xy: "x \ M" "y \ M" have "x = M1 #> ?inv x" by (simp add: M_elem_map_eq \x \ M\) also have "... = M1 #> ?inv y" proof (rule coset_mult_inv1 [OF in_H_imp_eq [OF coset_join1]]) show "H #> ?inv x \ inv (?inv y) = H" by (simp add: H_into_carrier_G M_elem_map_carrier xy coset_mult_inv2 eq subsetI) qed (simp_all add: H_is_subgroup M_elem_map_carrier xy) also have "... = y" using M_elem_map_eq \y \ M\ by simp finally show "x=y" . qed show "(\x\M. H #> ?inv x) \ M \ rcosets H" by (rule M_funcset_rcosets_H) qed end subsubsection \The Opposite Injection\ context sylow_central begin lemma H_elem_map: "H1 \ rcosets H \ \g. g \ carrier G \ H #> g = H1" by (auto simp: RCOSETS_def) lemmas H_elem_map_carrier = H_elem_map [THEN someI_ex, THEN conjunct1] lemmas H_elem_map_eq = H_elem_map [THEN someI_ex, THEN conjunct2] lemma rcosets_H_funcset_M: "(\C \ rcosets H. M1 #> (SOME g. g \ carrier G \ H #> g = C)) \ rcosets H \ M" using in_quotient_imp_closed [OF RelM_equiv M_in_quot _ M1_RelM_rcosetGM1g] by (simp add: M1_in_M H_elem_map_carrier RCOSETS_def) lemma inj_GmodH_M: "\g \ rcosets H\M. inj_on g (rcosets H)" proof let ?inv = "\x. SOME g. g \ carrier G \ H #> g = x" show "inj_on (\C\rcosets H. M1 #> ?inv C) (rcosets H)" proof (rule inj_onI, simp) fix x y assume eq: "M1 #> ?inv x = M1 #> ?inv y" and xy: "x \ rcosets H" "y \ rcosets H" have "x = H #> ?inv x" by (simp add: H_elem_map_eq \x \ rcosets H\) also have "... = H #> ?inv y" proof (rule coset_mult_inv1 [OF coset_join2]) show "?inv x \ inv (?inv y) \ carrier G" by (simp add: H_elem_map_carrier \x \ rcosets H\ \y \ rcosets H\) then show "(?inv x) \ inv (?inv y) \ H" by (simp add: H_I H_elem_map_carrier xy coset_mult_inv2 eq) show "H \ carrier G" by (simp add: H_is_subgroup subgroup.subset) qed (simp_all add: H_is_subgroup H_elem_map_carrier xy) also have "... = y" by (simp add: H_elem_map_eq \y \ rcosets H\) finally show "x=y" . qed show "(\C\rcosets H. M1 #> ?inv C) \ rcosets H \ M" using rcosets_H_funcset_M by blast qed lemma calM_subset_PowG: "calM \ Pow (carrier G)" by (auto simp: calM_def) lemma finite_M: "finite M" by (metis M_subset_calM finite_calM rev_finite_subset) lemma cardMeqIndexH: "card M = card (rcosets H)" using inj_M_GmodH inj_GmodH_M by (blast intro: card_bij finite_M H_is_subgroup rcosets_subset_PowG [THEN finite_subset]) lemma index_lem: "card M * card H = order G" by (simp add: cardMeqIndexH lagrange H_is_subgroup) lemma card_H_eq: "card H = p^a" proof (rule antisym) show "p^a \ card H" proof (rule dvd_imp_le) show "p ^ a dvd card H" apply (rule div_combine [OF prime_imp_prime_elem[OF prime_p] not_dvd_M]) by (simp add: index_lem multiplicity_dvd order_G power_add) show "0 < card H" by (blast intro: subgroup.finite_imp_card_positive H_is_subgroup) qed next show "card H \ p^a" using M1_inj_H card_M1 card_inj finite_M1 by fastforce qed end lemma (in sylow) sylow_thm: "\H. subgroup H G \ card H = p^a" proof - obtain M where M: "M \ calM // RelM" "\ (p ^ Suc (multiplicity p m) dvd card M)" using lemma_A1 by blast then obtain M1 where "M1 \ M" by (metis existsM1inM) define H where "H \ {g. g \ carrier G \ M1 #> g = M1}" with M \M1 \ M\ interpret sylow_central G p a m calM RelM H M1 M by unfold_locales (auto simp add: H_def calM_def RelM_def) show ?thesis using H_is_subgroup card_H_eq by blast qed text \Needed because the locale's automatic definition refers to \<^term>\semigroup G\ and \<^term>\group_axioms G\ rather than simply to \<^term>\group G\.\ lemma sylow_eq: "sylow G p a m \ group G \ sylow_axioms G p a m" by (simp add: sylow_def group_def) subsection \Sylow's Theorem\ theorem sylow_thm: "\prime p; group G; order G = (p^a) * m; finite (carrier G)\ \ \H. subgroup H G \ card H = p^a" by (rule sylow.sylow_thm [of G p a m]) (simp add: sylow_eq sylow_axioms_def) end diff --git a/src/HOL/Algebra/UnivPoly.thy b/src/HOL/Algebra/UnivPoly.thy --- a/src/HOL/Algebra/UnivPoly.thy +++ b/src/HOL/Algebra/UnivPoly.thy @@ -1,1839 +1,1839 @@ (* Title: HOL/Algebra/UnivPoly.thy Author: Clemens Ballarin, started 9 December 1996 Copyright: Clemens Ballarin Contributions, in particular on long division, by Jesus Aransay. *) theory UnivPoly imports Module RingHom begin section \Univariate Polynomials\ text \ Polynomials are formalised as modules with additional operations for extracting coefficients from polynomials and for obtaining monomials from coefficients and exponents (record \up_ring\). The carrier set is a set of bounded functions from Nat to the coefficient domain. Bounded means that these functions return zero above a certain bound (the degree). There is a chapter on the - formalisation of polynomials in the PhD thesis @{cite "Ballarin:1999"}, + formalisation of polynomials in the PhD thesis \<^cite>\"Ballarin:1999"\, which was implemented with axiomatic type classes. This was later ported to Locales. \ subsection \The Constructor for Univariate Polynomials\ text \ Functions with finite support. \ locale bound = fixes z :: 'a and n :: nat and f :: "nat => 'a" assumes bound: "!!m. n < m \ f m = z" declare bound.intro [intro!] and bound.bound [dest] lemma bound_below: assumes bound: "bound z m f" and nonzero: "f n \ z" shows "n \ m" proof (rule classical) assume "\ ?thesis" then have "m < n" by arith with bound have "f n = z" .. with nonzero show ?thesis by contradiction qed record ('a, 'p) up_ring = "('a, 'p) module" + monom :: "['a, nat] => 'p" coeff :: "['p, nat] => 'a" definition up :: "('a, 'm) ring_scheme => (nat => 'a) set" where "up R = {f. f \ UNIV \ carrier R \ (\n. bound \\<^bsub>R\<^esub> n f)}" definition UP :: "('a, 'm) ring_scheme => ('a, nat => 'a) up_ring" where "UP R = \ carrier = up R, mult = (\p\up R. \q\up R. \n. \\<^bsub>R\<^esub>i \ {..n}. p i \\<^bsub>R\<^esub> q (n-i)), one = (\i. if i=0 then \\<^bsub>R\<^esub> else \\<^bsub>R\<^esub>), zero = (\i. \\<^bsub>R\<^esub>), add = (\p\up R. \q\up R. \i. p i \\<^bsub>R\<^esub> q i), smult = (\a\carrier R. \p\up R. \i. a \\<^bsub>R\<^esub> p i), monom = (\a\carrier R. \n i. if i=n then a else \\<^bsub>R\<^esub>), coeff = (\p\up R. \n. p n)\" text \ Properties of the set of polynomials \<^term>\up\. \ lemma mem_upI [intro]: "[| \n. f n \ carrier R; \n. bound (zero R) n f |] ==> f \ up R" by (simp add: up_def Pi_def) lemma mem_upD [dest]: "f \ up R ==> f n \ carrier R" by (simp add: up_def Pi_def) context ring begin lemma bound_upD [dest]: "f \ up R \ \n. bound \ n f" by (simp add: up_def) lemma up_one_closed: "(\n. if n = 0 then \ else \) \ up R" using up_def by force lemma up_smult_closed: "[| a \ carrier R; p \ up R |] ==> (\i. a \ p i) \ up R" by force lemma up_add_closed: "[| p \ up R; q \ up R |] ==> (\i. p i \ q i) \ up R" proof fix n assume "p \ up R" and "q \ up R" then show "p n \ q n \ carrier R" by auto next assume UP: "p \ up R" "q \ up R" show "\n. bound \ n (\i. p i \ q i)" proof - from UP obtain n where boundn: "bound \ n p" by fast from UP obtain m where boundm: "bound \ m q" by fast have "bound \ (max n m) (\i. p i \ q i)" proof fix i assume "max n m < i" with boundn and boundm and UP show "p i \ q i = \" by fastforce qed then show ?thesis .. qed qed lemma up_a_inv_closed: "p \ up R ==> (\i. \ (p i)) \ up R" proof assume R: "p \ up R" then obtain n where "bound \ n p" by auto then have "bound \ n (\i. \ p i)" by (simp add: bound_def minus_equality) then show "\n. bound \ n (\i. \ p i)" by auto qed auto lemma up_minus_closed: "[| p \ up R; q \ up R |] ==> (\i. p i \ q i) \ up R" unfolding a_minus_def using mem_upD [of p R] mem_upD [of q R] up_add_closed up_a_inv_closed by auto lemma up_mult_closed: "[| p \ up R; q \ up R |] ==> (\n. \i \ {..n}. p i \ q (n-i)) \ up R" proof fix n assume "p \ up R" "q \ up R" then show "(\i \ {..n}. p i \ q (n-i)) \ carrier R" by (simp add: mem_upD funcsetI) next assume UP: "p \ up R" "q \ up R" show "\n. bound \ n (\n. \i \ {..n}. p i \ q (n-i))" proof - from UP obtain n where boundn: "bound \ n p" by fast from UP obtain m where boundm: "bound \ m q" by fast have "bound \ (n + m) (\n. \i \ {..n}. p i \ q (n - i))" proof fix k assume bound: "n + m < k" { fix i have "p i \ q (k-i) = \" proof (cases "n < i") case True with boundn have "p i = \" by auto moreover from UP have "q (k-i) \ carrier R" by auto ultimately show ?thesis by simp next case False with bound have "m < k-i" by arith with boundm have "q (k-i) = \" by auto moreover from UP have "p i \ carrier R" by auto ultimately show ?thesis by simp qed } then show "(\i \ {..k}. p i \ q (k-i)) = \" by (simp add: Pi_def) qed then show ?thesis by fast qed qed end subsection \Effect of Operations on Coefficients\ locale UP = fixes R (structure) and P (structure) defines P_def: "P == UP R" locale UP_ring = UP + R?: ring R locale UP_cring = UP + R?: cring R sublocale UP_cring < UP_ring by intro_locales [1] (rule P_def) locale UP_domain = UP + R?: "domain" R sublocale UP_domain < UP_cring by intro_locales [1] (rule P_def) context UP begin text \Temporarily declare @{thm P_def} as simp rule.\ declare P_def [simp] lemma up_eqI: assumes prem: "!!n. coeff P p n = coeff P q n" and R: "p \ carrier P" "q \ carrier P" shows "p = q" proof fix x from prem and R show "p x = q x" by (simp add: UP_def) qed lemma coeff_closed [simp]: "p \ carrier P ==> coeff P p n \ carrier R" by (auto simp add: UP_def) end context UP_ring begin (* Theorems generalised from commutative rings to rings by Jesus Aransay. *) lemma coeff_monom [simp]: "a \ carrier R ==> coeff P (monom P a m) n = (if m=n then a else \)" proof - assume R: "a \ carrier R" then have "(\n. if n = m then a else \) \ up R" using up_def by force with R show ?thesis by (simp add: UP_def) qed lemma coeff_zero [simp]: "coeff P \\<^bsub>P\<^esub> n = \" by (auto simp add: UP_def) lemma coeff_one [simp]: "coeff P \\<^bsub>P\<^esub> n = (if n=0 then \ else \)" using up_one_closed by (simp add: UP_def) lemma coeff_smult [simp]: "[| a \ carrier R; p \ carrier P |] ==> coeff P (a \\<^bsub>P\<^esub> p) n = a \ coeff P p n" by (simp add: UP_def up_smult_closed) lemma coeff_add [simp]: "[| p \ carrier P; q \ carrier P |] ==> coeff P (p \\<^bsub>P\<^esub> q) n = coeff P p n \ coeff P q n" by (simp add: UP_def up_add_closed) lemma coeff_mult [simp]: "[| p \ carrier P; q \ carrier P |] ==> coeff P (p \\<^bsub>P\<^esub> q) n = (\i \ {..n}. coeff P p i \ coeff P q (n-i))" by (simp add: UP_def up_mult_closed) end subsection \Polynomials Form a Ring.\ context UP_ring begin text \Operations are closed over \<^term>\P\.\ lemma UP_mult_closed [simp]: "[| p \ carrier P; q \ carrier P |] ==> p \\<^bsub>P\<^esub> q \ carrier P" by (simp add: UP_def up_mult_closed) lemma UP_one_closed [simp]: "\\<^bsub>P\<^esub> \ carrier P" by (simp add: UP_def up_one_closed) lemma UP_zero_closed [intro, simp]: "\\<^bsub>P\<^esub> \ carrier P" by (auto simp add: UP_def) lemma UP_a_closed [intro, simp]: "[| p \ carrier P; q \ carrier P |] ==> p \\<^bsub>P\<^esub> q \ carrier P" by (simp add: UP_def up_add_closed) lemma monom_closed [simp]: "a \ carrier R ==> monom P a n \ carrier P" by (auto simp add: UP_def up_def Pi_def) lemma UP_smult_closed [simp]: "[| a \ carrier R; p \ carrier P |] ==> a \\<^bsub>P\<^esub> p \ carrier P" by (simp add: UP_def up_smult_closed) end declare (in UP) P_def [simp del] text \Algebraic ring properties\ context UP_ring begin lemma UP_a_assoc: assumes R: "p \ carrier P" "q \ carrier P" "r \ carrier P" shows "(p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> r = p \\<^bsub>P\<^esub> (q \\<^bsub>P\<^esub> r)" by (rule up_eqI, simp add: a_assoc R, simp_all add: R) lemma UP_l_zero [simp]: assumes R: "p \ carrier P" shows "\\<^bsub>P\<^esub> \\<^bsub>P\<^esub> p = p" by (rule up_eqI, simp_all add: R) lemma UP_l_neg_ex: assumes R: "p \ carrier P" shows "\q \ carrier P. q \\<^bsub>P\<^esub> p = \\<^bsub>P\<^esub>" proof - let ?q = "\i. \ (p i)" from R have closed: "?q \ carrier P" by (simp add: UP_def P_def up_a_inv_closed) from R have coeff: "!!n. coeff P ?q n = \ (coeff P p n)" by (simp add: UP_def P_def up_a_inv_closed) show ?thesis proof show "?q \\<^bsub>P\<^esub> p = \\<^bsub>P\<^esub>" by (auto intro!: up_eqI simp add: R closed coeff R.l_neg) qed (rule closed) qed lemma UP_a_comm: assumes R: "p \ carrier P" "q \ carrier P" shows "p \\<^bsub>P\<^esub> q = q \\<^bsub>P\<^esub> p" by (rule up_eqI, simp add: a_comm R, simp_all add: R) lemma UP_m_assoc: assumes R: "p \ carrier P" "q \ carrier P" "r \ carrier P" shows "(p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> r = p \\<^bsub>P\<^esub> (q \\<^bsub>P\<^esub> r)" proof (rule up_eqI) fix n { fix k and a b c :: "nat=>'a" assume R: "a \ UNIV \ carrier R" "b \ UNIV \ carrier R" "c \ UNIV \ carrier R" then have "k <= n ==> (\j \ {..k}. (\i \ {..j}. a i \ b (j-i)) \ c (n-j)) = (\j \ {..k}. a j \ (\i \ {..k-j}. b i \ c (n-j-i)))" (is "_ \ ?eq k") proof (induct k) case 0 then show ?case by (simp add: Pi_def m_assoc) next case (Suc k) then have "k <= n" by arith from this R have "?eq k" by (rule Suc) with R show ?case by (simp cong: finsum_cong add: Suc_diff_le Pi_def l_distr r_distr m_assoc) (simp cong: finsum_cong add: Pi_def a_ac finsum_ldistr m_assoc) qed } with R show "coeff P ((p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> r) n = coeff P (p \\<^bsub>P\<^esub> (q \\<^bsub>P\<^esub> r)) n" by (simp add: Pi_def) qed (simp_all add: R) lemma UP_r_one [simp]: assumes R: "p \ carrier P" shows "p \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> = p" proof (rule up_eqI) fix n show "coeff P (p \\<^bsub>P\<^esub> \\<^bsub>P\<^esub>) n = coeff P p n" proof (cases n) case 0 { with R show ?thesis by simp } next case Suc { (*JE: in the locale UP_cring the proof was solved only with "by (simp del: finsum_Suc add: finsum_Suc2 Pi_def)", but I did not get it to work here*) fix nn assume Succ: "n = Suc nn" have "coeff P (p \\<^bsub>P\<^esub> \\<^bsub>P\<^esub>) (Suc nn) = coeff P p (Suc nn)" proof - have "coeff P (p \\<^bsub>P\<^esub> \\<^bsub>P\<^esub>) (Suc nn) = (\i\{..Suc nn}. coeff P p i \ (if Suc nn \ i then \ else \))" using R by simp also have "\ = coeff P p (Suc nn) \ (if Suc nn \ Suc nn then \ else \) \ (\i\{..nn}. coeff P p i \ (if Suc nn \ i then \ else \))" using finsum_Suc [of "(\i::nat. coeff P p i \ (if Suc nn \ i then \ else \))" "nn"] unfolding Pi_def using R by simp also have "\ = coeff P p (Suc nn) \ (if Suc nn \ Suc nn then \ else \)" proof - have "(\i\{..nn}. coeff P p i \ (if Suc nn \ i then \ else \)) = (\i\{..nn}. \)" using finsum_cong [of "{..nn}" "{..nn}" "(\i::nat. coeff P p i \ (if Suc nn \ i then \ else \))" "(\i::nat. \)"] using R unfolding Pi_def by simp also have "\ = \" by simp finally show ?thesis using r_zero R by simp qed also have "\ = coeff P p (Suc nn)" using R by simp finally show ?thesis by simp qed then show ?thesis using Succ by simp } qed qed (simp_all add: R) lemma UP_l_one [simp]: assumes R: "p \ carrier P" shows "\\<^bsub>P\<^esub> \\<^bsub>P\<^esub> p = p" proof (rule up_eqI) fix n show "coeff P (\\<^bsub>P\<^esub> \\<^bsub>P\<^esub> p) n = coeff P p n" proof (cases n) case 0 with R show ?thesis by simp next case Suc with R show ?thesis by (simp del: finsum_Suc add: finsum_Suc2 Pi_def) qed qed (simp_all add: R) lemma UP_l_distr: assumes R: "p \ carrier P" "q \ carrier P" "r \ carrier P" shows "(p \\<^bsub>P\<^esub> q) \\<^bsub>P\<^esub> r = (p \\<^bsub>P\<^esub> r) \\<^bsub>P\<^esub> (q \\<^bsub>P\<^esub> r)" by (rule up_eqI) (simp add: l_distr R Pi_def, simp_all add: R) lemma UP_r_distr: assumes R: "p \ carrier P" "q \ carrier P" "r \ carrier P" shows "r \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q) = (r \\<^bsub>P\<^esub> p) \\<^bsub>P\<^esub> (r \\<^bsub>P\<^esub> q)" by (rule up_eqI) (simp add: r_distr R Pi_def, simp_all add: R) theorem UP_ring: "ring P" by (auto intro!: ringI abelian_groupI monoidI UP_a_assoc) (auto intro: UP_a_comm UP_l_neg_ex UP_m_assoc UP_l_distr UP_r_distr) end subsection \Polynomials Form a Commutative Ring.\ context UP_cring begin lemma UP_m_comm: assumes R1: "p \ carrier P" and R2: "q \ carrier P" shows "p \\<^bsub>P\<^esub> q = q \\<^bsub>P\<^esub> p" proof (rule up_eqI) fix n { fix k and a b :: "nat=>'a" assume R: "a \ UNIV \ carrier R" "b \ UNIV \ carrier R" then have "k <= n ==> (\i \ {..k}. a i \ b (n-i)) = (\i \ {..k}. a (k-i) \ b (i+n-k))" (is "_ \ ?eq k") proof (induct k) case 0 then show ?case by (simp add: Pi_def) next case (Suc k) then show ?case by (subst (2) finsum_Suc2) (simp add: Pi_def a_comm)+ qed } note l = this from R1 R2 show "coeff P (p \\<^bsub>P\<^esub> q) n = coeff P (q \\<^bsub>P\<^esub> p) n" unfolding coeff_mult [OF R1 R2, of n] unfolding coeff_mult [OF R2 R1, of n] using l [of "(\i. coeff P p i)" "(\i. coeff P q i)" "n"] by (simp add: Pi_def m_comm) qed (simp_all add: R1 R2) subsection \Polynomials over a commutative ring for a commutative ring\ theorem UP_cring: "cring P" using UP_ring unfolding cring_def by (auto intro!: comm_monoidI UP_m_assoc UP_m_comm) end context UP_ring begin lemma UP_a_inv_closed [intro, simp]: "p \ carrier P ==> \\<^bsub>P\<^esub> p \ carrier P" by (rule abelian_group.a_inv_closed [OF ring.is_abelian_group [OF UP_ring]]) lemma coeff_a_inv [simp]: assumes R: "p \ carrier P" shows "coeff P (\\<^bsub>P\<^esub> p) n = \ (coeff P p n)" proof - from R coeff_closed UP_a_inv_closed have "coeff P (\\<^bsub>P\<^esub> p) n = \ coeff P p n \ (coeff P p n \ coeff P (\\<^bsub>P\<^esub> p) n)" by algebra also from R have "... = \ (coeff P p n)" by (simp del: coeff_add add: coeff_add [THEN sym] abelian_group.r_neg [OF ring.is_abelian_group [OF UP_ring]]) finally show ?thesis . qed end sublocale UP_ring < P?: ring P using UP_ring . sublocale UP_cring < P?: cring P using UP_cring . subsection \Polynomials Form an Algebra\ context UP_ring begin lemma UP_smult_l_distr: "[| a \ carrier R; b \ carrier R; p \ carrier P |] ==> (a \ b) \\<^bsub>P\<^esub> p = a \\<^bsub>P\<^esub> p \\<^bsub>P\<^esub> b \\<^bsub>P\<^esub> p" by (rule up_eqI) (simp_all add: R.l_distr) lemma UP_smult_r_distr: "[| a \ carrier R; p \ carrier P; q \ carrier P |] ==> a \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q) = a \\<^bsub>P\<^esub> p \\<^bsub>P\<^esub> a \\<^bsub>P\<^esub> q" by (rule up_eqI) (simp_all add: R.r_distr) lemma UP_smult_assoc1: "[| a \ carrier R; b \ carrier R; p \ carrier P |] ==> (a \ b) \\<^bsub>P\<^esub> p = a \\<^bsub>P\<^esub> (b \\<^bsub>P\<^esub> p)" by (rule up_eqI) (simp_all add: R.m_assoc) lemma UP_smult_zero [simp]: "p \ carrier P ==> \ \\<^bsub>P\<^esub> p = \\<^bsub>P\<^esub>" by (rule up_eqI) simp_all lemma UP_smult_one [simp]: "p \ carrier P ==> \ \\<^bsub>P\<^esub> p = p" by (rule up_eqI) simp_all lemma UP_smult_assoc2: "[| a \ carrier R; p \ carrier P; q \ carrier P |] ==> (a \\<^bsub>P\<^esub> p) \\<^bsub>P\<^esub> q = a \\<^bsub>P\<^esub> (p \\<^bsub>P\<^esub> q)" by (rule up_eqI) (simp_all add: R.finsum_rdistr R.m_assoc Pi_def) end text \ Interpretation of lemmas from \<^term>\algebra\. \ lemma (in UP_cring) UP_algebra: "algebra R P" by (auto intro!: algebraI R.cring_axioms UP_cring UP_smult_l_distr UP_smult_r_distr UP_smult_assoc1 UP_smult_assoc2) sublocale UP_cring < algebra R P using UP_algebra . subsection \Further Lemmas Involving Monomials\ context UP_ring begin lemma monom_zero [simp]: "monom P \ n = \\<^bsub>P\<^esub>" by (simp add: UP_def P_def) lemma monom_mult_is_smult: assumes R: "a \ carrier R" "p \ carrier P" shows "monom P a 0 \\<^bsub>P\<^esub> p = a \\<^bsub>P\<^esub> p" proof (rule up_eqI) fix n show "coeff P (monom P a 0 \\<^bsub>P\<^esub> p) n = coeff P (a \\<^bsub>P\<^esub> p) n" proof (cases n) case 0 with R show ?thesis by simp next case Suc with R show ?thesis using R.finsum_Suc2 by (simp del: R.finsum_Suc add: Pi_def) qed qed (simp_all add: R) lemma monom_one [simp]: "monom P \ 0 = \\<^bsub>P\<^esub>" by (rule up_eqI) simp_all lemma monom_add [simp]: "[| a \ carrier R; b \ carrier R |] ==> monom P (a \ b) n = monom P a n \\<^bsub>P\<^esub> monom P b n" by (rule up_eqI) simp_all lemma monom_one_Suc: "monom P \ (Suc n) = monom P \ n \\<^bsub>P\<^esub> monom P \ 1" proof (rule up_eqI) fix k show "coeff P (monom P \ (Suc n)) k = coeff P (monom P \ n \\<^bsub>P\<^esub> monom P \ 1) k" proof (cases "k = Suc n") case True show ?thesis proof - fix m from True have less_add_diff: "!!i. [| n < i; i <= n + m |] ==> n + m - i < m" by arith from True have "coeff P (monom P \ (Suc n)) k = \" by simp also from True have "... = (\i \ {.. {n}. coeff P (monom P \ n) i \ coeff P (monom P \ 1) (k - i))" by (simp cong: R.finsum_cong add: Pi_def) also have "... = (\i \ {..n}. coeff P (monom P \ n) i \ coeff P (monom P \ 1) (k - i))" by (simp only: ivl_disj_un_singleton) also from True have "... = (\i \ {..n} \ {n<..k}. coeff P (monom P \ n) i \ coeff P (monom P \ 1) (k - i))" by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one order_less_imp_not_eq Pi_def) also from True have "... = coeff P (monom P \ n \\<^bsub>P\<^esub> monom P \ 1) k" by (simp add: ivl_disj_un_one) finally show ?thesis . qed next case False note neq = False let ?s = "\i. (if n = i then \ else \) \ (if Suc 0 = k - i then \ else \)" from neq have "coeff P (monom P \ (Suc n)) k = \" by simp also have "... = (\i \ {..k}. ?s i)" proof - have f1: "(\i \ {.." by (simp cong: R.finsum_cong add: Pi_def) from neq have f2: "(\i \ {n}. ?s i) = \" by (simp cong: R.finsum_cong add: Pi_def) arith have f3: "n < k ==> (\i \ {n<..k}. ?s i) = \" by (simp cong: R.finsum_cong add: order_less_imp_not_eq Pi_def) show ?thesis proof (cases "k < n") case True then show ?thesis by (simp cong: R.finsum_cong add: Pi_def) next case False then have n_le_k: "n <= k" by arith show ?thesis proof (cases "n = k") case True then have "\ = (\i \ {.. {n}. ?s i)" by (simp cong: R.finsum_cong add: Pi_def) also from True have "... = (\i \ {..k}. ?s i)" by (simp only: ivl_disj_un_singleton) finally show ?thesis . next case False with n_le_k have n_less_k: "n < k" by arith with neq have "\ = (\i \ {.. {n}. ?s i)" by (simp add: R.finsum_Un_disjoint f1 f2 Pi_def del: Un_insert_right) also have "... = (\i \ {..n}. ?s i)" by (simp only: ivl_disj_un_singleton) also from n_less_k neq have "... = (\i \ {..n} \ {n<..k}. ?s i)" by (simp add: R.finsum_Un_disjoint f3 ivl_disj_int_one Pi_def) also from n_less_k have "... = (\i \ {..k}. ?s i)" by (simp only: ivl_disj_un_one) finally show ?thesis . qed qed qed also have "... = coeff P (monom P \ n \\<^bsub>P\<^esub> monom P \ 1) k" by simp finally show ?thesis . qed qed (simp_all) lemma monom_one_Suc2: "monom P \ (Suc n) = monom P \ 1 \\<^bsub>P\<^esub> monom P \ n" proof (induct n) case 0 show ?case by simp next case Suc { fix k:: nat assume hypo: "monom P \ (Suc k) = monom P \ 1 \\<^bsub>P\<^esub> monom P \ k" then show "monom P \ (Suc (Suc k)) = monom P \ 1 \\<^bsub>P\<^esub> monom P \ (Suc k)" proof - have lhs: "monom P \ (Suc (Suc k)) = monom P \ 1 \\<^bsub>P\<^esub> monom P \ k \\<^bsub>P\<^esub> monom P \ 1" unfolding monom_one_Suc [of "Suc k"] unfolding hypo .. note cl = monom_closed [OF R.one_closed, of 1] note clk = monom_closed [OF R.one_closed, of k] have rhs: "monom P \ 1 \\<^bsub>P\<^esub> monom P \ (Suc k) = monom P \ 1 \\<^bsub>P\<^esub> monom P \ k \\<^bsub>P\<^esub> monom P \ 1" unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc [OF cl clk cl]] .. from lhs rhs show ?thesis by simp qed } qed text\The following corollary follows from lemmas @{thm "monom_one_Suc"} and @{thm "monom_one_Suc2"}, and is trivial in \<^term>\UP_cring\\ corollary monom_one_comm: shows "monom P \ k \\<^bsub>P\<^esub> monom P \ 1 = monom P \ 1 \\<^bsub>P\<^esub> monom P \ k" unfolding monom_one_Suc [symmetric] monom_one_Suc2 [symmetric] .. lemma monom_mult_smult: "[| a \ carrier R; b \ carrier R |] ==> monom P (a \ b) n = a \\<^bsub>P\<^esub> monom P b n" by (rule up_eqI) simp_all lemma monom_one_mult: "monom P \ (n + m) = monom P \ n \\<^bsub>P\<^esub> monom P \ m" proof (induct n) case 0 show ?case by simp next case Suc then show ?case unfolding add_Suc unfolding monom_one_Suc unfolding Suc.hyps using m_assoc monom_one_comm [of m] by simp qed lemma monom_one_mult_comm: "monom P \ n \\<^bsub>P\<^esub> monom P \ m = monom P \ m \\<^bsub>P\<^esub> monom P \ n" unfolding monom_one_mult [symmetric] by (rule up_eqI) simp_all lemma monom_mult [simp]: assumes a_in_R: "a \ carrier R" and b_in_R: "b \ carrier R" shows "monom P (a \ b) (n + m) = monom P a n \\<^bsub>P\<^esub> monom P b m" proof (rule up_eqI) fix k show "coeff P (monom P (a \ b) (n + m)) k = coeff P (monom P a n \\<^bsub>P\<^esub> monom P b m) k" proof (cases "n + m = k") case True { show ?thesis unfolding True [symmetric] coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"] coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m] using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\i. (if n = i then a else \) \ (if m = n + m - i then b else \))" "(\i. if n = i then a \ b else \)"] a_in_R b_in_R unfolding simp_implies_def using R.finsum_singleton [of n "{.. n + m}" "(\i. a \ b)"] unfolding Pi_def by auto } next case False { show ?thesis unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False) unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k] unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False using R.finsum_cong [of "{..k}" "{..k}" "(\i. (if n = i then a else \) \ (if m = k - i then b else \))" "(\i. \)"] unfolding Pi_def simp_implies_def using a_in_R b_in_R by force } qed qed (simp_all add: a_in_R b_in_R) lemma monom_a_inv [simp]: "a \ carrier R ==> monom P (\ a) n = \\<^bsub>P\<^esub> monom P a n" by (rule up_eqI) auto lemma monom_inj: "inj_on (\a. monom P a n) (carrier R)" proof (rule inj_onI) fix x y assume R: "x \ carrier R" "y \ carrier R" and eq: "monom P x n = monom P y n" then have "coeff P (monom P x n) n = coeff P (monom P y n) n" by simp with R show "x = y" by simp qed end subsection \The Degree Function\ definition deg :: "[('a, 'm) ring_scheme, nat => 'a] => nat" where "deg R p = (LEAST n. bound \\<^bsub>R\<^esub> n (coeff (UP R) p))" context UP_ring begin lemma deg_aboveI: "[| (!!m. n < m ==> coeff P p m = \); p \ carrier P |] ==> deg R p <= n" by (unfold deg_def P_def) (fast intro: Least_le) (* lemma coeff_bound_ex: "EX n. bound n (coeff p)" proof - have "(\n. coeff p n) : UP" by (simp add: coeff_def Rep_UP) then obtain n where "bound n (coeff p)" by (unfold UP_def) fast then show ?thesis .. qed lemma bound_coeff_obtain: assumes prem: "(!!n. bound n (coeff p) ==> P)" shows "P" proof - have "(\n. coeff p n) : UP" by (simp add: coeff_def Rep_UP) then obtain n where "bound n (coeff p)" by (unfold UP_def) fast with prem show P . qed *) lemma deg_aboveD: assumes "deg R p < m" and "p \ carrier P" shows "coeff P p m = \" proof - from \p \ carrier P\ obtain n where "bound \ n (coeff P p)" by (auto simp add: UP_def P_def) then have "bound \ (deg R p) (coeff P p)" by (auto simp: deg_def P_def dest: LeastI) from this and \deg R p < m\ show ?thesis .. qed lemma deg_belowI: assumes non_zero: "n \ 0 \ coeff P p n \ \" and R: "p \ carrier P" shows "n \ deg R p" \ \Logically, this is a slightly stronger version of @{thm [source] deg_aboveD}\ proof (cases "n=0") case True then show ?thesis by simp next case False then have "coeff P p n \ \" by (rule non_zero) then have "\ deg R p < n" by (fast dest: deg_aboveD intro: R) then show ?thesis by arith qed lemma lcoeff_nonzero_deg: assumes deg: "deg R p \ 0" and R: "p \ carrier P" shows "coeff P p (deg R p) \ \" proof - from R obtain m where "deg R p \ m" and m_coeff: "coeff P p m \ \" proof - have minus: "\(n::nat) m. n \ 0 \ (n - Suc 0 < m) = (n \ m)" by arith from deg have "deg R p - 1 < (LEAST n. bound \ n (coeff P p))" by (unfold deg_def P_def) simp then have "\ bound \ (deg R p - 1) (coeff P p)" by (rule not_less_Least) then have "\m. deg R p - 1 < m \ coeff P p m \ \" by (unfold bound_def) fast then have "\m. deg R p \ m \ coeff P p m \ \" by (simp add: deg minus) then show ?thesis by (auto intro: that) qed with deg_belowI R have "deg R p = m" by fastforce with m_coeff show ?thesis by simp qed lemma lcoeff_nonzero_nonzero: assumes deg: "deg R p = 0" and nonzero: "p \ \\<^bsub>P\<^esub>" and R: "p \ carrier P" shows "coeff P p 0 \ \" proof - have "\m. coeff P p m \ \" proof (rule classical) assume "\ ?thesis" with R have "p = \\<^bsub>P\<^esub>" by (auto intro: up_eqI) with nonzero show ?thesis by contradiction qed then obtain m where coeff: "coeff P p m \ \" .. from this and R have "m \ deg R p" by (rule deg_belowI) then have "m = 0" by (simp add: deg) with coeff show ?thesis by simp qed lemma lcoeff_nonzero: assumes neq: "p \ \\<^bsub>P\<^esub>" and R: "p \ carrier P" shows "coeff P p (deg R p) \ \" proof (cases "deg R p = 0") case True with neq R show ?thesis by (simp add: lcoeff_nonzero_nonzero) next case False with neq R show ?thesis by (simp add: lcoeff_nonzero_deg) qed lemma deg_eqI: "[| \m. n < m \ coeff P p m = \; \n. n \ 0 \ coeff P p n \ \; p \ carrier P |] ==> deg R p = n" by (fast intro: le_antisym deg_aboveI deg_belowI) text \Degree and polynomial operations\ lemma deg_add [simp]: "p \ carrier P \ q \ carrier P \ deg R (p \\<^bsub>P\<^esub> q) \ max (deg R p) (deg R q)" by(rule deg_aboveI)(simp_all add: deg_aboveD) lemma deg_monom_le: "a \ carrier R \ deg R (monom P a n) \ n" by (intro deg_aboveI) simp_all lemma deg_monom [simp]: "[| a \ \; a \ carrier R |] ==> deg R (monom P a n) = n" by (fastforce intro: le_antisym deg_aboveI deg_belowI) lemma deg_const [simp]: assumes R: "a \ carrier R" shows "deg R (monom P a 0) = 0" proof (rule le_antisym) show "deg R (monom P a 0) \ 0" by (rule deg_aboveI) (simp_all add: R) next show "0 \ deg R (monom P a 0)" by (rule deg_belowI) (simp_all add: R) qed lemma deg_zero [simp]: "deg R \\<^bsub>P\<^esub> = 0" proof (rule le_antisym) show "deg R \\<^bsub>P\<^esub> \ 0" by (rule deg_aboveI) simp_all next show "0 \ deg R \\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all qed lemma deg_one [simp]: "deg R \\<^bsub>P\<^esub> = 0" proof (rule le_antisym) show "deg R \\<^bsub>P\<^esub> \ 0" by (rule deg_aboveI) simp_all next show "0 \ deg R \\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all qed lemma deg_uminus [simp]: assumes R: "p \ carrier P" shows "deg R (\\<^bsub>P\<^esub> p) = deg R p" proof (rule le_antisym) show "deg R (\\<^bsub>P\<^esub> p) \ deg R p" by (simp add: deg_aboveI deg_aboveD R) next show "deg R p \ deg R (\\<^bsub>P\<^esub> p)" by (simp add: deg_belowI lcoeff_nonzero_deg inj_on_eq_iff [OF R.a_inv_inj, of _ "\", simplified] R) qed text\The following lemma is later \emph{overwritten} by the most specific one for domains, \deg_smult\.\ lemma deg_smult_ring [simp]: "[| a \ carrier R; p \ carrier P |] ==> deg R (a \\<^bsub>P\<^esub> p) \ (if a = \ then 0 else deg R p)" by (cases "a = \") (simp add: deg_aboveI deg_aboveD)+ end context UP_domain begin lemma deg_smult [simp]: assumes R: "a \ carrier R" "p \ carrier P" shows "deg R (a \\<^bsub>P\<^esub> p) = (if a = \ then 0 else deg R p)" proof (rule le_antisym) show "deg R (a \\<^bsub>P\<^esub> p) \ (if a = \ then 0 else deg R p)" using R by (rule deg_smult_ring) next show "(if a = \ then 0 else deg R p) \ deg R (a \\<^bsub>P\<^esub> p)" proof (cases "a = \") qed (simp, simp add: deg_belowI lcoeff_nonzero_deg integral_iff R) qed end context UP_ring begin lemma deg_mult_ring: assumes R: "p \ carrier P" "q \ carrier P" shows "deg R (p \\<^bsub>P\<^esub> q) \ deg R p + deg R q" proof (rule deg_aboveI) fix m assume boundm: "deg R p + deg R q < m" { fix k i assume boundk: "deg R p + deg R q < k" then have "coeff P p i \ coeff P q (k - i) = \" proof (cases "deg R p < i") case True then show ?thesis by (simp add: deg_aboveD R) next case False with boundk have "deg R q < k - i" by arith then show ?thesis by (simp add: deg_aboveD R) qed } with boundm R show "coeff P (p \\<^bsub>P\<^esub> q) m = \" by simp qed (simp add: R) end context UP_domain begin lemma deg_mult [simp]: "[| p \ \\<^bsub>P\<^esub>; q \ \\<^bsub>P\<^esub>; p \ carrier P; q \ carrier P |] ==> deg R (p \\<^bsub>P\<^esub> q) = deg R p + deg R q" proof (rule le_antisym) assume "p \ carrier P" " q \ carrier P" then show "deg R (p \\<^bsub>P\<^esub> q) \ deg R p + deg R q" by (rule deg_mult_ring) next let ?s = "(\i. coeff P p i \ coeff P q (deg R p + deg R q - i))" assume R: "p \ carrier P" "q \ carrier P" and nz: "p \ \\<^bsub>P\<^esub>" "q \ \\<^bsub>P\<^esub>" have less_add_diff: "!!(k::nat) n m. k < n ==> m < n + m - k" by arith show "deg R p + deg R q \ deg R (p \\<^bsub>P\<^esub> q)" proof (rule deg_belowI, simp add: R) have "(\i \ {.. deg R p + deg R q}. ?s i) = (\i \ {..< deg R p} \ {deg R p .. deg R p + deg R q}. ?s i)" by (simp only: ivl_disj_un_one) also have "... = (\i \ {deg R p .. deg R p + deg R q}. ?s i)" by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one deg_aboveD less_add_diff R Pi_def) also have "...= (\i \ {deg R p} \ {deg R p <.. deg R p + deg R q}. ?s i)" by (simp only: ivl_disj_un_singleton) also have "... = coeff P p (deg R p) \ coeff P q (deg R q)" by (simp cong: R.finsum_cong add: deg_aboveD R Pi_def) finally have "(\i \ {.. deg R p + deg R q}. ?s i) = coeff P p (deg R p) \ coeff P q (deg R q)" . with nz show "(\i \ {.. deg R p + deg R q}. ?s i) \ \" by (simp add: integral_iff lcoeff_nonzero R) qed (simp add: R) qed end text\The following lemmas also can be lifted to \<^term>\UP_ring\.\ context UP_ring begin lemma coeff_finsum: assumes fin: "finite A" shows "p \ A \ carrier P ==> coeff P (finsum P p A) k = (\i \ A. coeff P (p i) k)" using fin by induct (auto simp: Pi_def) lemma up_repr: assumes R: "p \ carrier P" shows "(\\<^bsub>P\<^esub> i \ {..deg R p}. monom P (coeff P p i) i) = p" proof (rule up_eqI) let ?s = "(\i. monom P (coeff P p i) i)" fix k from R have RR: "!!i. (if i = k then coeff P p i else \) \ carrier R" by simp show "coeff P (\\<^bsub>P\<^esub> i \ {..deg R p}. ?s i) k = coeff P p k" proof (cases "k \ deg R p") case True hence "coeff P (\\<^bsub>P\<^esub> i \ {..deg R p}. ?s i) k = coeff P (\\<^bsub>P\<^esub> i \ {..k} \ {k<..deg R p}. ?s i) k" by (simp only: ivl_disj_un_one) also from True have "... = coeff P (\\<^bsub>P\<^esub> i \ {..k}. ?s i) k" by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one order_less_imp_not_eq2 coeff_finsum R RR Pi_def) also have "... = coeff P (\\<^bsub>P\<^esub> i \ {.. {k}. ?s i) k" by (simp only: ivl_disj_un_singleton) also have "... = coeff P p k" by (simp cong: R.finsum_cong add: coeff_finsum deg_aboveD R RR Pi_def) finally show ?thesis . next case False hence "coeff P (\\<^bsub>P\<^esub> i \ {..deg R p}. ?s i) k = coeff P (\\<^bsub>P\<^esub> i \ {.. {deg R p}. ?s i) k" by (simp only: ivl_disj_un_singleton) also from False have "... = coeff P p k" by (simp cong: R.finsum_cong add: coeff_finsum deg_aboveD R Pi_def) finally show ?thesis . qed qed (simp_all add: R Pi_def) lemma up_repr_le: "[| deg R p <= n; p \ carrier P |] ==> (\\<^bsub>P\<^esub> i \ {..n}. monom P (coeff P p i) i) = p" proof - let ?s = "(\i. monom P (coeff P p i) i)" assume R: "p \ carrier P" and "deg R p <= n" then have "finsum P ?s {..n} = finsum P ?s ({..deg R p} \ {deg R p<..n})" by (simp only: ivl_disj_un_one) also have "... = finsum P ?s {..deg R p}" by (simp cong: P.finsum_cong add: P.finsum_Un_disjoint ivl_disj_int_one deg_aboveD R Pi_def) also have "... = p" using R by (rule up_repr) finally show ?thesis . qed end subsection \Polynomials over Integral Domains\ lemma domainI: assumes cring: "cring R" and one_not_zero: "one R \ zero R" and integral: "\a b. [| mult R a b = zero R; a \ carrier R; b \ carrier R |] ==> a = zero R \ b = zero R" shows "domain R" by (auto intro!: domain.intro domain_axioms.intro cring.axioms assms del: disjCI) context UP_domain begin lemma UP_one_not_zero: "\\<^bsub>P\<^esub> \ \\<^bsub>P\<^esub>" proof assume "\\<^bsub>P\<^esub> = \\<^bsub>P\<^esub>" hence "coeff P \\<^bsub>P\<^esub> 0 = (coeff P \\<^bsub>P\<^esub> 0)" by simp hence "\ = \" by simp with R.one_not_zero show "False" by contradiction qed lemma UP_integral: "[| p \\<^bsub>P\<^esub> q = \\<^bsub>P\<^esub>; p \ carrier P; q \ carrier P |] ==> p = \\<^bsub>P\<^esub> \ q = \\<^bsub>P\<^esub>" proof - fix p q assume pq: "p \\<^bsub>P\<^esub> q = \\<^bsub>P\<^esub>" and R: "p \ carrier P" "q \ carrier P" show "p = \\<^bsub>P\<^esub> \ q = \\<^bsub>P\<^esub>" proof (rule classical) assume c: "\ (p = \\<^bsub>P\<^esub> \ q = \\<^bsub>P\<^esub>)" with R have "deg R p + deg R q = deg R (p \\<^bsub>P\<^esub> q)" by simp also from pq have "... = 0" by simp finally have "deg R p + deg R q = 0" . then have f1: "deg R p = 0 \ deg R q = 0" by simp from f1 R have "p = (\\<^bsub>P\<^esub> i \ {..0}. monom P (coeff P p i) i)" by (simp only: up_repr_le) also from R have "... = monom P (coeff P p 0) 0" by simp finally have p: "p = monom P (coeff P p 0) 0" . from f1 R have "q = (\\<^bsub>P\<^esub> i \ {..0}. monom P (coeff P q i) i)" by (simp only: up_repr_le) also from R have "... = monom P (coeff P q 0) 0" by simp finally have q: "q = monom P (coeff P q 0) 0" . from R have "coeff P p 0 \ coeff P q 0 = coeff P (p \\<^bsub>P\<^esub> q) 0" by simp also from pq have "... = \" by simp finally have "coeff P p 0 \ coeff P q 0 = \" . with R have "coeff P p 0 = \ \ coeff P q 0 = \" by (simp add: R.integral_iff) with p q show "p = \\<^bsub>P\<^esub> \ q = \\<^bsub>P\<^esub>" by fastforce qed qed theorem UP_domain: "domain P" by (auto intro!: domainI UP_cring UP_one_not_zero UP_integral del: disjCI) end text \ Interpretation of theorems from \<^term>\domain\. \ sublocale UP_domain < "domain" P by intro_locales (rule domain.axioms UP_domain)+ subsection \The Evaluation Homomorphism and Universal Property\ (* alternative congruence rule (possibly more efficient) lemma (in abelian_monoid) finsum_cong2: "[| !!i. i \ A ==> f i \ carrier G = True; A = B; !!i. i \ B ==> f i = g i |] ==> finsum G f A = finsum G g B" sorry*) lemma (in abelian_monoid) boundD_carrier: "[| bound \ n f; n < m |] ==> f m \ carrier G" by auto context ring begin theorem diagonal_sum: "[| f \ {..n + m::nat} \ carrier R; g \ {..n + m} \ carrier R |] ==> (\k \ {..n + m}. \i \ {..k}. f i \ g (k - i)) = (\k \ {..n + m}. \i \ {..n + m - k}. f k \ g i)" proof - assume Rf: "f \ {..n + m} \ carrier R" and Rg: "g \ {..n + m} \ carrier R" { fix j have "j <= n + m ==> (\k \ {..j}. \i \ {..k}. f i \ g (k - i)) = (\k \ {..j}. \i \ {..j - k}. f k \ g i)" proof (induct j) case 0 from Rf Rg show ?case by (simp add: Pi_def) next case (Suc j) have R6: "!!i k. [| k <= j; i <= Suc j - k |] ==> g i \ carrier R" using Suc by (auto intro!: funcset_mem [OF Rg]) have R8: "!!i k. [| k <= Suc j; i <= k |] ==> g (k - i) \ carrier R" using Suc by (auto intro!: funcset_mem [OF Rg]) have R9: "!!i k. [| k <= Suc j |] ==> f k \ carrier R" using Suc by (auto intro!: funcset_mem [OF Rf]) have R10: "!!i k. [| k <= Suc j; i <= Suc j - k |] ==> g i \ carrier R" using Suc by (auto intro!: funcset_mem [OF Rg]) have R11: "g 0 \ carrier R" using Suc by (auto intro!: funcset_mem [OF Rg]) from Suc show ?case by (simp cong: finsum_cong add: Suc_diff_le a_ac Pi_def R6 R8 R9 R10 R11) qed } then show ?thesis by fast qed theorem cauchy_product: assumes bf: "bound \ n f" and bg: "bound \ m g" and Rf: "f \ {..n} \ carrier R" and Rg: "g \ {..m} \ carrier R" shows "(\k \ {..n + m}. \i \ {..k}. f i \ g (k - i)) = (\i \ {..n}. f i) \ (\i \ {..m}. g i)" (* State reverse direction? *) proof - have f: "!!x. f x \ carrier R" proof - fix x show "f x \ carrier R" using Rf bf boundD_carrier by (cases "x <= n") (auto simp: Pi_def) qed have g: "!!x. g x \ carrier R" proof - fix x show "g x \ carrier R" using Rg bg boundD_carrier by (cases "x <= m") (auto simp: Pi_def) qed from f g have "(\k \ {..n + m}. \i \ {..k}. f i \ g (k - i)) = (\k \ {..n + m}. \i \ {..n + m - k}. f k \ g i)" by (simp add: diagonal_sum Pi_def) also have "... = (\k \ {..n} \ {n<..n + m}. \i \ {..n + m - k}. f k \ g i)" by (simp only: ivl_disj_un_one) also from f g have "... = (\k \ {..n}. \i \ {..n + m - k}. f k \ g i)" by (simp cong: finsum_cong add: bound.bound [OF bf] finsum_Un_disjoint ivl_disj_int_one Pi_def) also from f g have "... = (\k \ {..n}. \i \ {..m} \ {m<..n + m - k}. f k \ g i)" by (simp cong: finsum_cong add: ivl_disj_un_one le_add_diff Pi_def) also from f g have "... = (\k \ {..n}. \i \ {..m}. f k \ g i)" by (simp cong: finsum_cong add: bound.bound [OF bg] finsum_Un_disjoint ivl_disj_int_one Pi_def) also from f g have "... = (\i \ {..n}. f i) \ (\i \ {..m}. g i)" by (simp add: finsum_ldistr diagonal_sum Pi_def, simp cong: finsum_cong add: finsum_rdistr Pi_def) finally show ?thesis . qed end lemma (in UP_ring) const_ring_hom: "(\a. monom P a 0) \ ring_hom R P" by (auto intro!: ring_hom_memI intro: up_eqI simp: monom_mult_is_smult) definition eval :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme, 'a => 'b, 'b, nat => 'a] => 'b" where "eval R S phi s = (\p \ carrier (UP R). \\<^bsub>S\<^esub>i \ {..deg R p}. phi (coeff (UP R) p i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" context UP begin lemma eval_on_carrier: fixes S (structure) shows "p \ carrier P ==> eval R S phi s p = (\\<^bsub>S\<^esub> i \ {..deg R p}. phi (coeff P p i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (unfold eval_def, fold P_def) simp lemma eval_extensional: "eval R S phi p \ extensional (carrier P)" by (unfold eval_def, fold P_def) simp end text \The universal property of the polynomial ring\ locale UP_pre_univ_prop = ring_hom_cring + UP_cring locale UP_univ_prop = UP_pre_univ_prop + fixes s and Eval assumes indet_img_carrier [simp, intro]: "s \ carrier S" defines Eval_def: "Eval == eval R S h s" text\JE: I have moved the following lemma from Ring.thy and lifted then to the locale \<^term>\ring_hom_ring\ from \<^term>\ring_hom_cring\.\ text\JE: I was considering using it in \eval_ring_hom\, but that property does not hold for non commutative rings, so maybe it is not that necessary.\ lemma (in ring_hom_ring) hom_finsum [simp]: "f \ A \ carrier R \ h (finsum R f A) = finsum S (h \ f) A" by (induct A rule: infinite_finite_induct, auto simp: Pi_def) context UP_pre_univ_prop begin theorem eval_ring_hom: assumes S: "s \ carrier S" shows "eval R S h s \ ring_hom P S" proof (rule ring_hom_memI) fix p assume R: "p \ carrier P" then show "eval R S h s p \ carrier S" by (simp only: eval_on_carrier) (simp add: S Pi_def) next fix p q assume R: "p \ carrier P" "q \ carrier P" then show "eval R S h s (p \\<^bsub>P\<^esub> q) = eval R S h s p \\<^bsub>S\<^esub> eval R S h s q" proof (simp only: eval_on_carrier P.a_closed) from S R have "(\\<^bsub>S \<^esub>i\{..deg R (p \\<^bsub>P\<^esub> q)}. h (coeff P (p \\<^bsub>P\<^esub> q) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) = (\\<^bsub>S \<^esub>i\{..deg R (p \\<^bsub>P\<^esub> q)} \ {deg R (p \\<^bsub>P\<^esub> q)<..max (deg R p) (deg R q)}. h (coeff P (p \\<^bsub>P\<^esub> q) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp cong: S.finsum_cong add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def del: coeff_add) also from R have "... = (\\<^bsub>S\<^esub> i \ {..max (deg R p) (deg R q)}. h (coeff P (p \\<^bsub>P\<^esub> q) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp add: ivl_disj_un_one) also from R S have "... = (\\<^bsub>S\<^esub>i\{..max (deg R p) (deg R q)}. h (coeff P p i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \\<^bsub>S\<^esub> (\\<^bsub>S\<^esub>i\{..max (deg R p) (deg R q)}. h (coeff P q i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp cong: S.finsum_cong add: S.l_distr deg_aboveD ivl_disj_int_one Pi_def) also have "... = (\\<^bsub>S\<^esub> i \ {..deg R p} \ {deg R p<..max (deg R p) (deg R q)}. h (coeff P p i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \\<^bsub>S\<^esub> (\\<^bsub>S\<^esub> i \ {..deg R q} \ {deg R q<..max (deg R p) (deg R q)}. h (coeff P q i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp only: ivl_disj_un_one max.cobounded1 max.cobounded2) also from R S have "... = (\\<^bsub>S\<^esub> i \ {..deg R p}. h (coeff P p i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \\<^bsub>S\<^esub> (\\<^bsub>S\<^esub> i \ {..deg R q}. h (coeff P q i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp cong: S.finsum_cong add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def) finally show "(\\<^bsub>S\<^esub>i \ {..deg R (p \\<^bsub>P\<^esub> q)}. h (coeff P (p \\<^bsub>P\<^esub> q) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) = (\\<^bsub>S\<^esub>i \ {..deg R p}. h (coeff P p i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \\<^bsub>S\<^esub> (\\<^bsub>S\<^esub>i \ {..deg R q}. h (coeff P q i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" . qed next show "eval R S h s \\<^bsub>P\<^esub> = \\<^bsub>S\<^esub>" by (simp only: eval_on_carrier UP_one_closed) simp next fix p q assume R: "p \ carrier P" "q \ carrier P" then show "eval R S h s (p \\<^bsub>P\<^esub> q) = eval R S h s p \\<^bsub>S\<^esub> eval R S h s q" proof (simp only: eval_on_carrier UP_mult_closed) from R S have "(\\<^bsub>S\<^esub> i \ {..deg R (p \\<^bsub>P\<^esub> q)}. h (coeff P (p \\<^bsub>P\<^esub> q) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) = (\\<^bsub>S\<^esub> i \ {..deg R (p \\<^bsub>P\<^esub> q)} \ {deg R (p \\<^bsub>P\<^esub> q)<..deg R p + deg R q}. h (coeff P (p \\<^bsub>P\<^esub> q) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp cong: S.finsum_cong add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def del: coeff_mult) also from R have "... = (\\<^bsub>S\<^esub> i \ {..deg R p + deg R q}. h (coeff P (p \\<^bsub>P\<^esub> q) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp only: ivl_disj_un_one deg_mult_ring) also from R S have "... = (\\<^bsub>S\<^esub> i \ {..deg R p + deg R q}. \\<^bsub>S\<^esub> k \ {..i}. h (coeff P p k) \\<^bsub>S\<^esub> h (coeff P q (i - k)) \\<^bsub>S\<^esub> (s [^]\<^bsub>S\<^esub> k \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> (i - k)))" by (simp cong: S.finsum_cong add: S.nat_pow_mult Pi_def S.m_ac S.finsum_rdistr) also from R S have "... = (\\<^bsub>S\<^esub> i\{..deg R p}. h (coeff P p i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \\<^bsub>S\<^esub> (\\<^bsub>S\<^esub> i\{..deg R q}. h (coeff P q i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp add: S.cauchy_product [THEN sym] bound.intro deg_aboveD S.m_ac Pi_def) finally show "(\\<^bsub>S\<^esub> i \ {..deg R (p \\<^bsub>P\<^esub> q)}. h (coeff P (p \\<^bsub>P\<^esub> q) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) = (\\<^bsub>S\<^esub> i \ {..deg R p}. h (coeff P p i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) \\<^bsub>S\<^esub> (\\<^bsub>S\<^esub> i \ {..deg R q}. h (coeff P q i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" . qed qed text \ The following lemma could be proved in \UP_cring\ with the additional assumption that \h\ is closed.\ lemma (in UP_pre_univ_prop) eval_const: "[| s \ carrier S; r \ carrier R |] ==> eval R S h s (monom P r 0) = h r" by (simp only: eval_on_carrier monom_closed) simp text \Further properties of the evaluation homomorphism.\ text \The following proof is complicated by the fact that in arbitrary rings one might have \<^term>\one R = zero R\.\ (* TODO: simplify by cases "one R = zero R" *) lemma (in UP_pre_univ_prop) eval_monom1: assumes S: "s \ carrier S" shows "eval R S h s (monom P \ 1) = s" proof (simp only: eval_on_carrier monom_closed R.one_closed) from S have "(\\<^bsub>S\<^esub> i\{..deg R (monom P \ 1)}. h (coeff P (monom P \ 1) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) = (\\<^bsub>S\<^esub> i\{..deg R (monom P \ 1)} \ {deg R (monom P \ 1)<..1}. h (coeff P (monom P \ 1) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp cong: S.finsum_cong del: coeff_monom add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def) also have "... = (\\<^bsub>S\<^esub> i \ {..1}. h (coeff P (monom P \ 1) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i)" by (simp only: ivl_disj_un_one deg_monom_le R.one_closed) also have "... = s" proof (cases "s = \\<^bsub>S\<^esub>") case True then show ?thesis by (simp add: Pi_def) next case False then show ?thesis by (simp add: S Pi_def) qed finally show "(\\<^bsub>S\<^esub> i \ {..deg R (monom P \ 1)}. h (coeff P (monom P \ 1) i) \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> i) = s" . qed end text \Interpretation of ring homomorphism lemmas.\ sublocale UP_univ_prop < ring_hom_cring P S Eval unfolding Eval_def by unfold_locales (fast intro: eval_ring_hom) lemma (in UP_cring) monom_pow: assumes R: "a \ carrier R" shows "(monom P a n) [^]\<^bsub>P\<^esub> m = monom P (a [^] m) (n * m)" proof (induct m) case 0 from R show ?case by simp next case Suc with R show ?case by (simp del: monom_mult add: monom_mult [THEN sym] add.commute) qed lemma (in ring_hom_cring) hom_pow [simp]: "x \ carrier R ==> h (x [^] n) = h x [^]\<^bsub>S\<^esub> (n::nat)" by (induct n) simp_all lemma (in UP_univ_prop) Eval_monom: "r \ carrier R ==> Eval (monom P r n) = h r \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> n" proof - assume R: "r \ carrier R" from R have "Eval (monom P r n) = Eval (monom P r 0 \\<^bsub>P\<^esub> (monom P \ 1) [^]\<^bsub>P\<^esub> n)" by (simp del: monom_mult add: monom_mult [THEN sym] monom_pow) also from R eval_monom1 [where s = s, folded Eval_def] have "... = h r \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> n" by (simp add: eval_const [where s = s, folded Eval_def]) finally show ?thesis . qed lemma (in UP_pre_univ_prop) eval_monom: assumes R: "r \ carrier R" and S: "s \ carrier S" shows "eval R S h s (monom P r n) = h r \\<^bsub>S\<^esub> s [^]\<^bsub>S\<^esub> n" proof - interpret UP_univ_prop R S h P s "eval R S h s" using UP_pre_univ_prop_axioms P_def R S by (auto intro: UP_univ_prop.intro UP_univ_prop_axioms.intro) from R show ?thesis by (rule Eval_monom) qed lemma (in UP_univ_prop) Eval_smult: "[| r \ carrier R; p \ carrier P |] ==> Eval (r \\<^bsub>P\<^esub> p) = h r \\<^bsub>S\<^esub> Eval p" proof - assume R: "r \ carrier R" and P: "p \ carrier P" then show ?thesis by (simp add: monom_mult_is_smult [THEN sym] eval_const [where s = s, folded Eval_def]) qed lemma ring_hom_cringI: assumes "cring R" and "cring S" and "h \ ring_hom R S" shows "ring_hom_cring R S h" by (fast intro: ring_hom_cring.intro ring_hom_cring_axioms.intro cring.axioms assms) context UP_pre_univ_prop begin lemma UP_hom_unique: assumes "ring_hom_cring P S Phi" assumes Phi: "Phi (monom P \ (Suc 0)) = s" "!!r. r \ carrier R ==> Phi (monom P r 0) = h r" assumes "ring_hom_cring P S Psi" assumes Psi: "Psi (monom P \ (Suc 0)) = s" "!!r. r \ carrier R ==> Psi (monom P r 0) = h r" and P: "p \ carrier P" and S: "s \ carrier S" shows "Phi p = Psi p" proof - interpret ring_hom_cring P S Phi by fact interpret ring_hom_cring P S Psi by fact have "Phi p = Phi (\\<^bsub>P \<^esub>i \ {..deg R p}. monom P (coeff P p i) 0 \\<^bsub>P\<^esub> monom P \ 1 [^]\<^bsub>P\<^esub> i)" by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult) also have "... = Psi (\\<^bsub>P \<^esub>i\{..deg R p}. monom P (coeff P p i) 0 \\<^bsub>P\<^esub> monom P \ 1 [^]\<^bsub>P\<^esub> i)" by (simp add: Phi Psi P Pi_def comp_def) also have "... = Psi p" by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult) finally show ?thesis . qed lemma ring_homD: assumes Phi: "Phi \ ring_hom P S" shows "ring_hom_cring P S Phi" by unfold_locales (rule Phi) theorem UP_universal_property: assumes S: "s \ carrier S" shows "\!Phi. Phi \ ring_hom P S \ extensional (carrier P) \ Phi (monom P \ 1) = s \ (\r \ carrier R. Phi (monom P r 0) = h r)" using S eval_monom1 apply (auto intro: eval_ring_hom eval_const eval_extensional) apply (rule extensionalityI) apply (auto intro: UP_hom_unique ring_homD) done end text\JE: The following lemma was added by me; it might be even lifted to a simpler locale\ context monoid begin lemma nat_pow_eone[simp]: assumes x_in_G: "x \ carrier G" shows "x [^] (1::nat) = x" using nat_pow_Suc [of x 0] unfolding nat_pow_0 [of x] unfolding l_one [OF x_in_G] by simp end context UP_ring begin abbreviation lcoeff :: "(nat =>'a) => 'a" where "lcoeff p == coeff P p (deg R p)" lemma lcoeff_nonzero2: assumes p_in_R: "p \ carrier P" and p_not_zero: "p \ \\<^bsub>P\<^esub>" shows "lcoeff p \ \" using lcoeff_nonzero [OF p_not_zero p_in_R] . subsection\The long division algorithm: some previous facts.\ lemma coeff_minus [simp]: assumes p: "p \ carrier P" and q: "q \ carrier P" shows "coeff P (p \\<^bsub>P\<^esub> q) n = coeff P p n \ coeff P q n" by (simp add: a_minus_def p q) lemma lcoeff_closed [simp]: assumes p: "p \ carrier P" shows "lcoeff p \ carrier R" using coeff_closed [OF p, of "deg R p"] by simp lemma deg_smult_decr: assumes a_in_R: "a \ carrier R" and f_in_P: "f \ carrier P" shows "deg R (a \\<^bsub>P\<^esub> f) \ deg R f" using deg_smult_ring [OF a_in_R f_in_P] by (cases "a = \", auto) lemma coeff_monom_mult: assumes R: "c \ carrier R" and P: "p \ carrier P" shows "coeff P (monom P c n \\<^bsub>P\<^esub> p) (m + n) = c \ (coeff P p m)" proof - have "coeff P (monom P c n \\<^bsub>P\<^esub> p) (m + n) = (\i\{..m + n}. (if n = i then c else \) \ coeff P p (m + n - i))" unfolding coeff_mult [OF monom_closed [OF R, of n] P, of "m + n"] unfolding coeff_monom [OF R, of n] by simp also have "(\i\{..m + n}. (if n = i then c else \) \ coeff P p (m + n - i)) = (\i\{..m + n}. (if n = i then c \ coeff P p (m + n - i) else \))" using R.finsum_cong [of "{..m + n}" "{..m + n}" "(\i::nat. (if n = i then c else \) \ coeff P p (m + n - i))" "(\i::nat. (if n = i then c \ coeff P p (m + n - i) else \))"] using coeff_closed [OF P] unfolding Pi_def simp_implies_def using R by auto also have "\ = c \ coeff P p m" using R.finsum_singleton [of n "{..m + n}" "(\i. c \ coeff P p (m + n - i))"] unfolding Pi_def using coeff_closed [OF P] using P R by auto finally show ?thesis by simp qed lemma deg_lcoeff_cancel: assumes p_in_P: "p \ carrier P" and q_in_P: "q \ carrier P" and r_in_P: "r \ carrier P" and deg_r_nonzero: "deg R r \ 0" and deg_R_p: "deg R p \ deg R r" and deg_R_q: "deg R q \ deg R r" and coeff_R_p_eq_q: "coeff P p (deg R r) = \\<^bsub>R\<^esub> (coeff P q (deg R r))" shows "deg R (p \\<^bsub>P\<^esub> q) < deg R r" proof - have deg_le: "deg R (p \\<^bsub>P\<^esub> q) \ deg R r" proof (rule deg_aboveI) fix m assume deg_r_le: "deg R r < m" show "coeff P (p \\<^bsub>P\<^esub> q) m = \" proof - have slp: "deg R p < m" and "deg R q < m" using deg_R_p deg_R_q using deg_r_le by auto then have max_sl: "max (deg R p) (deg R q) < m" by simp then have "deg R (p \\<^bsub>P\<^esub> q) < m" using deg_add [OF p_in_P q_in_P] by arith with deg_R_p deg_R_q show ?thesis using coeff_add [OF p_in_P q_in_P, of m] using deg_aboveD [of "p \\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp qed qed (simp add: p_in_P q_in_P) moreover have deg_ne: "deg R (p \\<^bsub>P\<^esub> q) \ deg R r" proof (rule ccontr) assume nz: "\ deg R (p \\<^bsub>P\<^esub> q) \ deg R r" then have deg_eq: "deg R (p \\<^bsub>P\<^esub> q) = deg R r" by simp from deg_r_nonzero have r_nonzero: "r \ \\<^bsub>P\<^esub>" by (cases "r = \\<^bsub>P\<^esub>", simp_all) have "coeff P (p \\<^bsub>P\<^esub> q) (deg R r) = \\<^bsub>R\<^esub>" using coeff_add [OF p_in_P q_in_P, of "deg R r"] using coeff_R_p_eq_q using coeff_closed [OF p_in_P, of "deg R r"] coeff_closed [OF q_in_P, of "deg R r"] by algebra with lcoeff_nonzero [OF r_nonzero r_in_P] and deg_eq show False using lcoeff_nonzero [of "p \\<^bsub>P\<^esub> q"] using p_in_P q_in_P using deg_r_nonzero by (cases "p \\<^bsub>P\<^esub> q \ \\<^bsub>P\<^esub>", auto) qed ultimately show ?thesis by simp qed lemma monom_deg_mult: assumes f_in_P: "f \ carrier P" and g_in_P: "g \ carrier P" and deg_le: "deg R g \ deg R f" and a_in_R: "a \ carrier R" shows "deg R (g \\<^bsub>P\<^esub> monom P a (deg R f - deg R g)) \ deg R f" using deg_mult_ring [OF g_in_P monom_closed [OF a_in_R, of "deg R f - deg R g"]] apply (cases "a = \") using g_in_P apply simp using deg_monom [OF _ a_in_R, of "deg R f - deg R g"] using deg_le by simp lemma deg_zero_impl_monom: assumes f_in_P: "f \ carrier P" and deg_f: "deg R f = 0" shows "f = monom P (coeff P f 0) 0" apply (rule up_eqI) using coeff_monom [OF coeff_closed [OF f_in_P], of 0 0] using f_in_P deg_f using deg_aboveD [of f _] by auto end subsection \The long division proof for commutative rings\ context UP_cring begin lemma exI3: assumes exist: "Pred x y z" shows "\ x y z. Pred x y z" using exist by blast text \Jacobson's Theorem 2.14\ lemma long_div_theorem: assumes g_in_P [simp]: "g \ carrier P" and f_in_P [simp]: "f \ carrier P" and g_not_zero: "g \ \\<^bsub>P\<^esub>" shows "\q r (k::nat). (q \ carrier P) \ (r \ carrier P) \ (lcoeff g)[^]\<^bsub>R\<^esub>k \\<^bsub>P\<^esub> f = g \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r \ (r = \\<^bsub>P\<^esub> \ deg R r < deg R g)" using f_in_P proof (induct "deg R f" arbitrary: "f" rule: nat_less_induct) case (1 f) note f_in_P [simp] = "1.prems" let ?pred = "(\ q r (k::nat). (q \ carrier P) \ (r \ carrier P) \ (lcoeff g)[^]\<^bsub>R\<^esub>k \\<^bsub>P\<^esub> f = g \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r \ (r = \\<^bsub>P\<^esub> \ deg R r < deg R g))" let ?lg = "lcoeff g" and ?lf = "lcoeff f" show ?case proof (cases "deg R f < deg R g") case True have "?pred \\<^bsub>P\<^esub> f 0" using True by force then show ?thesis by blast next case False then have deg_g_le_deg_f: "deg R g \ deg R f" by simp { let ?k = "1::nat" let ?f1 = "(g \\<^bsub>P\<^esub> (monom P (?lf) (deg R f - deg R g))) \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> (?lg \\<^bsub>P\<^esub> f)" let ?q = "monom P (?lf) (deg R f - deg R g)" have f1_in_carrier: "?f1 \ carrier P" and q_in_carrier: "?q \ carrier P" by simp_all show ?thesis proof (cases "deg R f = 0") case True { have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp have "?pred f \\<^bsub>P\<^esub> 1" using deg_zero_impl_monom [OF g_in_P deg_g] using sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]] using deg_g by simp then show ?thesis by blast } next case False note deg_f_nzero = False { have exist: "lcoeff g [^] ?k \\<^bsub>P\<^esub> f = g \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> ?f1" by (simp add: minus_add r_neg sym [ OF a_assoc [of "g \\<^bsub>P\<^esub> ?q" "\\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q)" "lcoeff g \\<^bsub>P\<^esub> f"]]) have deg_remainder_l_f: "deg R (\\<^bsub>P\<^esub> ?f1) < deg R f" proof (unfold deg_uminus [OF f1_in_carrier]) show "deg R ?f1 < deg R f" proof (rule deg_lcoeff_cancel) show "deg R (\\<^bsub>P\<^esub> (?lg \\<^bsub>P\<^esub> f)) \ deg R f" using deg_smult_ring [of ?lg f] using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp show "deg R (g \\<^bsub>P\<^esub> ?q) \ deg R f" by (simp add: monom_deg_mult [OF f_in_P g_in_P deg_g_le_deg_f, of ?lf]) show "coeff P (g \\<^bsub>P\<^esub> ?q) (deg R f) = \ coeff P (\\<^bsub>P\<^esub> (?lg \\<^bsub>P\<^esub> f)) (deg R f)" unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"] unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"] using R.finsum_cong' [of "{..deg R f}" "{..deg R f}" "(\i. coeff P g i \ (if deg R f - deg R g = deg R f - i then ?lf else \))" "(\i. if deg R g = i then coeff P g i \ ?lf else \)"] using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\i. coeff P g i \ ?lf)"] unfolding Pi_def using deg_g_le_deg_f by force qed (simp_all add: deg_f_nzero) qed then obtain q' r' k' where rem_desc: "?lg [^] (k'::nat) \\<^bsub>P\<^esub> (\\<^bsub>P\<^esub> ?f1) = g \\<^bsub>P\<^esub> q' \\<^bsub>P\<^esub> r'" and rem_deg: "(r' = \\<^bsub>P\<^esub> \ deg R r' < deg R g)" and q'_in_carrier: "q' \ carrier P" and r'_in_carrier: "r' \ carrier P" using "1.hyps" using f1_in_carrier by blast show ?thesis proof (rule exI3 [of _ "((?lg [^] k') \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI) show "(?lg [^] (Suc k')) \\<^bsub>P\<^esub> f = g \\<^bsub>P\<^esub> ((?lg [^] k') \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> q') \\<^bsub>P\<^esub> r'" proof - have "(?lg [^] (Suc k')) \\<^bsub>P\<^esub> f = (?lg [^] k') \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> ?f1)" using smult_assoc1 [OF _ _ f_in_P] using exist by simp also have "\ = (?lg [^] k') \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q) \\<^bsub>P\<^esub> ((?lg [^] k') \\<^bsub>P\<^esub> ( \\<^bsub>P\<^esub> ?f1))" using UP_smult_r_distr by simp also have "\ = (?lg [^] k') \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q) \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> q' \\<^bsub>P\<^esub> r')" unfolding rem_desc .. also have "\ = (?lg [^] k') \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q) \\<^bsub>P\<^esub> g \\<^bsub>P\<^esub> q' \\<^bsub>P\<^esub> r'" using sym [OF a_assoc [of "?lg [^] k' \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q)" "g \\<^bsub>P\<^esub> q'" "r'"]] using r'_in_carrier q'_in_carrier by simp also have "\ = (?lg [^] k') \\<^bsub>P\<^esub> (?q \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> q' \\<^bsub>P\<^esub> g \\<^bsub>P\<^esub> r'" using q'_in_carrier by (auto simp add: m_comm) also have "\ = (((?lg [^] k') \\<^bsub>P\<^esub> ?q) \\<^bsub>P\<^esub> g) \\<^bsub>P\<^esub> q' \\<^bsub>P\<^esub> g \\<^bsub>P\<^esub> r'" using smult_assoc2 q'_in_carrier "1.prems" by auto also have "\ = ((?lg [^] k') \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> q') \\<^bsub>P\<^esub> g \\<^bsub>P\<^esub> r'" using sym [OF l_distr] and q'_in_carrier by auto finally show ?thesis using m_comm q'_in_carrier by auto qed qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier) } qed } qed qed end text \The remainder theorem as corollary of the long division theorem.\ context UP_cring begin lemma deg_minus_monom: assumes a: "a \ carrier R" and R_not_trivial: "(carrier R \ {\})" shows "deg R (monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0) = 1" (is "deg R ?g = 1") proof - have "deg R ?g \ 1" proof (rule deg_aboveI) fix m assume "(1::nat) < m" then show "coeff P ?g m = \" using coeff_minus using a by auto algebra qed (simp add: a) moreover have "deg R ?g \ 1" proof (rule deg_belowI) show "coeff P ?g 1 \ \" using a using R.carrier_one_not_zero R_not_trivial by simp algebra qed (simp add: a) ultimately show ?thesis by simp qed lemma lcoeff_monom: assumes a: "a \ carrier R" and R_not_trivial: "(carrier R \ {\})" shows "lcoeff (monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0) = \" using deg_minus_monom [OF a R_not_trivial] using coeff_minus a by auto algebra lemma deg_nzero_nzero: assumes deg_p_nzero: "deg R p \ 0" shows "p \ \\<^bsub>P\<^esub>" using deg_zero deg_p_nzero by auto lemma deg_monom_minus: assumes a: "a \ carrier R" and R_not_trivial: "carrier R \ {\}" shows "deg R (monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0) = 1" (is "deg R ?g = 1") proof - have "deg R ?g \ 1" proof (rule deg_aboveI) fix m::nat assume "1 < m" then show "coeff P ?g m = \" using coeff_minus [OF monom_closed [OF R.one_closed, of 1] monom_closed [OF a, of 0], of m] using coeff_monom [OF R.one_closed, of 1 m] using coeff_monom [OF a, of 0 m] by auto algebra qed (simp add: a) moreover have "1 \ deg R ?g" proof (rule deg_belowI) show "coeff P ?g 1 \ \" using coeff_minus [OF monom_closed [OF R.one_closed, of 1] monom_closed [OF a, of 0], of 1] using coeff_monom [OF R.one_closed, of 1 1] using coeff_monom [OF a, of 0 1] using R_not_trivial using R.carrier_one_not_zero by auto algebra qed (simp add: a) ultimately show ?thesis by simp qed lemma eval_monom_expr: assumes a: "a \ carrier R" shows "eval R R id a (monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0) = \" (is "eval R R id a ?g = _") proof - interpret UP_pre_univ_prop R R id by unfold_locales simp have eval_ring_hom: "eval R R id a \ ring_hom P R" using eval_ring_hom [OF a] by simp interpret ring_hom_cring P R "eval R R id a" by unfold_locales (rule eval_ring_hom) have mon1_closed: "monom P \\<^bsub>R\<^esub> 1 \ carrier P" and mon0_closed: "monom P a 0 \ carrier P" and min_mon0_closed: "\\<^bsub>P\<^esub> monom P a 0 \ carrier P" using a R.a_inv_closed by auto have "eval R R id a ?g = eval R R id a (monom P \ 1) \ eval R R id a (monom P a 0)" by (simp add: a_minus_def mon0_closed) also have "\ = a \ a" using eval_monom [OF R.one_closed a, of 1] using eval_monom [OF a a, of 0] using a by simp also have "\ = \" using a by algebra finally show ?thesis by simp qed lemma remainder_theorem_exist: assumes f: "f \ carrier P" and a: "a \ carrier R" and R_not_trivial: "carrier R \ {\}" shows "\ q r. (q \ carrier P) \ (r \ carrier P) \ f = (monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0) \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r \ (deg R r = 0)" (is "\ q r. (q \ carrier P) \ (r \ carrier P) \ f = ?g \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r \ (deg R r = 0)") proof - let ?g = "monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0" from deg_minus_monom [OF a R_not_trivial] have deg_g_nzero: "deg R ?g \ 0" by simp have "\q r (k::nat). q \ carrier P \ r \ carrier P \ lcoeff ?g [^] k \\<^bsub>P\<^esub> f = ?g \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r \ (r = \\<^bsub>P\<^esub> \ deg R r < deg R ?g)" using long_div_theorem [OF _ f deg_nzero_nzero [OF deg_g_nzero]] a by auto then show ?thesis unfolding lcoeff_monom [OF a R_not_trivial] unfolding deg_monom_minus [OF a R_not_trivial] using smult_one [OF f] using deg_zero by force qed lemma remainder_theorem_expression: assumes f [simp]: "f \ carrier P" and a [simp]: "a \ carrier R" and q [simp]: "q \ carrier P" and r [simp]: "r \ carrier P" and R_not_trivial: "carrier R \ {\}" and f_expr: "f = (monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0) \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r" (is "f = ?g \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r" is "f = ?gq \\<^bsub>P\<^esub> r") and deg_r_0: "deg R r = 0" shows "r = monom P (eval R R id a f) 0" proof - interpret UP_pre_univ_prop R R id P by standard simp have eval_ring_hom: "eval R R id a \ ring_hom P R" using eval_ring_hom [OF a] by simp have "eval R R id a f = eval R R id a ?gq \\<^bsub>R\<^esub> eval R R id a r" unfolding f_expr using ring_hom_add [OF eval_ring_hom] by auto also have "\ = ((eval R R id a ?g) \ (eval R R id a q)) \\<^bsub>R\<^esub> eval R R id a r" using ring_hom_mult [OF eval_ring_hom] by auto also have "\ = \ \ eval R R id a r" unfolding eval_monom_expr [OF a] using eval_ring_hom unfolding ring_hom_def using q unfolding Pi_def by simp also have "\ = eval R R id a r" using eval_ring_hom unfolding ring_hom_def using r unfolding Pi_def by simp finally have eval_eq: "eval R R id a f = eval R R id a r" by simp from deg_zero_impl_monom [OF r deg_r_0] have "r = monom P (coeff P r 0) 0" by simp with eval_const [OF a, of "coeff P r 0"] eval_eq show ?thesis by auto qed corollary remainder_theorem: assumes f [simp]: "f \ carrier P" and a [simp]: "a \ carrier R" and R_not_trivial: "carrier R \ {\}" shows "\ q r. (q \ carrier P) \ (r \ carrier P) \ f = (monom P \\<^bsub>R\<^esub> 1 \\<^bsub>P\<^esub> monom P a 0) \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> monom P (eval R R id a f) 0" (is "\ q r. (q \ carrier P) \ (r \ carrier P) \ f = ?g \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> monom P (eval R R id a f) 0") proof - from remainder_theorem_exist [OF f a R_not_trivial] obtain q r where q_r: "q \ carrier P \ r \ carrier P \ f = ?g \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r" and deg_r: "deg R r = 0" by force with remainder_theorem_expression [OF f a _ _ R_not_trivial, of q r] show ?thesis by auto qed end subsection \Sample Application of Evaluation Homomorphism\ lemma UP_pre_univ_propI: assumes "cring R" and "cring S" and "h \ ring_hom R S" shows "UP_pre_univ_prop R S h" using assms by (auto intro!: UP_pre_univ_prop.intro ring_hom_cring.intro ring_hom_cring_axioms.intro UP_cring.intro) definition INTEG :: "int ring" where "INTEG = \carrier = UNIV, mult = (*), one = 1, zero = 0, add = (+)\" lemma INTEG_cring: "cring INTEG" by (unfold INTEG_def) (auto intro!: cringI abelian_groupI comm_monoidI left_minus distrib_right) lemma INTEG_id_eval: "UP_pre_univ_prop INTEG INTEG id" by (fast intro: UP_pre_univ_propI INTEG_cring id_ring_hom) text \ Interpretation now enables to import all theorems and lemmas valid in the context of homomorphisms between \<^term>\INTEG\ and \<^term>\UP INTEG\ globally. \ interpretation INTEG: UP_pre_univ_prop INTEG INTEG id "UP INTEG" using INTEG_id_eval by simp_all lemma INTEG_closed [intro, simp]: "z \ carrier INTEG" by (unfold INTEG_def) simp lemma INTEG_mult [simp]: "mult INTEG z w = z * w" by (unfold INTEG_def) simp lemma INTEG_pow [simp]: "pow INTEG z n = z ^ n" by (induct n) (simp_all add: INTEG_def nat_pow_def) lemma "eval INTEG INTEG id 10 (monom (UP INTEG) 5 2) = 500" by (simp add: INTEG.eval_monom) end diff --git a/src/HOL/Analysis/Continuous_Extension.thy b/src/HOL/Analysis/Continuous_Extension.thy --- a/src/HOL/Analysis/Continuous_Extension.thy +++ b/src/HOL/Analysis/Continuous_Extension.thy @@ -1,551 +1,551 @@ (* Title: HOL/Analysis/Continuous_Extension.thy Authors: LC Paulson, based on material from HOL Light *) section \Continuous Extensions of Functions\ theory Continuous_Extension imports Starlike begin subsection\Partitions of unity subordinate to locally finite open coverings\ text\A difference from HOL Light: all summations over infinite sets equal zero, so the "support" must be made explicit in the summation below!\ proposition subordinate_partition_of_unity: fixes S :: "'a::metric_space set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" and fin: "\x. x \ S \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" obtains F :: "['a set, 'a] \ real" where "\U. U \ \ \ continuous_on S (F U) \ (\x \ S. 0 \ F U x)" and "\x U. \U \ \; x \ S; x \ U\ \ F U x = 0" and "\x. x \ S \ supp_sum (\W. F W x) \ = 1" and "\x. x \ S \ \V. open V \ x \ V \ finite {U \ \. \x\V. F U x \ 0}" proof (cases "\W. W \ \ \ S \ W") case True then obtain W where "W \ \" "S \ W" by metis then show ?thesis by (rule_tac F = "\V x. if V = W then 1 else 0" in that) (auto simp: supp_sum_def support_on_def) next case False have nonneg: "0 \ supp_sum (\V. setdist {x} (S - V)) \" for x by (simp add: supp_sum_def sum_nonneg) have sd_pos: "0 < setdist {x} (S - V)" if "V \ \" "x \ S" "x \ V" for V x proof - have "closedin (top_of_set S) (S - V)" by (simp add: Diff_Diff_Int closedin_def opC openin_open_Int \V \ \\) with that False setdist_pos_le [of "{x}" "S - V"] show ?thesis using setdist_gt_0_closedin by fastforce qed have ss_pos: "0 < supp_sum (\V. setdist {x} (S - V)) \" if "x \ S" for x proof - obtain U where "U \ \" "x \ U" using \x \ S\ \S \ \\\ by blast obtain V where "open V" "x \ V" "finite {U \ \. U \ V \ {}}" using \x \ S\ fin by blast then have *: "finite {A \ \. \ S \ A \ x \ closure (S - A)}" using closure_def that by (blast intro: rev_finite_subset) have "x \ closure (S - U)" using \U \ \\ \x \ U\ opC open_Int_closure_eq_empty by fastforce then show ?thesis apply (simp add: setdist_eq_0_sing_1 supp_sum_def support_on_def) apply (rule ordered_comm_monoid_add_class.sum_pos2 [OF *, of U]) using \U \ \\ \x \ U\ False apply (auto simp: sd_pos that) done qed define F where "F \ \W x. if x \ S then setdist {x} (S - W) / supp_sum (\V. setdist {x} (S - V)) \ else 0" show ?thesis proof (rule_tac F = F in that) have "continuous_on S (F U)" if "U \ \" for U proof - have *: "continuous_on S (\x. supp_sum (\V. setdist {x} (S - V)) \)" proof (clarsimp simp add: continuous_on_eq_continuous_within) fix x assume "x \ S" then obtain X where "open X" and x: "x \ S \ X" and finX: "finite {U \ \. U \ X \ {}}" using assms by blast then have OSX: "openin (top_of_set S) (S \ X)" by blast have sumeq: "\x. x \ S \ X \ (\V | V \ \ \ V \ X \ {}. setdist {x} (S - V)) = supp_sum (\V. setdist {x} (S - V)) \" apply (simp add: supp_sum_def) apply (rule sum.mono_neutral_right [OF finX]) apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff) apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE) done show "continuous (at x within S) (\x. supp_sum (\V. setdist {x} (S - V)) \)" apply (rule continuous_transform_within_openin [where f = "\x. (sum (\V. setdist {x} (S - V)) {V \ \. V \ X \ {}})" and S ="S \ X"]) apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+ apply (simp add: sumeq) done qed show ?thesis apply (simp add: F_def) apply (rule continuous_intros *)+ using ss_pos apply force done qed moreover have "\U \ \; x \ S\ \ 0 \ F U x" for U x using nonneg [of x] by (simp add: F_def field_split_simps) ultimately show "\U. U \ \ \ continuous_on S (F U) \ (\x\S. 0 \ F U x)" by metis next show "\x U. \U \ \; x \ S; x \ U\ \ F U x = 0" by (simp add: setdist_eq_0_sing_1 closure_def F_def) next show "supp_sum (\W. F W x) \ = 1" if "x \ S" for x using that ss_pos [OF that] by (simp add: F_def field_split_simps supp_sum_divide_distrib [symmetric]) next show "\V. open V \ x \ V \ finite {U \ \. \x\V. F U x \ 0}" if "x \ S" for x using fin [OF that] that by (fastforce simp: setdist_eq_0_sing_1 closure_def F_def elim!: rev_finite_subset) qed qed subsection\Urysohn's Lemma for Euclidean Spaces\ text \For Euclidean spaces the proof is easy using distances.\ lemma Urysohn_both_ne: assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" "S \ {}" "T \ {}" "a \ b" obtains f :: "'a::euclidean_space \ 'b::real_normed_vector" where "continuous_on U f" "\x. x \ U \ f x \ closed_segment a b" "\x. x \ U \ (f x = a \ x \ S)" "\x. x \ U \ (f x = b \ x \ T)" proof - have S0: "\x. x \ U \ setdist {x} S = 0 \ x \ S" using \S \ {}\ US setdist_eq_0_closedin by auto have T0: "\x. x \ U \ setdist {x} T = 0 \ x \ T" using \T \ {}\ UT setdist_eq_0_closedin by auto have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \ U" for x proof - have "\ (setdist {x} S = 0 \ setdist {x} T = 0)" using assms by (metis IntI empty_iff setdist_eq_0_closedin that) then show ?thesis by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le) qed define f where "f \ \x. a + (setdist {x} S / (setdist {x} S + setdist {x} T)) *\<^sub>R (b - a)" show ?thesis proof (rule_tac f = f in that) show "continuous_on U f" using sdpos unfolding f_def by (intro continuous_intros | force)+ show "f x \ closed_segment a b" if "x \ U" for x unfolding f_def apply (simp add: closed_segment_def) apply (rule_tac x="(setdist {x} S / (setdist {x} S + setdist {x} T))" in exI) using sdpos that apply (simp add: algebra_simps) done show "\x. x \ U \ (f x = a \ x \ S)" using S0 \a \ b\ f_def sdpos by force show "(f x = b \ x \ T)" if "x \ U" for x proof - have "f x = b \ (setdist {x} S / (setdist {x} S + setdist {x} T)) = 1" unfolding f_def apply (rule iffI) apply (metis \a \ b\ add_diff_cancel_left' eq_iff_diff_eq_0 pth_1 real_vector.scale_right_imp_eq, force) done also have "... \ setdist {x} T = 0 \ setdist {x} S \ 0" using sdpos that by (simp add: field_split_simps) linarith also have "... \ x \ T" using \S \ {}\ \T \ {}\ \S \ T = {}\ that by (force simp: S0 T0) finally show ?thesis . qed qed qed proposition Urysohn_local_strong: assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" "a \ b" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on U f" "\x. x \ U \ f x \ closed_segment a b" "\x. x \ U \ (f x = a \ x \ S)" "\x. x \ U \ (f x = b \ x \ T)" proof (cases "S = {}") case True show ?thesis proof (cases "T = {}") case True show ?thesis proof (rule_tac f = "\x. midpoint a b" in that) show "continuous_on U (\x. midpoint a b)" by (intro continuous_intros) show "midpoint a b \ closed_segment a b" using csegment_midpoint_subset by blast show "(midpoint a b = a) = (x \ S)" for x using \S = {}\ \a \ b\ by simp show "(midpoint a b = b) = (x \ T)" for x using \T = {}\ \a \ b\ by simp qed next case False show ?thesis proof (cases "T = U") case True with \S = {}\ \a \ b\ show ?thesis by (rule_tac f = "\x. b" in that) (auto) next case False with UT closedin_subset obtain c where c: "c \ U" "c \ T" by fastforce obtain f where f: "continuous_on U f" "\x. x \ U \ f x \ closed_segment (midpoint a b) b" "\x. x \ U \ (f x = midpoint a b \ x = c)" "\x. x \ U \ (f x = b \ x \ T)" apply (rule Urysohn_both_ne [of U "{c}" T "midpoint a b" "b"]) using c \T \ {}\ assms apply simp_all done show ?thesis apply (rule_tac f=f in that) using \S = {}\ \T \ {}\ f csegment_midpoint_subset notin_segment_midpoint [OF \a \ b\] apply force+ done qed qed next case False show ?thesis proof (cases "T = {}") case True show ?thesis proof (cases "S = U") case True with \T = {}\ \a \ b\ show ?thesis by (rule_tac f = "\x. a" in that) (auto) next case False with US closedin_subset obtain c where c: "c \ U" "c \ S" by fastforce obtain f where f: "continuous_on U f" "\x. x \ U \ f x \ closed_segment a (midpoint a b)" "\x. x \ U \ (f x = midpoint a b \ x = c)" "\x. x \ U \ (f x = a \ x \ S)" apply (rule Urysohn_both_ne [of U S "{c}" a "midpoint a b"]) using c \S \ {}\ assms apply simp_all apply (metis midpoint_eq_endpoint) done show ?thesis apply (rule_tac f=f in that) using \S \ {}\ \T = {}\ f \a \ b\ apply simp_all apply (metis (no_types) closed_segment_commute csegment_midpoint_subset midpoint_sym subset_iff) apply (metis closed_segment_commute midpoint_sym notin_segment_midpoint) done qed next case False show ?thesis using Urysohn_both_ne [OF US UT \S \ T = {}\ \S \ {}\ \T \ {}\ \a \ b\] that by blast qed qed lemma Urysohn_local: assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on U f" "\x. x \ U \ f x \ closed_segment a b" "\x. x \ S \ f x = a" "\x. x \ T \ f x = b" proof (cases "a = b") case True then show ?thesis by (rule_tac f = "\x. b" in that) (auto) next case False then show ?thesis apply (rule Urysohn_local_strong [OF assms]) apply (erule that, assumption) apply (meson US closedin_singleton closedin_trans) apply (meson UT closedin_singleton closedin_trans) done qed lemma Urysohn_strong: assumes US: "closed S" and UT: "closed T" and "S \ T = {}" "a \ b" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on UNIV f" "\x. f x \ closed_segment a b" "\x. f x = a \ x \ S" "\x. f x = b \ x \ T" using assms by (auto intro: Urysohn_local_strong [of UNIV S T]) proposition Urysohn: assumes US: "closed S" and UT: "closed T" and "S \ T = {}" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on UNIV f" "\x. f x \ closed_segment a b" "\x. x \ S \ f x = a" "\x. x \ T \ f x = b" using assms by (auto intro: Urysohn_local [of UNIV S T a b]) subsection\Dugundji's Extension Theorem and Tietze Variants\ -text \See \cite{dugundji}.\ +text \See \<^cite>\"dugundji"\.\ lemma convex_supp_sum: assumes "convex S" and 1: "supp_sum u I = 1" and "\i. i \ I \ 0 \ u i \ (u i = 0 \ f i \ S)" shows "supp_sum (\i. u i *\<^sub>R f i) I \ S" proof - have fin: "finite {i \ I. u i \ 0}" using 1 sum.infinite by (force simp: supp_sum_def support_on_def) then have "supp_sum (\i. u i *\<^sub>R f i) I = sum (\i. u i *\<^sub>R f i) {i \ I. u i \ 0}" by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def) also have "... \ S" using 1 assms by (force simp: supp_sum_def support_on_def intro: convex_sum [OF fin \convex S\]) finally show ?thesis . qed theorem Dugundji: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "convex C" "C \ {}" and cloin: "closedin (top_of_set U) S" and contf: "continuous_on S f" and "f ` S \ C" obtains g where "continuous_on U g" "g ` U \ C" "\x. x \ S \ g x = f x" proof (cases "S = {}") case True then show thesis apply (rule_tac g="\x. SOME y. y \ C" in that) apply (rule continuous_intros) apply (meson all_not_in_conv \C \ {}\ image_subsetI someI_ex, simp) done next case False then have sd_pos: "\x. \x \ U; x \ S\ \ 0 < setdist {x} S" using setdist_eq_0_closedin [OF cloin] le_less setdist_pos_le by fastforce define \ where "\ = {ball x (setdist {x} S / 2) |x. x \ U - S}" have [simp]: "\T. T \ \ \ open T" by (auto simp: \_def) have USS: "U - S \ \\" by (auto simp: sd_pos \_def) obtain \ where USsub: "U - S \ \\" and nbrhd: "\U. U \ \ \ open U \ (\T. T \ \ \ U \ T)" and fin: "\x. x \ U - S \ \V. open V \ x \ V \ finite {U. U \ \ \ U \ V \ {}}" by (rule paracompact [OF USS]) auto have "\v a. v \ U \ v \ S \ a \ S \ T \ ball v (setdist {v} S / 2) \ dist v a \ 2 * setdist {v} S" if "T \ \" for T proof - obtain v where v: "T \ ball v (setdist {v} S / 2)" "v \ U" "v \ S" using \T \ \\ nbrhd by (force simp: \_def) then obtain a where "a \ S" "dist v a < 2 * setdist {v} S" using setdist_ltE [of "{v}" S "2 * setdist {v} S"] using False sd_pos by force with v show ?thesis apply (rule_tac x=v in exI) apply (rule_tac x=a in exI, auto) done qed then obtain \ \ where VA: "\T. T \ \ \ \ T \ U \ \ T \ S \ \ T \ S \ T \ ball (\ T) (setdist {\ T} S / 2) \ dist (\ T) (\ T) \ 2 * setdist {\ T} S" by metis have sdle: "setdist {\ T} S \ 2 * setdist {v} S" if "T \ \" "v \ T" for T v using setdist_Lipschitz [of "\ T" S v] VA [OF \T \ \\] \v \ T\ by auto have d6: "dist a (\ T) \ 6 * dist a v" if "T \ \" "v \ T" "a \ S" for T v a proof - have "dist (\ T) v < setdist {\ T} S / 2" using that VA mem_ball by blast also have "\ \ setdist {v} S" using sdle [OF \T \ \\ \v \ T\] by simp also have vS: "setdist {v} S \ dist a v" by (simp add: setdist_le_dist setdist_sym \a \ S\) finally have VTV: "dist (\ T) v < dist a v" . have VTS: "setdist {\ T} S \ 2 * dist a v" using sdle that vS by force have "dist a (\ T) \ dist a v + dist v (\ T) + dist (\ T) (\ T)" by (metis add.commute add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le) also have "\ \ dist a v + dist a v + dist (\ T) (\ T)" using VTV by (simp add: dist_commute) also have "\ \ 2 * dist a v + 2 * setdist {\ T} S" using VA [OF \T \ \\] by auto finally show ?thesis using VTS by linarith qed obtain H :: "['a set, 'a] \ real" where Hcont: "\Z. Z \ \ \ continuous_on (U-S) (H Z)" and Hge0: "\Z x. \Z \ \; x \ U-S\ \ 0 \ H Z x" and Heq0: "\x Z. \Z \ \; x \ U-S; x \ Z\ \ H Z x = 0" and H1: "\x. x \ U-S \ supp_sum (\W. H W x) \ = 1" and Hfin: "\x. x \ U-S \ \V. open V \ x \ V \ finite {U \ \. \x\V. H U x \ 0}" apply (rule subordinate_partition_of_unity [OF USsub _ fin]) using nbrhd by auto define g where "g \ \x. if x \ S then f x else supp_sum (\T. H T x *\<^sub>R f(\ T)) \" show ?thesis proof (rule that) show "continuous_on U g" proof (clarsimp simp: continuous_on_eq_continuous_within) fix a assume "a \ U" show "continuous (at a within U) g" proof (cases "a \ S") case True show ?thesis proof (clarsimp simp add: continuous_within_topological) fix W assume "open W" "g a \ W" then obtain e where "0 < e" and e: "ball (f a) e \ W" using openE True g_def by auto have "continuous (at a within S) f" using True contf continuous_on_eq_continuous_within by blast then obtain d where "0 < d" and d: "\x. \x \ S; dist x a < d\ \ dist (f x) (f a) < e" using continuous_within_eps_delta \0 < e\ by force have "g y \ ball (f a) e" if "y \ U" and y: "y \ ball a (d / 6)" for y proof (cases "y \ S") case True then have "dist (f a) (f y) < e" by (metis ball_divide_subset_numeral dist_commute in_mono mem_ball y d) then show ?thesis by (simp add: True g_def) next case False have *: "dist (f (\ T)) (f a) < e" if "T \ \" "H T y \ 0" for T proof - have "y \ T" using Heq0 that False \y \ U\ by blast have "dist (\ T) a < d" using d6 [OF \T \ \\ \y \ T\ \a \ S\] y by (simp add: dist_commute mult.commute) then show ?thesis using VA [OF \T \ \\] by (auto simp: d) qed have "supp_sum (\T. H T y *\<^sub>R f (\ T)) \ \ ball (f a) e" apply (rule convex_supp_sum [OF convex_ball]) apply (simp_all add: False H1 Hge0 \y \ U\) by (metis dist_commute *) then show ?thesis by (simp add: False g_def) qed then show "\A. open A \ a \ A \ (\y\U. y \ A \ g y \ W)" apply (rule_tac x = "ball a (d / 6)" in exI) using e \0 < d\ by fastforce qed next case False obtain N where N: "open N" "a \ N" and finN: "finite {U \ \. \a\N. H U a \ 0}" using Hfin False \a \ U\ by auto have oUS: "openin (top_of_set U) (U - S)" using cloin by (simp add: openin_diff) have HcontU: "continuous (at a within U) (H T)" if "T \ \" for T using Hcont [OF \T \ \\] False \a \ U\ \T \ \\ apply (simp add: continuous_on_eq_continuous_within continuous_within) apply (rule Lim_transform_within_set) using oUS apply (force simp: eventually_at openin_contains_ball dist_commute dest!: bspec)+ done show ?thesis proof (rule continuous_transform_within_openin [OF _ oUS]) show "continuous (at a within U) (\x. supp_sum (\T. H T x *\<^sub>R f (\ T)) \)" proof (rule continuous_transform_within_openin) show "continuous (at a within U) (\x. \T\{U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T))" by (force intro: continuous_intros HcontU)+ next show "openin (top_of_set U) ((U - S) \ N)" using N oUS openin_trans by blast next show "a \ (U - S) \ N" using False \a \ U\ N by blast next show "\x. x \ (U - S) \ N \ (\T \ {U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T)) = supp_sum (\T. H T x *\<^sub>R f (\ T)) \" by (auto simp: supp_sum_def support_on_def intro: sum.mono_neutral_right [OF finN]) qed next show "a \ U - S" using False \a \ U\ by blast next show "\x. x \ U - S \ supp_sum (\T. H T x *\<^sub>R f (\ T)) \ = g x" by (simp add: g_def) qed qed qed show "g ` U \ C" using \f ` S \ C\ VA by (fastforce simp: g_def Hge0 intro!: convex_supp_sum [OF \convex C\] H1) show "\x. x \ S \ g x = f x" by (simp add: g_def) qed qed corollary Tietze: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "0 \ B" and "\x. x \ S \ norm(f x) \ B" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ norm(g x) \ B" using assms by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f]) corollary\<^marker>\tag unimportant\ Tietze_closed_interval: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::euclidean_space" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "cbox a b \ {}" and "\x. x \ S \ f x \ cbox a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ g x \ cbox a b" apply (rule Dugundji [of "cbox a b" U S f]) using assms by auto corollary\<^marker>\tag unimportant\ Tietze_closed_interval_1: fixes f :: "'a::{metric_space,second_countable_topology} \ real" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "a \ b" and "\x. x \ S \ f x \ cbox a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ g x \ cbox a b" apply (rule Dugundji [of "cbox a b" U S f]) using assms by (auto simp: image_subset_iff) corollary\<^marker>\tag unimportant\ Tietze_open_interval: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::euclidean_space" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "box a b \ {}" and "\x. x \ S \ f x \ box a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ g x \ box a b" apply (rule Dugundji [of "box a b" U S f]) using assms by auto corollary\<^marker>\tag unimportant\ Tietze_open_interval_1: fixes f :: "'a::{metric_space,second_countable_topology} \ real" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "a < b" and no: "\x. x \ S \ f x \ box a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ g x \ box a b" apply (rule Dugundji [of "box a b" U S f]) using assms by (auto simp: image_subset_iff) corollary\<^marker>\tag unimportant\ Tietze_unbounded: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "continuous_on S f" and "closedin (top_of_set U) S" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" apply (rule Dugundji [of UNIV U S f]) using assms by auto end diff --git a/src/HOL/Analysis/Infinite_Sum.thy b/src/HOL/Analysis/Infinite_Sum.thy --- a/src/HOL/Analysis/Infinite_Sum.thy +++ b/src/HOL/Analysis/Infinite_Sum.thy @@ -1,2433 +1,2433 @@ (* Title: HOL/Analysis/Infinite_Sum.thy Author: Dominique Unruh, University of Tartu Manuel Eberl, University of Innsbruck A theory of sums over possibly infinite sets. *) section \Infinite sums\ \<^latex>\\label{section:Infinite_Sum}\ text \In this theory, we introduce the definition of infinite sums, i.e., sums ranging over an infinite, potentially uncountable index set with no particular ordering. (This is different from series. Those are sums indexed by natural numbers, and the order of the index set matters.) Our definition is quite standard: $s:=\sum_{x\in A} f(x)$ is the limit of finite sums $s_F:=\sum_{x\in F} f(x)$ for increasing $F$. That is, $s$ is the limit of the net $s_F$ where $F$ are finite subsets of $A$ ordered by inclusion. We believe that this is the standard definition for such sums. -See, e.g., Definition 4.11 in \cite{conway2013course}. +See, e.g., Definition 4.11 in \<^cite>\"conway2013course"\. This definition is quite general: it is well-defined whenever $f$ takes values in some commutative monoid endowed with a Hausdorff topology. (Examples are reals, complex numbers, normed vector spaces, and more.)\ theory Infinite_Sum imports Elementary_Topology "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Complex_Order" begin subsection \Definition and syntax\ definition has_sum :: \('a \ 'b :: {comm_monoid_add, topological_space}) \ 'a set \ 'b \ bool\ where \has_sum f A x \ (sum f \ x) (finite_subsets_at_top A)\ definition summable_on :: "('a \ 'b::{comm_monoid_add, topological_space}) \ 'a set \ bool" (infixr "summable'_on" 46) where "f summable_on A \ (\x. has_sum f A x)" definition infsum :: "('a \ 'b::{comm_monoid_add,t2_space}) \ 'a set \ 'b" where "infsum f A = (if f summable_on A then Lim (finite_subsets_at_top A) (sum f) else 0)" abbreviation abs_summable_on :: "('a \ 'b::real_normed_vector) \ 'a set \ bool" (infixr "abs'_summable'_on" 46) where "f abs_summable_on A \ (\x. norm (f x)) summable_on A" syntax (ASCII) "_infsum" :: "pttrn \ 'a set \ 'b \ 'b::topological_comm_monoid_add" ("(3INFSUM (_/:_)./ _)" [0, 51, 10] 10) syntax "_infsum" :: "pttrn \ 'a set \ 'b \ 'b::topological_comm_monoid_add" ("(2\\<^sub>\(_/\_)./ _)" [0, 51, 10] 10) translations \ \Beware of argument permutation!\ "\\<^sub>\i\A. b" \ "CONST infsum (\i. b) A" syntax (ASCII) "_univinfsum" :: "pttrn \ 'a \ 'a" ("(3INFSUM _./ _)" [0, 10] 10) syntax "_univinfsum" :: "pttrn \ 'a \ 'a" ("(2\\<^sub>\_./ _)" [0, 10] 10) translations "\\<^sub>\x. t" \ "CONST infsum (\x. t) (CONST UNIV)" syntax (ASCII) "_qinfsum" :: "pttrn \ bool \ 'a \ 'a" ("(3INFSUM _ |/ _./ _)" [0, 0, 10] 10) syntax "_qinfsum" :: "pttrn \ bool \ 'a \ 'a" ("(2\\<^sub>\_ | (_)./ _)" [0, 0, 10] 10) translations "\\<^sub>\x|P. t" => "CONST infsum (\x. t) {x. P}" print_translation \ let fun sum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] = if x <> y then raise Match else let val x' = Syntax_Trans.mark_bound_body (x, Tx); val t' = subst_bound (x', t); val P' = subst_bound (x', P); in Syntax.const @{syntax_const "_qinfsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t' end | sum_tr' _ = raise Match; in [(@{const_syntax infsum}, K sum_tr')] end \ subsection \General properties\ lemma infsumI: fixes f g :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \has_sum f A x\ shows \infsum f A = x\ by (metis assms finite_subsets_at_top_neq_bot infsum_def summable_on_def has_sum_def tendsto_Lim) lemma infsum_eqI: fixes f g :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \x = y\ assumes \has_sum f A x\ assumes \has_sum g B y\ shows \infsum f A = infsum g B\ by (metis assms(1) assms(2) assms(3) finite_subsets_at_top_neq_bot infsum_def summable_on_def has_sum_def tendsto_Lim) lemma infsum_eqI': fixes f g :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \\x. has_sum f A x \ has_sum g B x\ shows \infsum f A = infsum g B\ by (metis assms infsum_def infsum_eqI summable_on_def) lemma infsum_not_exists: fixes f :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \\ f summable_on A\ shows \infsum f A = 0\ by (simp add: assms infsum_def) lemma summable_iff_has_sum_infsum: "f summable_on A \ has_sum f A (infsum f A)" using infsumI summable_on_def by blast lemma has_sum_infsum[simp]: assumes \f summable_on S\ shows \has_sum f S (infsum f S)\ using assms by (auto simp: summable_on_def infsum_def has_sum_def tendsto_Lim) lemma has_sum_cong_neutral: fixes f g :: \'a \ 'b::{comm_monoid_add, topological_space}\ assumes \\x. x\T-S \ g x = 0\ assumes \\x. x\S-T \ f x = 0\ assumes \\x. x\S\T \ f x = g x\ shows "has_sum f S x \ has_sum g T x" proof - have \eventually P (filtermap (sum f) (finite_subsets_at_top S)) = eventually P (filtermap (sum g) (finite_subsets_at_top T))\ for P proof assume \eventually P (filtermap (sum f) (finite_subsets_at_top S))\ then obtain F0 where \finite F0\ and \F0 \ S\ and F0_P: \\F. finite F \ F \ S \ F \ F0 \ P (sum f F)\ by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) define F0' where \F0' = F0 \ T\ have [simp]: \finite F0'\ \F0' \ T\ by (simp_all add: F0'_def \finite F0\) have \P (sum g F)\ if \finite F\ \F \ T\ \F \ F0'\ for F proof - have \P (sum f ((F\S) \ (F0\S)))\ by (intro F0_P) (use \F0 \ S\ \finite F0\ that in auto) also have \sum f ((F\S) \ (F0\S)) = sum g F\ by (intro sum.mono_neutral_cong) (use that \finite F0\ F0'_def assms in auto) finally show ?thesis . qed with \F0' \ T\ \finite F0'\ show \eventually P (filtermap (sum g) (finite_subsets_at_top T))\ by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) next assume \eventually P (filtermap (sum g) (finite_subsets_at_top T))\ then obtain F0 where \finite F0\ and \F0 \ T\ and F0_P: \\F. finite F \ F \ T \ F \ F0 \ P (sum g F)\ by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) define F0' where \F0' = F0 \ S\ have [simp]: \finite F0'\ \F0' \ S\ by (simp_all add: F0'_def \finite F0\) have \P (sum f F)\ if \finite F\ \F \ S\ \F \ F0'\ for F proof - have \P (sum g ((F\T) \ (F0\T)))\ by (intro F0_P) (use \F0 \ T\ \finite F0\ that in auto) also have \sum g ((F\T) \ (F0\T)) = sum f F\ by (intro sum.mono_neutral_cong) (use that \finite F0\ F0'_def assms in auto) finally show ?thesis . qed with \F0' \ S\ \finite F0'\ show \eventually P (filtermap (sum f) (finite_subsets_at_top S))\ by (metis (no_types, lifting) eventually_filtermap eventually_finite_subsets_at_top) qed then have tendsto_x: "(sum f \ x) (finite_subsets_at_top S) \ (sum g \ x) (finite_subsets_at_top T)" for x by (simp add: le_filter_def filterlim_def) then show ?thesis by (simp add: has_sum_def) qed lemma summable_on_cong_neutral: fixes f g :: \'a \ 'b::{comm_monoid_add, topological_space}\ assumes \\x. x\T-S \ g x = 0\ assumes \\x. x\S-T \ f x = 0\ assumes \\x. x\S\T \ f x = g x\ shows "f summable_on S \ g summable_on T" using has_sum_cong_neutral[of T S g f, OF assms] by (simp add: summable_on_def) lemma infsum_cong_neutral: fixes f g :: \'a \ 'b::{comm_monoid_add, t2_space}\ assumes \\x. x\T-S \ g x = 0\ assumes \\x. x\S-T \ f x = 0\ assumes \\x. x\S\T \ f x = g x\ shows \infsum f S = infsum g T\ by (smt (verit, best) assms has_sum_cong_neutral infsum_eqI') lemma has_sum_cong: assumes "\x. x\A \ f x = g x" shows "has_sum f A x \ has_sum g A x" using assms by (intro has_sum_cong_neutral) auto lemma summable_on_cong: assumes "\x. x\A \ f x = g x" shows "f summable_on A \ g summable_on A" by (metis assms summable_on_def has_sum_cong) lemma infsum_cong: assumes "\x. x\A \ f x = g x" shows "infsum f A = infsum g A" using assms infsum_eqI' has_sum_cong by blast lemma summable_on_cofin_subset: fixes f :: "'a \ 'b::topological_ab_group_add" assumes "f summable_on A" and [simp]: "finite F" shows "f summable_on (A - F)" proof - from assms(1) obtain x where lim_f: "(sum f \ x) (finite_subsets_at_top A)" unfolding summable_on_def has_sum_def by auto define F' where "F' = F\A" with assms have "finite F'" and "A-F = A-F'" by auto have "filtermap ((\)F') (finite_subsets_at_top (A-F)) \ finite_subsets_at_top A" proof (rule filter_leI) fix P assume "eventually P (finite_subsets_at_top A)" then obtain X where [simp]: "finite X" and XA: "X \ A" and P: "\Y. finite Y \ X \ Y \ Y \ A \ P Y" unfolding eventually_finite_subsets_at_top by auto define X' where "X' = X-F" hence [simp]: "finite X'" and [simp]: "X' \ A-F" using XA by auto hence "finite Y \ X' \ Y \ Y \ A - F \ P (F' \ Y)" for Y using P XA unfolding X'_def using F'_def \finite F'\ by blast thus "eventually P (filtermap ((\) F') (finite_subsets_at_top (A - F)))" unfolding eventually_filtermap eventually_finite_subsets_at_top by (rule_tac x=X' in exI, simp) qed with lim_f have "(sum f \ x) (filtermap ((\)F') (finite_subsets_at_top (A-F)))" using tendsto_mono by blast have "((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A - F))" if "((sum f \ (\) F') \ x) (finite_subsets_at_top (A - F))" using that unfolding o_def by auto hence "((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A-F))" using tendsto_compose_filtermap [symmetric] by (simp add: \(sum f \ x) (filtermap ((\) F') (finite_subsets_at_top (A - F)))\ tendsto_compose_filtermap) have "\Y. finite Y \ Y \ A - F \ sum f (F' \ Y) = sum f F' + sum f Y" by (metis Diff_disjoint Int_Diff \A - F = A - F'\ \finite F'\ inf.orderE sum.union_disjoint) hence "\\<^sub>F x in finite_subsets_at_top (A - F). sum f (F' \ x) = sum f F' + sum f x" unfolding eventually_finite_subsets_at_top using exI [where x = "{}"] by (simp add: \\P. P {} \ \x. P x\) hence "((\G. sum f F' + sum f G) \ x) (finite_subsets_at_top (A-F))" using tendsto_cong [THEN iffD1 , rotated] \((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A - F))\ by fastforce hence "((\G. sum f F' + sum f G) \ sum f F' + (x-sum f F')) (finite_subsets_at_top (A-F))" by simp hence "(sum f \ x - sum f F') (finite_subsets_at_top (A-F))" using tendsto_add_const_iff by blast thus "f summable_on (A - F)" unfolding summable_on_def has_sum_def by auto qed lemma fixes f :: "'a \ 'b::{topological_ab_group_add}" assumes \has_sum f B b\ and \has_sum f A a\ and AB: "A \ B" shows has_sum_Diff: "has_sum f (B - A) (b - a)" proof - have finite_subsets1: "finite_subsets_at_top (B - A) \ filtermap (\F. F - A) (finite_subsets_at_top B)" proof (rule filter_leI) fix P assume "eventually P (filtermap (\F. F - A) (finite_subsets_at_top B))" then obtain X where "finite X" and "X \ B" and P: "finite Y \ X \ Y \ Y \ B \ P (Y - A)" for Y unfolding eventually_filtermap eventually_finite_subsets_at_top by auto hence "finite (X-A)" and "X-A \ B - A" by auto moreover have "finite Y \ X-A \ Y \ Y \ B - A \ P Y" for Y using P[where Y="Y\X"] \finite X\ \X \ B\ by (metis Diff_subset Int_Diff Un_Diff finite_Un inf.orderE le_sup_iff sup.orderE sup_ge2) ultimately show "eventually P (finite_subsets_at_top (B - A))" unfolding eventually_finite_subsets_at_top by meson qed have finite_subsets2: "filtermap (\F. F \ A) (finite_subsets_at_top B) \ finite_subsets_at_top A" apply (rule filter_leI) using assms unfolding eventually_filtermap eventually_finite_subsets_at_top by (metis Int_subset_iff finite_Int inf_le2 subset_trans) from assms(1) have limB: "(sum f \ b) (finite_subsets_at_top B)" using has_sum_def by auto from assms(2) have limA: "(sum f \ a) (finite_subsets_at_top A)" using has_sum_def by blast have "((\F. sum f (F\A)) \ a) (finite_subsets_at_top B)" proof (subst asm_rl [of "(\F. sum f (F\A)) = sum f \ (\F. F\A)"]) show "(\F. sum f (F \ A)) = sum f \ (\F. F \ A)" unfolding o_def by auto show "((sum f \ (\F. F \ A)) \ a) (finite_subsets_at_top B)" unfolding o_def using tendsto_compose_filtermap finite_subsets2 limA tendsto_mono \(\F. sum f (F \ A)) = sum f \ (\F. F \ A)\ by fastforce qed with limB have "((\F. sum f F - sum f (F\A)) \ b - a) (finite_subsets_at_top B)" using tendsto_diff by blast have "sum f X - sum f (X \ A) = sum f (X - A)" if "finite X" and "X \ B" for X :: "'a set" using that by (metis add_diff_cancel_left' sum.Int_Diff) hence "\\<^sub>F x in finite_subsets_at_top B. sum f x - sum f (x \ A) = sum f (x - A)" by (rule eventually_finite_subsets_at_top_weakI) hence "((\F. sum f (F-A)) \ b - a) (finite_subsets_at_top B)" using tendsto_cong [THEN iffD1 , rotated] \((\F. sum f F - sum f (F \ A)) \ b - a) (finite_subsets_at_top B)\ by fastforce hence "(sum f \ b - a) (filtermap (\F. F-A) (finite_subsets_at_top B))" by (subst tendsto_compose_filtermap[symmetric], simp add: o_def) thus ?thesis using finite_subsets1 has_sum_def tendsto_mono by blast qed lemma fixes f :: "'a \ 'b::{topological_ab_group_add}" assumes "f summable_on B" and "f summable_on A" and "A \ B" shows summable_on_Diff: "f summable_on (B-A)" by (meson assms summable_on_def has_sum_Diff) lemma fixes f :: "'a \ 'b::{topological_ab_group_add,t2_space}" assumes "f summable_on B" and "f summable_on A" and AB: "A \ B" shows infsum_Diff: "infsum f (B - A) = infsum f B - infsum f A" by (metis AB assms has_sum_Diff infsumI summable_on_def) lemma has_sum_mono_neutral: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" (* Does this really require a linorder topology? (Instead of order topology.) *) assumes \has_sum f A a\ and "has_sum g B b" assumes \\x. x \ A\B \ f x \ g x\ assumes \\x. x \ A-B \ f x \ 0\ assumes \\x. x \ B-A \ g x \ 0\ shows "a \ b" proof - define f' g' where \f' x = (if x \ A then f x else 0)\ and \g' x = (if x \ B then g x else 0)\ for x have [simp]: \f summable_on A\ \g summable_on B\ using assms(1,2) summable_on_def by auto have \has_sum f' (A\B) a\ by (smt (verit, best) DiffE IntE Un_iff f'_def assms(1) has_sum_cong_neutral) then have f'_lim: \(sum f' \ a) (finite_subsets_at_top (A\B))\ by (meson has_sum_def) have \has_sum g' (A\B) b\ by (smt (verit, best) DiffD1 DiffD2 IntE UnCI g'_def assms(2) has_sum_cong_neutral) then have g'_lim: \(sum g' \ b) (finite_subsets_at_top (A\B))\ using has_sum_def by blast have "\X i. \X \ A \ B; i \ X\ \ f' i \ g' i" using assms by (auto simp: f'_def g'_def) then have \\\<^sub>F x in finite_subsets_at_top (A \ B). sum f' x \ sum g' x\ by (intro eventually_finite_subsets_at_top_weakI sum_mono) then show ?thesis using f'_lim finite_subsets_at_top_neq_bot g'_lim tendsto_le by blast qed lemma infsum_mono_neutral: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" assumes "f summable_on A" and "g summable_on B" assumes \\x. x \ A\B \ f x \ g x\ assumes \\x. x \ A-B \ f x \ 0\ assumes \\x. x \ B-A \ g x \ 0\ shows "infsum f A \ infsum g B" by (rule has_sum_mono_neutral[of f A _ g B _]) (use assms in \auto intro: has_sum_infsum\) lemma has_sum_mono: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" assumes "has_sum f A x" and "has_sum g A y" assumes \\x. x \ A \ f x \ g x\ shows "x \ y" using assms has_sum_mono_neutral by force lemma infsum_mono: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" assumes "f summable_on A" and "g summable_on A" assumes \\x. x \ A \ f x \ g x\ shows "infsum f A \ infsum g A" by (meson assms has_sum_infsum has_sum_mono) lemma has_sum_finite[simp]: assumes "finite F" shows "has_sum f F (sum f F)" using assms by (auto intro: tendsto_Lim simp: finite_subsets_at_top_finite infsum_def has_sum_def principal_eq_bot_iff) lemma summable_on_finite[simp]: fixes f :: \'a \ 'b::{comm_monoid_add,topological_space}\ assumes "finite F" shows "f summable_on F" using assms summable_on_def has_sum_finite by blast lemma infsum_finite[simp]: assumes "finite F" shows "infsum f F = sum f F" using assms by (auto intro: tendsto_Lim simp: finite_subsets_at_top_finite infsum_def principal_eq_bot_iff) lemma has_sum_finite_approximation: fixes f :: "'a \ 'b::{comm_monoid_add,metric_space}" assumes "has_sum f A x" and "\ > 0" shows "\F. finite F \ F \ A \ dist (sum f F) x \ \" proof - have "(sum f \ x) (finite_subsets_at_top A)" by (meson assms(1) has_sum_def) hence *: "\\<^sub>F F in (finite_subsets_at_top A). dist (sum f F) x < \" using assms(2) by (rule tendstoD) thus ?thesis unfolding eventually_finite_subsets_at_top by fastforce qed lemma infsum_finite_approximation: fixes f :: "'a \ 'b::{comm_monoid_add,metric_space}" assumes "f summable_on A" and "\ > 0" shows "\F. finite F \ F \ A \ dist (sum f F) (infsum f A) \ \" proof - from assms have "has_sum f A (infsum f A)" by (simp add: summable_iff_has_sum_infsum) from this and \\ > 0\ show ?thesis by (rule has_sum_finite_approximation) qed lemma abs_summable_summable: fixes f :: \'a \ 'b :: banach\ assumes \f abs_summable_on A\ shows \f summable_on A\ proof - from assms obtain L where lim: \(sum (\x. norm (f x)) \ L) (finite_subsets_at_top A)\ unfolding has_sum_def summable_on_def by blast then have *: \cauchy_filter (filtermap (sum (\x. norm (f x))) (finite_subsets_at_top A))\ by (auto intro!: nhds_imp_cauchy_filter simp: filterlim_def) have \\P. eventually P (finite_subsets_at_top A) \ (\F F'. P F \ P F' \ dist (sum f F) (sum f F') < e)\ if \e>0\ for e proof - define d P where \d = e/4\ and \P F \ finite F \ F \ A \ dist (sum (\x. norm (f x)) F) L < d\ for F then have \d > 0\ by (simp add: d_def that) have ev_P: \eventually P (finite_subsets_at_top A)\ using lim by (auto simp add: P_def[abs_def] \0 < d\ eventually_conj_iff eventually_finite_subsets_at_top_weakI tendsto_iff) moreover have \dist (sum f F1) (sum f F2) < e\ if \P F1\ and \P F2\ for F1 F2 proof - from ev_P obtain F' where \finite F'\ and \F' \ A\ and P_sup_F': \finite F \ F \ F' \ F \ A \ P F\ for F by atomize_elim (simp add: eventually_finite_subsets_at_top) define F where \F = F' \ F1 \ F2\ have \finite F\ and \F \ A\ using F_def P_def[abs_def] that \finite F'\ \F' \ A\ by auto have dist_F: \dist (sum (\x. norm (f x)) F) L < d\ by (metis F_def \F \ A\ P_def P_sup_F' \finite F\ le_supE order_refl) have dist_F_subset: \dist (sum f F) (sum f F') < 2*d\ if F': \F' \ F\ \P F'\ for F' proof - have \dist (sum f F) (sum f F') = norm (sum f (F-F'))\ unfolding dist_norm using \finite F\ F' by (subst sum_diff) auto also have \\ \ norm (\x\F-F'. norm (f x))\ by (rule order.trans[OF sum_norm_le[OF order.refl]]) auto also have \\ = dist (\x\F. norm (f x)) (\x\F'. norm (f x))\ unfolding dist_norm using \finite F\ F' by (subst sum_diff) auto also have \\ < 2 * d\ using dist_F F' unfolding P_def dist_norm real_norm_def by linarith finally show \dist (sum f F) (sum f F') < 2*d\ . qed have \dist (sum f F1) (sum f F2) \ dist (sum f F) (sum f F1) + dist (sum f F) (sum f F2)\ by (rule dist_triangle3) also have \\ < 2 * d + 2 * d\ by (intro add_strict_mono dist_F_subset that) (auto simp: F_def) also have \\ \ e\ by (auto simp: d_def) finally show \dist (sum f F1) (sum f F2) < e\ . qed then show ?thesis using ev_P by blast qed then have \cauchy_filter (filtermap (sum f) (finite_subsets_at_top A))\ by (simp add: cauchy_filter_metric_filtermap) moreover have "complete (UNIV::'b set)" by (meson Cauchy_convergent UNIV_I complete_def convergent_def) ultimately obtain L' where \(sum f \ L') (finite_subsets_at_top A)\ using complete_uniform[where S=UNIV] by (force simp add: filterlim_def) then show ?thesis using summable_on_def has_sum_def by blast qed text \The converse of @{thm [source] abs_summable_summable} does not hold: Consider the Hilbert space of square-summable sequences. Let $e_i$ denote the sequence with 1 in the $i$th position and 0 elsewhere. Let $f(i) := e_i/i$ for $i\geq1$. We have \<^term>\\ f abs_summable_on UNIV\ because $\lVert f(i)\rVert=1/i$ and thus the sum over $\lVert f(i)\rVert$ diverges. On the other hand, we have \<^term>\f summable_on UNIV\; the limit is the sequence with $1/i$ in the $i$th position. (We have not formalized this separating example here because to the best of our knowledge, this Hilbert space has not been formalized in Isabelle/HOL yet.)\ lemma norm_has_sum_bound: fixes f :: "'b \ 'a::real_normed_vector" and A :: "'b set" assumes "has_sum (\x. norm (f x)) A n" assumes "has_sum f A a" shows "norm a \ n" proof - have "norm a \ n + \" if "\>0" for \ proof- have "\F. norm (a - sum f F) \ \ \ finite F \ F \ A" using has_sum_finite_approximation[where A=A and f=f and \="\"] assms \0 < \\ by (metis dist_commute dist_norm) then obtain F where "norm (a - sum f F) \ \" and "finite F" and "F \ A" by (simp add: atomize_elim) hence "norm a \ norm (sum f F) + \" by (metis add.commute diff_add_cancel dual_order.refl norm_triangle_mono) also have "\ \ sum (\x. norm (f x)) F + \" using norm_sum by auto also have "\ \ n + \" proof (intro add_right_mono [OF has_sum_mono_neutral]) show "has_sum (\x. norm (f x)) F (\x\F. norm (f x))" by (simp add: \finite F\) qed (use \F \ A\ assms in auto) finally show ?thesis by assumption qed thus ?thesis using linordered_field_class.field_le_epsilon by blast qed lemma norm_infsum_bound: fixes f :: "'b \ 'a::real_normed_vector" and A :: "'b set" assumes "f abs_summable_on A" shows "norm (infsum f A) \ infsum (\x. norm (f x)) A" proof (cases "f summable_on A") case True have "has_sum (\x. norm (f x)) A (\\<^sub>\x\A. norm (f x))" by (simp add: assms) then show ?thesis by (metis True has_sum_infsum norm_has_sum_bound) next case False obtain t where t_def: "(sum (\x. norm (f x)) \ t) (finite_subsets_at_top A)" using assms unfolding summable_on_def has_sum_def by blast have sumpos: "sum (\x. norm (f x)) X \ 0" for X by (simp add: sum_nonneg) have tgeq0:"t \ 0" proof(rule ccontr) define S::"real set" where "S = {s. s < 0}" assume "\ 0 \ t" hence "t < 0" by simp hence "t \ S" unfolding S_def by blast moreover have "open S" proof- have "closed {s::real. s \ 0}" using Elementary_Topology.closed_sequential_limits[where S = "{s::real. s \ 0}"] by (metis Lim_bounded2 mem_Collect_eq) moreover have "{s::real. s \ 0} = UNIV - S" unfolding S_def by auto ultimately have "closed (UNIV - S)" by simp thus ?thesis by (simp add: Compl_eq_Diff_UNIV open_closed) qed ultimately have "\\<^sub>F X in finite_subsets_at_top A. (\x\X. norm (f x)) \ S" using t_def unfolding tendsto_def by blast hence "\X. (\x\X. norm (f x)) \ S" by (metis (no_types, lifting) eventually_mono filterlim_iff finite_subsets_at_top_neq_bot tendsto_Lim) then obtain X where "(\x\X. norm (f x)) \ S" by blast hence "(\x\X. norm (f x)) < 0" unfolding S_def by auto thus False by (simp add: leD sumpos) qed have "\!h. (sum (\x. norm (f x)) \ h) (finite_subsets_at_top A)" using t_def finite_subsets_at_top_neq_bot tendsto_unique by blast hence "t = (Topological_Spaces.Lim (finite_subsets_at_top A) (sum (\x. norm (f x))))" using t_def unfolding Topological_Spaces.Lim_def by (metis the_equality) hence "Lim (finite_subsets_at_top A) (sum (\x. norm (f x))) \ 0" using tgeq0 by blast thus ?thesis unfolding infsum_def using False by auto qed lemma infsum_tendsto: assumes \f summable_on S\ shows \((\F. sum f F) \ infsum f S) (finite_subsets_at_top S)\ using assms by (simp flip: has_sum_def) lemma has_sum_0: assumes \\x. x\M \ f x = 0\ shows \has_sum f M 0\ by (metis assms finite.intros(1) has_sum_cong has_sum_cong_neutral has_sum_finite sum.neutral_const) lemma summable_on_0: assumes \\x. x\M \ f x = 0\ shows \f summable_on M\ using assms summable_on_def has_sum_0 by blast lemma infsum_0: assumes \\x. x\M \ f x = 0\ shows \infsum f M = 0\ by (metis assms finite_subsets_at_top_neq_bot infsum_def has_sum_0 has_sum_def tendsto_Lim) text \Variants of @{thm [source] infsum_0} etc. suitable as simp-rules\ lemma infsum_0_simp[simp]: \infsum (\_. 0) M = 0\ by (simp_all add: infsum_0) lemma summable_on_0_simp[simp]: \(\_. 0) summable_on M\ by (simp_all add: summable_on_0) lemma has_sum_0_simp[simp]: \has_sum (\_. 0) M 0\ by (simp_all add: has_sum_0) lemma has_sum_add: fixes f g :: "'a \ 'b::{topological_comm_monoid_add}" assumes \has_sum f A a\ assumes \has_sum g A b\ shows \has_sum (\x. f x + g x) A (a + b)\ proof - from assms have lim_f: \(sum f \ a) (finite_subsets_at_top A)\ and lim_g: \(sum g \ b) (finite_subsets_at_top A)\ by (simp_all add: has_sum_def) then have lim: \(sum (\x. f x + g x) \ a + b) (finite_subsets_at_top A)\ unfolding sum.distrib by (rule tendsto_add) then show ?thesis by (simp_all add: has_sum_def) qed lemma summable_on_add: fixes f g :: "'a \ 'b::{topological_comm_monoid_add}" assumes \f summable_on A\ assumes \g summable_on A\ shows \(\x. f x + g x) summable_on A\ by (metis (full_types) assms(1) assms(2) summable_on_def has_sum_add) lemma infsum_add: fixes f g :: "'a \ 'b::{topological_comm_monoid_add, t2_space}" assumes \f summable_on A\ assumes \g summable_on A\ shows \infsum (\x. f x + g x) A = infsum f A + infsum g A\ proof - have \has_sum (\x. f x + g x) A (infsum f A + infsum g A)\ by (simp add: assms(1) assms(2) has_sum_add) then show ?thesis using infsumI by blast qed lemma has_sum_Un_disjoint: fixes f :: "'a \ 'b::topological_comm_monoid_add" assumes "has_sum f A a" assumes "has_sum f B b" assumes disj: "A \ B = {}" shows \has_sum f (A \ B) (a + b)\ proof - define fA fB where \fA x = (if x \ A then f x else 0)\ and \fB x = (if x \ A then f x else 0)\ for x have fA: \has_sum fA (A \ B) a\ by (smt (verit, ccfv_SIG) DiffD1 DiffD2 UnCI fA_def assms(1) has_sum_cong_neutral inf_sup_absorb) have fB: \has_sum fB (A \ B) b\ by (smt (verit, best) DiffD1 DiffD2 IntE Un_iff fB_def assms(2) disj disjoint_iff has_sum_cong_neutral) have fAB: \f x = fA x + fB x\ for x unfolding fA_def fB_def by simp show ?thesis unfolding fAB using fA fB by (rule has_sum_add) qed lemma summable_on_Un_disjoint: fixes f :: "'a \ 'b::topological_comm_monoid_add" assumes "f summable_on A" assumes "f summable_on B" assumes disj: "A \ B = {}" shows \f summable_on (A \ B)\ by (meson assms(1) assms(2) disj summable_on_def has_sum_Un_disjoint) lemma infsum_Un_disjoint: fixes f :: "'a \ 'b::{topological_comm_monoid_add, t2_space}" assumes "f summable_on A" assumes "f summable_on B" assumes disj: "A \ B = {}" shows \infsum f (A \ B) = infsum f A + infsum f B\ by (intro infsumI has_sum_Un_disjoint has_sum_infsum assms) lemma norm_summable_imp_has_sum: fixes f :: "nat \ 'a :: banach" assumes "summable (\n. norm (f n))" and "f sums S" shows "has_sum f (UNIV :: nat set) S" unfolding has_sum_def tendsto_iff eventually_finite_subsets_at_top proof (safe, goal_cases) case (1 \) from assms(1) obtain S' where S': "(\n. norm (f n)) sums S'" by (auto simp: summable_def) with 1 obtain N where N: "\n. n \ N \ \S' - (\i < \" by (auto simp: tendsto_iff eventually_at_top_linorder sums_def dist_norm abs_minus_commute) show ?case proof (rule exI[of _ "{..n. if n \ Y then 0 else f n) sums (S - sum f Y)" by (intro sums_If_finite_set'[OF \f sums S\]) (auto simp: sum_negf) hence "S - sum f Y = (\n. if n \ Y then 0 else f n)" by (simp add: sums_iff) also have "norm \ \ (\n. norm (if n \ Y then 0 else f n))" by (rule summable_norm[OF summable_comparison_test'[OF assms(1)]]) auto also have "\ \ (\n. if n < N then 0 else norm (f n))" using 2 by (intro suminf_le summable_comparison_test'[OF assms(1)]) auto also have "(\n. if n \ {..in. if n < N then 0 else norm (f n)) = S' - (\ii \S' - (\i" by simp also have "\ < \" by (rule N) auto finally show ?case by (simp add: dist_norm norm_minus_commute) qed auto qed lemma norm_summable_imp_summable_on: fixes f :: "nat \ 'a :: banach" assumes "summable (\n. norm (f n))" shows "f summable_on UNIV" using norm_summable_imp_has_sum[OF assms, of "suminf f"] assms by (auto simp: sums_iff summable_on_def dest: summable_norm_cancel) text \The following lemma indeed needs a complete space (as formalized by the premise \<^term>\complete UNIV\). The following two counterexamples show this: \begin{itemize} \item Consider the real vector space $V$ of sequences with finite support, and with the $\ell_2$-norm (sum of squares). Let $e_i$ denote the sequence with a $1$ at position $i$. Let $f : \mathbb Z \to V$ be defined as $f(n) := e_{\lvert n\rvert} / n$ (with $f(0) := 0$). We have that $\sum_{n\in\mathbb Z} f(n) = 0$ (it even converges absolutely). But $\sum_{n\in\mathbb N} f(n)$ does not exist (it would converge against a sequence with infinite support). \item Let $f$ be a positive rational valued function such that $\sum_{x\in B} f(x)$ is $\sqrt 2$ and $\sum_{x\in A} f(x)$ is 1 (over the reals, with $A\subseteq B$). Then $\sum_{x\in B} f(x)$ does not exist over the rationals. But $\sum_{x\in A} f(x)$ exists. \end{itemize} The lemma also requires uniform continuity of the addition. And example of a topological group with continuous but not uniformly continuous addition would be the positive reals with the usual multiplication as the addition. We do not know whether the lemma would also hold for such topological groups.\ lemma summable_on_subset: fixes A B and f :: \'a \ 'b::{ab_group_add, uniform_space}\ assumes \complete (UNIV :: 'b set)\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'b,y). x+y)\ assumes \f summable_on A\ assumes \B \ A\ shows \f summable_on B\ proof - let ?filter_fB = \filtermap (sum f) (finite_subsets_at_top B)\ from \f summable_on A\ obtain S where \(sum f \ S) (finite_subsets_at_top A)\ (is \(sum f \ S) ?filter_A\) using summable_on_def has_sum_def by blast then have cauchy_fA: \cauchy_filter (filtermap (sum f) (finite_subsets_at_top A))\ (is \cauchy_filter ?filter_fA\) by (auto intro!: nhds_imp_cauchy_filter simp: filterlim_def) have \cauchy_filter (filtermap (sum f) (finite_subsets_at_top B))\ proof (unfold cauchy_filter_def, rule filter_leI) fix E :: \('b\'b) \ bool\ assume \eventually E uniformity\ then obtain E' where \eventually E' uniformity\ and E'E'E: \E' (x, y) \ E' (y, z) \ E (x, z)\ for x y z using uniformity_trans by blast obtain D where \eventually D uniformity\ and DE: \D (x, y) \ E' (x+c, y+c)\ for x y c using plus_cont \eventually E' uniformity\ unfolding uniformly_continuous_on_uniformity filterlim_def le_filter_def uniformity_prod_def by (auto simp: case_prod_beta eventually_filtermap eventually_prod_same uniformity_refl) have DE': "E' (x, y)" if "D (x + c, y + c)" for x y c using DE[of "x + c" "y + c" "-c"] that by simp from \eventually D uniformity\ and cauchy_fA have \eventually D (?filter_fA \\<^sub>F ?filter_fA)\ unfolding cauchy_filter_def le_filter_def by simp then obtain P1 P2 where ev_P1: \eventually (\F. P1 (sum f F)) ?filter_A\ and ev_P2: \eventually (\F. P2 (sum f F)) ?filter_A\ and P1P2E: \P1 x \ P2 y \ D (x, y)\ for x y unfolding eventually_prod_filter eventually_filtermap by auto from ev_P1 obtain F1 where F1: \finite F1\ \F1 \ A\ \\F. F\F1 \ finite F \ F\A \ P1 (sum f F)\ by (metis eventually_finite_subsets_at_top) from ev_P2 obtain F2 where F2: \finite F2\ \F2 \ A\ \\F. F\F2 \ finite F \ F\A \ P2 (sum f F)\ by (metis eventually_finite_subsets_at_top) define F0 F0A F0B where \F0 \ F1 \ F2\ and \F0A \ F0 - B\ and \F0B \ F0 \ B\ have [simp]: \finite F0\ \F0 \ A\ using \F1 \ A\ \F2 \ A\ \finite F1\ \finite F2\ unfolding F0_def by blast+ have *: "E' (sum f F1', sum f F2')" if "F1'\F0B" "F2'\F0B" "finite F1'" "finite F2'" "F1'\B" "F2'\B" for F1' F2' proof (intro DE'[where c = "sum f F0A"] P1P2E) have "P1 (sum f (F1' \ F0A))" using that assms F1(1,2) F2(1,2) by (intro F1) (auto simp: F0A_def F0B_def F0_def) thus "P1 (sum f F1' + sum f F0A)" by (subst (asm) sum.union_disjoint) (use that in \auto simp: F0A_def\) next have "P2 (sum f (F2' \ F0A))" using that assms F1(1,2) F2(1,2) by (intro F2) (auto simp: F0A_def F0B_def F0_def) thus "P2 (sum f F2' + sum f F0A)" by (subst (asm) sum.union_disjoint) (use that in \auto simp: F0A_def\) qed show \eventually E (?filter_fB \\<^sub>F ?filter_fB)\ unfolding eventually_prod_filter proof (safe intro!: exI) show "eventually (\x. E' (x, sum f F0B)) (filtermap (sum f) (finite_subsets_at_top B))" and "eventually (\x. E' (sum f F0B, x)) (filtermap (sum f) (finite_subsets_at_top B))" unfolding eventually_filtermap eventually_finite_subsets_at_top by (rule exI[of _ F0B]; use * in \force simp: F0B_def\)+ next show "E (x, y)" if "E' (x, sum f F0B)" and "E' (sum f F0B, y)" for x y using E'E'E that by blast qed qed then obtain x where \?filter_fB \ nhds x\ using cauchy_filter_complete_converges[of ?filter_fB UNIV] \complete (UNIV :: _)\ by (auto simp: filtermap_bot_iff) then have \(sum f \ x) (finite_subsets_at_top B)\ by (auto simp: filterlim_def) then show ?thesis by (auto simp: summable_on_def has_sum_def) qed text \A special case of @{thm [source] summable_on_subset} for Banach spaces with less premises.\ lemma summable_on_subset_banach: fixes A B and f :: \'a \ 'b::banach\ assumes \f summable_on A\ assumes \B \ A\ shows \f summable_on B\ by (rule summable_on_subset[OF _ _ assms]) (auto simp: complete_def convergent_def dest!: Cauchy_convergent) lemma has_sum_empty[simp]: \has_sum f {} 0\ by (meson ex_in_conv has_sum_0) lemma summable_on_empty[simp]: \f summable_on {}\ by auto lemma infsum_empty[simp]: \infsum f {} = 0\ by simp lemma sum_has_sum: fixes f :: "'a \ 'b::topological_comm_monoid_add" assumes finite: \finite A\ assumes conv: \\a. a \ A \ has_sum f (B a) (s a)\ assumes disj: \\a a'. a\A \ a'\A \ a\a' \ B a \ B a' = {}\ shows \has_sum f (\a\A. B a) (sum s A)\ using assms finite conv proof (induction) case empty then show ?case by simp next case (insert x A) have \has_sum f (B x) (s x)\ by (simp add: insert.prems) moreover have IH: \has_sum f (\a\A. B a) (sum s A)\ using insert by simp ultimately have \has_sum f (B x \ (\a\A. B a)) (s x + sum s A)\ using insert by (intro has_sum_Un_disjoint) auto then show ?case using insert.hyps by auto qed lemma summable_on_finite_union_disjoint: fixes f :: "'a \ 'b::topological_comm_monoid_add" assumes finite: \finite A\ assumes conv: \\a. a \ A \ f summable_on (B a)\ assumes disj: \\a a'. a\A \ a'\A \ a\a' \ B a \ B a' = {}\ shows \f summable_on (\a\A. B a)\ using finite conv disj by induction (auto intro!: summable_on_Un_disjoint) lemma sum_infsum: fixes f :: "'a \ 'b::{topological_comm_monoid_add, t2_space}" assumes finite: \finite A\ assumes conv: \\a. a \ A \ f summable_on (B a)\ assumes disj: \\a a'. a\A \ a'\A \ a\a' \ B a \ B a' = {}\ shows \sum (\a. infsum f (B a)) A = infsum f (\a\A. B a)\ by (rule sym, rule infsumI) (use sum_has_sum[of A f B \\a. infsum f (B a)\] assms in auto) text \The lemmas \infsum_comm_additive_general\ and \infsum_comm_additive\ (and variants) below both state that the infinite sum commutes with a continuous additive function. \infsum_comm_additive_general\ is stated more for more general type classes at the expense of a somewhat less compact formulation of the premises. E.g., by avoiding the constant \<^const>\additive\ which introduces an additional sort constraint (group instead of monoid). For example, extended reals (\<^typ>\ereal\, \<^typ>\ennreal\) are not covered by \infsum_comm_additive\.\ lemma has_sum_comm_additive_general: fixes f :: \'b :: {comm_monoid_add,topological_space} \ 'c :: {comm_monoid_add,topological_space}\ assumes f_sum: \\F. finite F \ F \ S \ sum (f \ g) F = f (sum g F)\ \ \Not using \<^const>\additive\ because it would add sort constraint \<^class>\ab_group_add\\ assumes cont: \f \x\ f x\ \ \For \<^class>\t2_space\, this is equivalent to \isCont f x\ by @{thm [source] isCont_def}.\ assumes infsum: \has_sum g S x\ shows \has_sum (f \ g) S (f x)\ proof - have \(sum g \ x) (finite_subsets_at_top S)\ using infsum has_sum_def by blast then have \((f \ sum g) \ f x) (finite_subsets_at_top S)\ by (meson cont filterlim_def tendsto_at_iff_tendsto_nhds tendsto_compose_filtermap tendsto_mono) then have \(sum (f \ g) \ f x) (finite_subsets_at_top S)\ using tendsto_cong f_sum by (simp add: Lim_transform_eventually eventually_finite_subsets_at_top_weakI) then show \has_sum (f \ g) S (f x)\ using has_sum_def by blast qed lemma summable_on_comm_additive_general: fixes f :: \'b :: {comm_monoid_add,topological_space} \ 'c :: {comm_monoid_add,topological_space}\ assumes \\F. finite F \ F \ S \ sum (f \ g) F = f (sum g F)\ \ \Not using \<^const>\additive\ because it would add sort constraint \<^class>\ab_group_add\\ assumes \\x. has_sum g S x \ f \x\ f x\ \ \For \<^class>\t2_space\, this is equivalent to \isCont f x\ by @{thm [source] isCont_def}.\ assumes \g summable_on S\ shows \(f \ g) summable_on S\ by (meson assms summable_on_def has_sum_comm_additive_general has_sum_def infsum_tendsto) lemma infsum_comm_additive_general: fixes f :: \'b :: {comm_monoid_add,t2_space} \ 'c :: {comm_monoid_add,t2_space}\ assumes f_sum: \\F. finite F \ F \ S \ sum (f \ g) F = f (sum g F)\ \ \Not using \<^const>\additive\ because it would add sort constraint \<^class>\ab_group_add\\ assumes \isCont f (infsum g S)\ assumes \g summable_on S\ shows \infsum (f \ g) S = f (infsum g S)\ using assms by (intro infsumI has_sum_comm_additive_general has_sum_infsum) (auto simp: isCont_def) lemma has_sum_comm_additive: fixes f :: \'b :: {ab_group_add,topological_space} \ 'c :: {ab_group_add,topological_space}\ assumes \additive f\ assumes \f \x\ f x\ \ \For \<^class>\t2_space\, this is equivalent to \isCont f x\ by @{thm [source] isCont_def}.\ assumes infsum: \has_sum g S x\ shows \has_sum (f \ g) S (f x)\ using assms by (intro has_sum_comm_additive_general has_sum_infsum) (auto simp: isCont_def additive.sum) lemma summable_on_comm_additive: fixes f :: \'b :: {ab_group_add,t2_space} \ 'c :: {ab_group_add,topological_space}\ assumes \additive f\ assumes \isCont f (infsum g S)\ assumes \g summable_on S\ shows \(f \ g) summable_on S\ by (meson assms(1) assms(2) assms(3) summable_on_def has_sum_comm_additive has_sum_infsum isContD) lemma infsum_comm_additive: fixes f :: \'b :: {ab_group_add,t2_space} \ 'c :: {ab_group_add,t2_space}\ assumes \additive f\ assumes \isCont f (infsum g S)\ assumes \g summable_on S\ shows \infsum (f \ g) S = f (infsum g S)\ by (rule infsum_comm_additive_general; auto simp: assms additive.sum) lemma nonneg_bdd_above_has_sum: fixes f :: \'a \ 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ assumes \bdd_above (sum f ` {F. F\A \ finite F})\ shows \has_sum f A (SUP F\{F. finite F \ F\A}. sum f F)\ proof - have \(sum f \ (SUP F\{F. finite F \ F\A}. sum f F)) (finite_subsets_at_top A)\ proof (rule order_tendstoI) fix a assume \a < (SUP F\{F. finite F \ F\A}. sum f F)\ then obtain F where \a < sum f F\ and \finite F\ and \F \ A\ by (metis (mono_tags, lifting) Collect_cong Collect_empty_eq assms(2) empty_subsetI finite.emptyI less_cSUP_iff mem_Collect_eq) show \\\<^sub>F x in finite_subsets_at_top A. a < sum f x\ unfolding eventually_finite_subsets_at_top proof (rule exI[of _ F], safe) fix Y assume Y: "finite Y" "F \ Y" "Y \ A" have "a < sum f F" by fact also have "\ \ sum f Y" using assms Y by (intro sum_mono2) auto finally show "a < sum f Y" . qed (use \finite F\ \F \ A\ in auto) next fix a assume *: \(SUP F\{F. finite F \ F\A}. sum f F) < a\ have \sum f F < a\ if \F\A\ and \finite F\ for F proof - have "sum f F \ (SUP F\{F. finite F \ F\A}. sum f F)" by (rule cSUP_upper) (use that assms(2) in \auto simp: conj_commute\) also have "\ < a" by fact finally show ?thesis . qed then show \\\<^sub>F x in finite_subsets_at_top A. sum f x < a\ by (rule eventually_finite_subsets_at_top_weakI) qed then show ?thesis using has_sum_def by blast qed lemma nonneg_bdd_above_summable_on: fixes f :: \'a \ 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ assumes \bdd_above (sum f ` {F. F\A \ finite F})\ shows \f summable_on A\ using assms(1) assms(2) summable_on_def nonneg_bdd_above_has_sum by blast lemma nonneg_bdd_above_infsum: fixes f :: \'a \ 'b :: {conditionally_complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ assumes \bdd_above (sum f ` {F. F\A \ finite F})\ shows \infsum f A = (SUP F\{F. finite F \ F\A}. sum f F)\ using assms by (auto intro!: infsumI nonneg_bdd_above_has_sum) lemma nonneg_has_sum_complete: fixes f :: \'a \ 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ shows \has_sum f A (SUP F\{F. finite F \ F\A}. sum f F)\ using assms nonneg_bdd_above_has_sum by blast lemma nonneg_summable_on_complete: fixes f :: \'a \ 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ shows \f summable_on A\ using assms nonneg_bdd_above_summable_on by blast lemma nonneg_infsum_complete: fixes f :: \'a \ 'b :: {complete_linorder, ordered_comm_monoid_add, linorder_topology}\ assumes \\x. x\A \ f x \ 0\ shows \infsum f A = (SUP F\{F. finite F \ F\A}. sum f F)\ using assms nonneg_bdd_above_infsum by blast lemma has_sum_nonneg: fixes f :: "'a \ 'b::{ordered_comm_monoid_add,linorder_topology}" assumes "has_sum f M a" and "\x. x \ M \ 0 \ f x" shows "a \ 0" by (metis (no_types, lifting) DiffD1 assms(1) assms(2) empty_iff has_sum_0 has_sum_mono_neutral order_refl) lemma infsum_nonneg: fixes f :: "'a \ 'b::{ordered_comm_monoid_add,linorder_topology}" assumes "\x. x \ M \ 0 \ f x" shows "infsum f M \ 0" (is "?lhs \ _") by (metis assms has_sum_infsum has_sum_nonneg infsum_not_exists linorder_linear) lemma has_sum_mono2: fixes f :: "'a \ 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" assumes "has_sum f A S" "has_sum f B S'" "A \ B" assumes "\x. x \ B - A \ f x \ 0" shows "S \ S'" proof - have "has_sum f (B - A) (S' - S)" by (rule has_sum_Diff) fact+ hence "S' - S \ 0" by (rule has_sum_nonneg) (use assms(4) in auto) thus ?thesis by (metis add_0 add_mono_thms_linordered_semiring(3) diff_add_cancel) qed lemma infsum_mono2: fixes f :: "'a \ 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" assumes "f summable_on A" "f summable_on B" "A \ B" assumes "\x. x \ B - A \ f x \ 0" shows "infsum f A \ infsum f B" by (rule has_sum_mono2[OF has_sum_infsum has_sum_infsum]) (use assms in auto) lemma finite_sum_le_has_sum: fixes f :: "'a \ 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" assumes "has_sum f A S" "finite B" "B \ A" assumes "\x. x \ A - B \ f x \ 0" shows "sum f B \ S" proof (rule has_sum_mono2) show "has_sum f A S" by fact show "has_sum f B (sum f B)" by (rule has_sum_finite) fact+ qed (use assms in auto) lemma finite_sum_le_infsum: fixes f :: "'a \ 'b::{topological_ab_group_add, ordered_comm_monoid_add,linorder_topology}" assumes "f summable_on A" "finite B" "B \ A" assumes "\x. x \ A - B \ f x \ 0" shows "sum f B \ infsum f A" by (rule finite_sum_le_has_sum[OF has_sum_infsum]) (use assms in auto) lemma has_sum_reindex: assumes \inj_on h A\ shows \has_sum g (h ` A) x \ has_sum (g \ h) A x\ proof - have \has_sum g (h ` A) x \ (sum g \ x) (finite_subsets_at_top (h ` A))\ by (simp add: has_sum_def) also have \\ \ ((\F. sum g (h ` F)) \ x) (finite_subsets_at_top A)\ by (metis assms filterlim_filtermap filtermap_image_finite_subsets_at_top) also have \\ \ (sum (g \ h) \ x) (finite_subsets_at_top A)\ proof (intro tendsto_cong eventually_finite_subsets_at_top_weakI sum.reindex) show "\X. \finite X; X \ A\ \ inj_on h X" using assms subset_inj_on by blast qed also have \\ \ has_sum (g \ h) A x\ by (simp add: has_sum_def) finally show ?thesis . qed lemma summable_on_reindex: assumes \inj_on h A\ shows \g summable_on (h ` A) \ (g \ h) summable_on A\ by (simp add: assms summable_on_def has_sum_reindex) lemma infsum_reindex: assumes \inj_on h A\ shows \infsum g (h ` A) = infsum (g \ h) A\ by (metis (no_types, opaque_lifting) assms finite_subsets_at_top_neq_bot infsum_def summable_on_reindex has_sum_def has_sum_infsum has_sum_reindex tendsto_Lim) lemma summable_on_reindex_bij_betw: assumes "bij_betw g A B" shows "(\x. f (g x)) summable_on A \ f summable_on B" proof - have "inj_on g A" using assms bij_betw_imp_inj_on by blast then have \(\x. f (g x)) summable_on A \ f summable_on g ` A\ by (metis (mono_tags, lifting) comp_apply summable_on_cong summable_on_reindex) also have \\ \ f summable_on B\ using assms bij_betw_imp_surj_on by blast finally show ?thesis . qed lemma infsum_reindex_bij_betw: assumes "bij_betw g A B" shows "infsum (\x. f (g x)) A = infsum f B" proof - have \infsum (\x. f (g x)) A = infsum f (g ` A)\ by (metis (mono_tags, lifting) assms bij_betw_imp_inj_on infsum_cong infsum_reindex o_def) also have \\ = infsum f B\ using assms bij_betw_imp_surj_on by blast finally show ?thesis . qed lemma sum_uniformity: assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'b::{uniform_space,comm_monoid_add},y). x+y)\ assumes \eventually E uniformity\ obtains D where \eventually D uniformity\ and \\M::'a set. \f f' :: 'a \ 'b. card M \ n \ (\m\M. D (f m, f' m)) \ E (sum f M, sum f' M)\ proof (atomize_elim, insert \eventually E uniformity\, induction n arbitrary: E rule:nat_induct) case 0 then show ?case by (metis card_eq_0_iff equals0D le_zero_eq sum.infinite sum.not_neutral_contains_not_neutral uniformity_refl) next case (Suc n) from plus_cont[unfolded uniformly_continuous_on_uniformity filterlim_def le_filter_def, rule_format, OF Suc.prems] obtain D1 D2 where \eventually D1 uniformity\ and \eventually D2 uniformity\ and D1D2E: \D1 (x, y) \ D2 (x', y') \ E (x + x', y + y')\ for x y x' y' apply atomize_elim by (auto simp: eventually_prod_filter case_prod_beta uniformity_prod_def eventually_filtermap) from Suc.IH[OF \eventually D2 uniformity\] obtain D3 where \eventually D3 uniformity\ and D3: \card M \ n \ (\m\M. D3 (f m, f' m)) \ D2 (sum f M, sum f' M)\ for M :: \'a set\ and f f' by metis define D where \D x \ D1 x \ D3 x\ for x have \eventually D uniformity\ using D_def \eventually D1 uniformity\ \eventually D3 uniformity\ eventually_elim2 by blast have \E (sum f M, sum f' M)\ if \card M \ Suc n\ and DM: \\m\M. D (f m, f' m)\ for M :: \'a set\ and f f' proof (cases \card M = 0\) case True then show ?thesis by (metis Suc.prems card_eq_0_iff sum.empty sum.infinite uniformity_refl) next case False with \card M \ Suc n\ obtain N x where \card N \ n\ and \x \ N\ and \M = insert x N\ by (metis card_Suc_eq less_Suc_eq_0_disj less_Suc_eq_le) from DM have \\m. m\N \ D (f m, f' m)\ using \M = insert x N\ by blast with D3[OF \card N \ n\] have D2_N: \D2 (sum f N, sum f' N)\ using D_def by blast from DM have \D (f x, f' x)\ using \M = insert x N\ by blast then have \D1 (f x, f' x)\ by (simp add: D_def) with D2_N have \E (f x + sum f N, f' x + sum f' N)\ using D1D2E by presburger then show \E (sum f M, sum f' M)\ by (metis False \M = insert x N\ \x \ N\ card.infinite finite_insert sum.insert) qed with \eventually D uniformity\ show ?case by auto qed lemma has_sum_Sigma: fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::{comm_monoid_add,uniform_space}\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes summableAB: "has_sum f (Sigma A B) a" assumes summableB: \\x. x\A \ has_sum (\y. f (x, y)) (B x) (b x)\ shows "has_sum b A a" proof - define F FB FA where \F = finite_subsets_at_top (Sigma A B)\ and \FB x = finite_subsets_at_top (B x)\ and \FA = finite_subsets_at_top A\ for x from summableB have sum_b: \(sum (\y. f (x, y)) \ b x) (FB x)\ if \x \ A\ for x using FB_def[abs_def] has_sum_def that by auto from summableAB have sum_S: \(sum f \ a) F\ using F_def has_sum_def by blast have finite_proj: \finite {b| b. (a,b) \ H}\ if \finite H\ for H :: \('a\'b) set\ and a by (metis (no_types, lifting) finite_imageI finite_subset image_eqI mem_Collect_eq snd_conv subsetI that) have \(sum b \ a) FA\ proof (rule tendsto_iff_uniformity[THEN iffD2, rule_format]) fix E :: \('c \ 'c) \ bool\ assume \eventually E uniformity\ then obtain D where D_uni: \eventually D uniformity\ and DDE': \\x y z. D (x, y) \ D (y, z) \ E (x, z)\ by (metis (no_types, lifting) \eventually E uniformity\ uniformity_transE) from sum_S obtain G where \finite G\ and \G \ Sigma A B\ and G_sum: \G \ H \ H \ Sigma A B \ finite H \ D (sum f H, a)\ for H unfolding tendsto_iff_uniformity by (metis (mono_tags, lifting) D_uni F_def eventually_finite_subsets_at_top) have \finite (fst ` G)\ and \fst ` G \ A\ using \finite G\ \G \ Sigma A B\ by auto thm uniformity_prod_def define Ga where \Ga a = {b. (a,b) \ G}\ for a have Ga_fin: \finite (Ga a)\ and Ga_B: \Ga a \ B a\ for a using \finite G\ \G \ Sigma A B\ finite_proj by (auto simp: Ga_def finite_proj) have \E (sum b M, a)\ if \M \ fst ` G\ and \finite M\ and \M \ A\ for M proof - define FMB where \FMB = finite_subsets_at_top (Sigma M B)\ have \eventually (\H. D (\a\M. b a, \(a,b)\H. f (a,b))) FMB\ proof - obtain D' where D'_uni: \eventually D' uniformity\ and \card M' \ card M \ (\m\M'. D' (g m, g' m)) \ D (sum g M', sum g' M')\ for M' :: \'a set\ and g g' using sum_uniformity[OF plus_cont \eventually D uniformity\] by blast then have D'_sum_D: \(\m\M. D' (g m, g' m)) \ D (sum g M, sum g' M)\ for g g' by auto obtain Ha where \Ha a \ Ga a\ and Ha_fin: \finite (Ha a)\ and Ha_B: \Ha a \ B a\ and D'_sum_Ha: \Ha a \ L \ L \ B a \ finite L \ D' (b a, sum (\b. f (a,b)) L)\ if \a \ A\ for a L proof - from sum_b[unfolded tendsto_iff_uniformity, rule_format, OF _ D'_uni[THEN uniformity_sym]] obtain Ha0 where \finite (Ha0 a)\ and \Ha0 a \ B a\ and \Ha0 a \ L \ L \ B a \ finite L \ D' (b a, sum (\b. f (a,b)) L)\ if \a \ A\ for a L unfolding FB_def eventually_finite_subsets_at_top unfolding prod.case by metis moreover define Ha where \Ha a = Ha0 a \ Ga a\ for a ultimately show ?thesis using that[where Ha=Ha] using Ga_fin Ga_B by auto qed have \D (\a\M. b a, \(a,b)\H. f (a,b))\ if \finite H\ and \H \ Sigma M B\ and \H \ Sigma M Ha\ for H proof - define Ha' where \Ha' a = {b| b. (a,b) \ H}\ for a have [simp]: \finite (Ha' a)\ and [simp]: \Ha' a \ Ha a\ and [simp]: \Ha' a \ B a\ if \a \ M\ for a unfolding Ha'_def using \finite H\ \H \ Sigma M B\ \Sigma M Ha \ H\ that finite_proj by auto have \Sigma M Ha' = H\ using that by (auto simp: Ha'_def) then have *: \(\(a,b)\H. f (a,b)) = (\a\M. \b\Ha' a. f (a,b))\ by (simp add: \finite M\ sum.Sigma) have \D' (b a, sum (\b. f (a,b)) (Ha' a))\ if \a \ M\ for a using D'_sum_Ha \M \ A\ that by auto then have \D (\a\M. b a, \a\M. sum (\b. f (a,b)) (Ha' a))\ by (rule_tac D'_sum_D, auto) with * show ?thesis by auto qed moreover have \Sigma M Ha \ Sigma M B\ using Ha_B \M \ A\ by auto ultimately show ?thesis unfolding FMB_def eventually_finite_subsets_at_top by (intro exI[of _ "Sigma M Ha"]) (use Ha_fin that(2,3) in \fastforce intro!: finite_SigmaI\) qed moreover have \eventually (\H. D (\(a,b)\H. f (a,b), a)) FMB\ unfolding FMB_def eventually_finite_subsets_at_top proof (rule exI[of _ G], safe) fix Y assume Y: "finite Y" "G \ Y" "Y \ Sigma M B" have "Y \ Sigma A B" using Y \M \ A\ by blast thus "D (\(a,b)\Y. f (a, b), a)" using G_sum[of Y] Y by auto qed (use \finite G\ \G \ Sigma A B\ that in auto) ultimately have \\\<^sub>F x in FMB. E (sum b M, a)\ by eventually_elim (use DDE' in auto) then show \E (sum b M, a)\ by (rule eventually_const[THEN iffD1, rotated]) (force simp: FMB_def) qed then show \\\<^sub>F x in FA. E (sum b x, a)\ using \finite (fst ` G)\ and \fst ` G \ A\ by (auto intro!: exI[of _ \fst ` G\] simp add: FA_def eventually_finite_subsets_at_top) qed then show ?thesis by (simp add: FA_def has_sum_def) qed lemma summable_on_Sigma: fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::{comm_monoid_add, t2_space, uniform_space}\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes summableAB: "(\(x,y). f x y) summable_on (Sigma A B)" assumes summableB: \\x. x\A \ (f x) summable_on (B x)\ shows \(\x. infsum (f x) (B x)) summable_on A\ proof - from summableAB obtain a where a: \has_sum (\(x,y). f x y) (Sigma A B) a\ using has_sum_infsum by blast from summableB have b: \\x. x\A \ has_sum (f x) (B x) (infsum (f x) (B x))\ by (auto intro!: has_sum_infsum) show ?thesis using plus_cont a b by (auto intro: has_sum_Sigma[where f=\\(x,y). f x y\, simplified] simp: summable_on_def) qed lemma infsum_Sigma: fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::{comm_monoid_add, t2_space, uniform_space}\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes summableAB: "f summable_on (Sigma A B)" assumes summableB: \\x. x\A \ (\y. f (x, y)) summable_on (B x)\ shows "infsum f (Sigma A B) = infsum (\x. infsum (\y. f (x, y)) (B x)) A" proof - from summableAB have a: \has_sum f (Sigma A B) (infsum f (Sigma A B))\ using has_sum_infsum by blast from summableB have b: \\x. x\A \ has_sum (\y. f (x, y)) (B x) (infsum (\y. f (x, y)) (B x))\ by (auto intro!: has_sum_infsum) show ?thesis using plus_cont a b by (auto intro: infsumI[symmetric] has_sum_Sigma simp: summable_on_def) qed lemma infsum_Sigma': fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::{comm_monoid_add, t2_space, uniform_space}\ assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes summableAB: "(\(x,y). f x y) summable_on (Sigma A B)" assumes summableB: \\x. x\A \ (f x) summable_on (B x)\ shows \infsum (\x. infsum (f x) (B x)) A = infsum (\(x,y). f x y) (Sigma A B)\ using infsum_Sigma[of \\(x,y). f x y\ A B] using assms by auto text \A special case of @{thm [source] infsum_Sigma} etc. for Banach spaces. It has less premises.\ lemma fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::banach\ assumes [simp]: "(\(x,y). f x y) summable_on (Sigma A B)" shows infsum_Sigma'_banach: \infsum (\x. infsum (f x) (B x)) A = infsum (\(x,y). f x y) (Sigma A B)\ (is ?thesis1) and summable_on_Sigma_banach: \(\x. infsum (f x) (B x)) summable_on A\ (is ?thesis2) proof - have fsum: \(f x) summable_on (B x)\ if \x \ A\ for x proof - from assms have \(\(x,y). f x y) summable_on (Pair x ` B x)\ by (meson image_subset_iff summable_on_subset_banach mem_Sigma_iff that) then have \((\(x,y). f x y) \ Pair x) summable_on (B x)\ by (metis summable_on_reindex inj_on_def prod.inject) then show ?thesis by (auto simp: o_def) qed show ?thesis1 using fsum assms infsum_Sigma' isUCont_plus by blast show ?thesis2 using fsum assms isUCont_plus summable_on_Sigma by blast qed lemma infsum_Sigma_banach: fixes A :: "'a set" and B :: "'a \ 'b set" and f :: \'a \ 'b \ 'c::banach\ assumes [simp]: "f summable_on (Sigma A B)" shows \infsum (\x. infsum (\y. f (x,y)) (B x)) A = infsum f (Sigma A B)\ using assms by (subst infsum_Sigma'_banach) auto lemma infsum_swap: fixes A :: "'a set" and B :: "'b set" fixes f :: "'a \ 'b \ 'c::{comm_monoid_add,t2_space,uniform_space}" assumes plus_cont: \uniformly_continuous_on UNIV (\(x::'c,y). x+y)\ assumes \(\(x, y). f x y) summable_on (A \ B)\ assumes \\a. a\A \ (f a) summable_on B\ assumes \\b. b\B \ (\a. f a b) summable_on A\ shows \infsum (\x. infsum (\y. f x y) B) A = infsum (\y. infsum (\x. f x y) A) B\ proof - have "(\(x, y). f y x) \ prod.swap summable_on A \ B" by (simp add: assms(2) summable_on_cong) then have fyx: \(\(x, y). f y x) summable_on (B \ A)\ by (metis has_sum_reindex infsum_reindex inj_swap product_swap summable_iff_has_sum_infsum) have \infsum (\x. infsum (\y. f x y) B) A = infsum (\(x,y). f x y) (A \ B)\ using assms infsum_Sigma' by blast also have \\ = infsum (\(x,y). f y x) (B \ A)\ apply (subst product_swap[symmetric]) apply (subst infsum_reindex) using assms by (auto simp: o_def) also have \\ = infsum (\y. infsum (\x. f x y) A) B\ by (smt (verit) fyx assms(1) assms(4) infsum_Sigma' infsum_cong) finally show ?thesis . qed lemma infsum_swap_banach: fixes A :: "'a set" and B :: "'b set" fixes f :: "'a \ 'b \ 'c::banach" assumes \(\(x, y). f x y) summable_on (A \ B)\ shows "infsum (\x. infsum (\y. f x y) B) A = infsum (\y. infsum (\x. f x y) A) B" proof - have \
: \(\(x, y). f y x) summable_on (B \ A)\ by (metis (mono_tags, lifting) assms case_swap inj_swap o_apply product_swap summable_on_cong summable_on_reindex) have \infsum (\x. infsum (\y. f x y) B) A = infsum (\(x,y). f x y) (A \ B)\ using assms infsum_Sigma'_banach by blast also have \\ = infsum (\(x,y). f y x) (B \ A)\ apply (subst product_swap[symmetric]) apply (subst infsum_reindex) using assms by (auto simp: o_def) also have \\ = infsum (\y. infsum (\x. f x y) A) B\ by (metis (mono_tags, lifting) \
infsum_Sigma'_banach infsum_cong) finally show ?thesis . qed lemma nonneg_infsum_le_0D: fixes f :: "'a \ 'b::{topological_ab_group_add,ordered_ab_group_add,linorder_topology}" assumes "infsum f A \ 0" and abs_sum: "f summable_on A" and nneg: "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" proof (rule ccontr) assume \f x \ 0\ have ex: \f summable_on (A-{x})\ by (rule summable_on_cofin_subset) (use assms in auto) have pos: \infsum f (A - {x}) \ 0\ by (rule infsum_nonneg) (use nneg in auto) have [trans]: \x \ y \ y > z \ x > z\ for x y z :: 'b by auto have \infsum f A = infsum f (A-{x}) + infsum f {x}\ by (subst infsum_Un_disjoint[symmetric]) (use assms ex in \auto simp: insert_absorb\) also have \\ \ infsum f {x}\ (is \_ \ \\) using pos by (rule add_increasing) simp also have \\ = f x\ (is \_ = \\) by (subst infsum_finite) auto also have \\ > 0\ using \f x \ 0\ assms(4) nneg by fastforce finally show False using assms by auto qed lemma nonneg_has_sum_le_0D: fixes f :: "'a \ 'b::{topological_ab_group_add,ordered_ab_group_add,linorder_topology}" assumes "has_sum f A a" \a \ 0\ and nneg: "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" by (metis assms(1) assms(2) assms(4) infsumI nonneg_infsum_le_0D summable_on_def nneg) lemma has_sum_cmult_left: fixes f :: "'a \ 'b :: {topological_semigroup_mult, semiring_0}" assumes \has_sum f A a\ shows "has_sum (\x. f x * c) A (a * c)" proof - from assms have \(sum f \ a) (finite_subsets_at_top A)\ using has_sum_def by blast then have \((\F. sum f F * c) \ a * c) (finite_subsets_at_top A)\ by (simp add: tendsto_mult_right) then have \(sum (\x. f x * c) \ a * c) (finite_subsets_at_top A)\ by (metis (mono_tags) tendsto_cong eventually_finite_subsets_at_top_weakI sum_distrib_right) then show ?thesis using infsumI has_sum_def by blast qed lemma infsum_cmult_left: fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, semiring_0}" assumes \c \ 0 \ f summable_on A\ shows "infsum (\x. f x * c) A = infsum f A * c" proof (cases \c=0\) case True then show ?thesis by auto next case False then have \has_sum f A (infsum f A)\ by (simp add: assms) then show ?thesis by (auto intro!: infsumI has_sum_cmult_left) qed lemma summable_on_cmult_left: fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, semiring_0}" assumes \f summable_on A\ shows "(\x. f x * c) summable_on A" using assms summable_on_def has_sum_cmult_left by blast lemma has_sum_cmult_right: fixes f :: "'a \ 'b :: {topological_semigroup_mult, semiring_0}" assumes \has_sum f A a\ shows "has_sum (\x. c * f x) A (c * a)" proof - from assms have \(sum f \ a) (finite_subsets_at_top A)\ using has_sum_def by blast then have \((\F. c * sum f F) \ c * a) (finite_subsets_at_top A)\ by (simp add: tendsto_mult_left) then have \(sum (\x. c * f x) \ c * a) (finite_subsets_at_top A)\ by (metis (mono_tags, lifting) tendsto_cong eventually_finite_subsets_at_top_weakI sum_distrib_left) then show ?thesis using infsumI has_sum_def by blast qed lemma infsum_cmult_right: fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, semiring_0}" assumes \c \ 0 \ f summable_on A\ shows \infsum (\x. c * f x) A = c * infsum f A\ using assms has_sum_cmult_right infsumI summable_iff_has_sum_infsum by fastforce lemma summable_on_cmult_right: fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, semiring_0}" assumes \f summable_on A\ shows "(\x. c * f x) summable_on A" using assms summable_on_def has_sum_cmult_right by blast lemma summable_on_cmult_left': fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, division_ring}" assumes \c \ 0\ shows "(\x. f x * c) summable_on A \ f summable_on A" proof assume \f summable_on A\ then show \(\x. f x * c) summable_on A\ by (rule summable_on_cmult_left) next assume \(\x. f x * c) summable_on A\ then have \(\x. f x * c * inverse c) summable_on A\ by (rule summable_on_cmult_left) then show \f summable_on A\ by (metis (no_types, lifting) assms summable_on_cong mult.assoc mult.right_neutral right_inverse) qed lemma summable_on_cmult_right': fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, division_ring}" assumes \c \ 0\ shows "(\x. c * f x) summable_on A \ f summable_on A" proof assume \f summable_on A\ then show \(\x. c * f x) summable_on A\ by (rule summable_on_cmult_right) next assume \(\x. c * f x) summable_on A\ then have \(\x. inverse c * (c * f x)) summable_on A\ by (rule summable_on_cmult_right) then show \f summable_on A\ by (metis (no_types, lifting) assms summable_on_cong left_inverse mult.assoc mult.left_neutral) qed lemma infsum_cmult_left': fixes f :: "'a \ 'b :: {t2_space, topological_semigroup_mult, division_ring}" shows "infsum (\x. f x * c) A = infsum f A * c" by (metis (full_types) infsum_cmult_left infsum_not_exists mult_eq_0_iff summable_on_cmult_left') lemma infsum_cmult_right': fixes f :: "'a \ 'b :: {t2_space,topological_semigroup_mult,division_ring}" shows "infsum (\x. c * f x) A = c * infsum f A" by (metis (full_types) infsum_cmult_right infsum_not_exists mult_eq_0_iff summable_on_cmult_right') lemma has_sum_constant[simp]: assumes \finite F\ shows \has_sum (\_. c) F (of_nat (card F) * c)\ by (metis assms has_sum_finite sum_constant) lemma infsum_constant[simp]: assumes \finite F\ shows \infsum (\_. c) F = of_nat (card F) * c\ by (simp add: assms) lemma infsum_diverge_constant: \ \This probably does not really need all of \<^class>\archimedean_field\ but Isabelle/HOL has no type class such as, e.g., "archimedean ring".\ fixes c :: \'a::{archimedean_field, comm_monoid_add, linorder_topology, topological_semigroup_mult}\ assumes \infinite A\ and \c \ 0\ shows \\ (\_. c) summable_on A\ proof (rule notI) assume \(\_. c) summable_on A\ then have \(\_. inverse c * c) summable_on A\ by (rule summable_on_cmult_right) then have [simp]: \(\_. 1::'a) summable_on A\ using assms by auto have \infsum (\_. 1) A \ d\ for d :: 'a proof - obtain n :: nat where \of_nat n \ d\ by (meson real_arch_simple) from assms obtain F where \F \ A\ and \finite F\ and \card F = n\ by (meson infinite_arbitrarily_large) note \d \ of_nat n\ also have \of_nat n = infsum (\_. 1::'a) F\ by (simp add: \card F = n\ \finite F\) also have \\ \ infsum (\_. 1::'a) A\ apply (rule infsum_mono_neutral) using \finite F\ \F \ A\ by auto finally show ?thesis . qed then show False by (meson linordered_field_no_ub not_less) qed lemma has_sum_constant_archimedean[simp]: \ \This probably does not really need all of \<^class>\archimedean_field\ but Isabelle/HOL has no type class such as, e.g., "archimedean ring".\ fixes c :: \'a::{archimedean_field, comm_monoid_add, linorder_topology, topological_semigroup_mult}\ shows \infsum (\_. c) A = of_nat (card A) * c\ by (metis infsum_0 infsum_constant infsum_diverge_constant infsum_not_exists sum.infinite sum_constant) lemma has_sum_uminus: fixes f :: \'a \ 'b::topological_ab_group_add\ shows \has_sum (\x. - f x) A a \ has_sum f A (- a)\ by (auto simp add: sum_negf[abs_def] tendsto_minus_cancel_left has_sum_def) lemma summable_on_uminus: fixes f :: \'a \ 'b::topological_ab_group_add\ shows\(\x. - f x) summable_on A \ f summable_on A\ by (metis summable_on_def has_sum_uminus verit_minus_simplify(4)) lemma infsum_uminus: fixes f :: \'a \ 'b::{topological_ab_group_add, t2_space}\ shows \infsum (\x. - f x) A = - infsum f A\ by (metis (full_types) add.inverse_inverse add.inverse_neutral infsumI infsum_def has_sum_infsum has_sum_uminus) lemma has_sum_le_finite_sums: fixes a :: \'a::{comm_monoid_add,topological_space,linorder_topology}\ assumes \has_sum f A a\ assumes \\F. finite F \ F \ A \ sum f F \ b\ shows \a \ b\ by (metis assms eventually_finite_subsets_at_top_weakI finite_subsets_at_top_neq_bot has_sum_def tendsto_upperbound) lemma infsum_le_finite_sums: fixes b :: \'a::{comm_monoid_add,topological_space,linorder_topology}\ assumes \f summable_on A\ assumes \\F. finite F \ F \ A \ sum f F \ b\ shows \infsum f A \ b\ by (meson assms(1) assms(2) has_sum_infsum has_sum_le_finite_sums) lemma summable_on_scaleR_left [intro]: fixes c :: \'a :: real_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "(\x. f x *\<^sub>R c) summable_on A" proof (cases \c = 0\) case False then have "(\y. y *\<^sub>R c) \ f summable_on A" using assms by (auto simp add: scaleR_left.additive_axioms summable_on_comm_additive) then show ?thesis by (metis (mono_tags, lifting) comp_apply summable_on_cong) qed auto lemma summable_on_scaleR_right [intro]: fixes f :: \'a \ 'b :: real_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "(\x. c *\<^sub>R f x) summable_on A" proof (cases \c = 0\) case False then have "(*\<^sub>R) c \ f summable_on A" using assms by (auto simp add: scaleR_right.additive_axioms summable_on_comm_additive) then show ?thesis by (metis (mono_tags, lifting) comp_apply summable_on_cong) qed auto lemma infsum_scaleR_left: fixes c :: \'a :: real_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "infsum (\x. f x *\<^sub>R c) A = infsum f A *\<^sub>R c" proof (cases \c = 0\) case False then have "infsum ((\y. y *\<^sub>R c) \ f) A = infsum f A *\<^sub>R c" using assms by (auto simp add: scaleR_left.additive_axioms infsum_comm_additive) then show ?thesis by (metis (mono_tags, lifting) comp_apply infsum_cong) qed auto lemma infsum_scaleR_right: fixes f :: \'a \ 'b :: real_normed_vector\ shows "infsum (\x. c *\<^sub>R f x) A = c *\<^sub>R infsum f A" proof - consider (summable) \f summable_on A\ | (c0) \c = 0\ | (not_summable) \\ f summable_on A\ \c \ 0\ by auto then show ?thesis proof cases case summable then have "infsum ((*\<^sub>R) c \ f) A = c *\<^sub>R infsum f A" by (auto simp add: scaleR_right.additive_axioms infsum_comm_additive) then show ?thesis by (metis (mono_tags, lifting) comp_apply infsum_cong) next case c0 then show ?thesis by auto next case not_summable have \\ (\x. c *\<^sub>R f x) summable_on A\ proof (rule notI) assume \(\x. c *\<^sub>R f x) summable_on A\ then have \(\x. inverse c *\<^sub>R c *\<^sub>R f x) summable_on A\ using summable_on_scaleR_right by blast with not_summable show False by simp qed then show ?thesis by (simp add: infsum_not_exists not_summable(1)) qed qed lemma infsum_Un_Int: fixes f :: "'a \ 'b::{topological_ab_group_add, t2_space}" assumes "f summable_on A - B" "f summable_on B - A" \f summable_on A \ B\ shows "infsum f (A \ B) = infsum f A + infsum f B - infsum f (A \ B)" proof - have \f summable_on A\ using assms by (metis Diff_Diff_Int Diff_disjoint Un_Diff_Int summable_on_Un_disjoint) have \infsum f (A \ B) = infsum f A + infsum f (B - A)\ by (metis Diff_disjoint Un_Diff_cancel \f summable_on A\ assms(2) infsum_Un_disjoint) moreover have \infsum f (B - A \ A \ B) = infsum f (B - A) + infsum f (A \ B)\ using assms by (metis Int_Diff_disjoint inf_commute infsum_Un_disjoint) ultimately show ?thesis by (metis Un_Diff_Int add_diff_cancel_right' add_diff_eq inf_commute) qed lemma inj_combinator': assumes "x \ F" shows \inj_on (\(g, y). g(x := y)) (Pi\<^sub>E F B \ B x)\ proof - have "inj_on ((\(y, g). g(x := y)) \ prod.swap) (Pi\<^sub>E F B \ B x)" using inj_combinator[of x F B] assms by (intro comp_inj_on) (auto simp: product_swap) thus ?thesis by (simp add: o_def) qed lemma infsum_prod_PiE: \ \See also \infsum_prod_PiE_abs\ below with incomparable premises.\ fixes f :: "'a \ 'b \ 'c :: {comm_monoid_mult, topological_semigroup_mult, division_ring, banach}" assumes finite: "finite A" assumes "\x. x \ A \ f x summable_on B x" assumes "(\g. \x\A. f x (g x)) summable_on (PiE A B)" shows "infsum (\g. \x\A. f x (g x)) (PiE A B) = (\x\A. infsum (f x) (B x))" proof (use finite assms(2-) in induction) case empty then show ?case by auto next case (insert x F) have pi: \Pi\<^sub>E (insert x F) B = (\(g,y). g(x:=y)) ` (Pi\<^sub>E F B \ B x)\ unfolding PiE_insert_eq by (subst swap_product [symmetric]) (simp add: image_image case_prod_unfold) have prod: \(\x'\F. f x' ((p(x:=y)) x')) = (\x'\F. f x' (p x'))\ for p y by (rule prod.cong) (use insert.hyps in auto) have inj: \inj_on (\(g, y). g(x := y)) (Pi\<^sub>E F B \ B x)\ using \x \ F\ by (rule inj_combinator') have summable1: \(\g. \x\insert x F. f x (g x)) summable_on Pi\<^sub>E (insert x F) B\ using insert.prems(2) . also have \Pi\<^sub>E (insert x F) B = (\(g,y). g(x:=y)) ` (Pi\<^sub>E F B \ B x)\ by (simp only: pi) also have "(\g. \x\insert x F. f x (g x)) summable_on \ \ ((\g. \x\insert x F. f x (g x)) \ (\(g,y). g(x:=y))) summable_on (Pi\<^sub>E F B \ B x)" using inj by (rule summable_on_reindex) also have "(\z\F. f z ((g(x := y)) z)) = (\z\F. f z (g z))" for g y using insert.hyps by (intro prod.cong) auto hence "((\g. \x\insert x F. f x (g x)) \ (\(g,y). g(x:=y))) = (\(p, y). f x y * (\x'\F. f x' (p x')))" using insert.hyps by (auto simp: fun_eq_iff cong: prod.cong_simp) finally have summable2: \(\(p, y). f x y * (\x'\F. f x' (p x'))) summable_on Pi\<^sub>E F B \ B x\ . then have \(\p. \\<^sub>\y\B x. f x y * (\x'\F. f x' (p x'))) summable_on Pi\<^sub>E F B\ by (rule summable_on_Sigma_banach) then have \(\p. (\\<^sub>\y\B x. f x y) * (\x'\F. f x' (p x'))) summable_on Pi\<^sub>E F B\ by (metis (mono_tags, lifting) infsum_cmult_left' infsum_cong summable_on_cong) then have summable3: \(\p. (\x'\F. f x' (p x'))) summable_on Pi\<^sub>E F B\ if \(\\<^sub>\y\B x. f x y) \ 0\ using summable_on_cmult_right' that by blast have \(\\<^sub>\g\Pi\<^sub>E (insert x F) B. \x\insert x F. f x (g x)) = (\\<^sub>\(p,y)\Pi\<^sub>E F B \ B x. \x'\insert x F. f x' ((p(x:=y)) x'))\ by (smt (verit, ccfv_SIG) comp_apply infsum_cong infsum_reindex inj pi prod.cong split_def) also have \\ = (\\<^sub>\(p, y)\Pi\<^sub>E F B \ B x. f x y * (\x'\F. f x' ((p(x:=y)) x')))\ using insert.hyps by auto also have \\ = (\\<^sub>\(p, y)\Pi\<^sub>E F B \ B x. f x y * (\x'\F. f x' (p x')))\ using prod by presburger also have \\ = (\\<^sub>\p\Pi\<^sub>E F B. \\<^sub>\y\B x. f x y * (\x'\F. f x' (p x')))\ using infsum_Sigma'_banach summable2 by force also have \\ = (\\<^sub>\y\B x. f x y) * (\\<^sub>\p\Pi\<^sub>E F B. \x'\F. f x' (p x'))\ by (smt (verit) infsum_cmult_left' infsum_cmult_right' infsum_cong) also have \\ = (\x\insert x F. infsum (f x) (B x))\ using insert summable3 by auto finally show ?case by simp qed lemma infsum_prod_PiE_abs: \ \See also @{thm [source] infsum_prod_PiE} above with incomparable premises.\ fixes f :: "'a \ 'b \ 'c :: {banach, real_normed_div_algebra, comm_semiring_1}" assumes finite: "finite A" assumes "\x. x \ A \ f x abs_summable_on B x" shows "infsum (\g. \x\A. f x (g x)) (PiE A B) = (\x\A. infsum (f x) (B x))" proof (use finite assms(2) in induction) case empty then show ?case by auto next case (insert x A) have pi: \Pi\<^sub>E (insert x F) B = (\(g,y). g(x:=y)) ` (Pi\<^sub>E F B \ B x)\ for x F and B :: "'a \ 'b set" unfolding PiE_insert_eq by (subst swap_product [symmetric]) (simp add: image_image case_prod_unfold) have prod: \(\x'\A. f x' ((p(x:=y)) x')) = (\x'\A. f x' (p x'))\ for p y by (rule prod.cong) (use insert.hyps in auto) have inj: \inj_on (\(g, y). g(x := y)) (Pi\<^sub>E A B \ B x)\ using \x \ A\ by (rule inj_combinator') define s where \s x = infsum (\y. norm (f x y)) (B x)\ for x have *: \(\p\P. norm (\x\F. f x (p x))) \ prod s F\ if P: \P \ Pi\<^sub>E F B\ and [simp]: \finite P\ \finite F\ and sum: \\x. x \ F \ f x abs_summable_on B x\ for P F proof - define B' where \B' x = {p x| p. p\P}\ for x have fin_B'[simp]: \finite (B' x)\ for x using that by (auto simp: B'_def) have [simp]: \finite (Pi\<^sub>E F B')\ by (simp add: finite_PiE) have [simp]: \P \ Pi\<^sub>E F B'\ using that by (auto simp: B'_def) have B'B: \B' x \ B x\ if \x \ F\ for x unfolding B'_def using P that by auto have s_bound: \(\y\B' x. norm (f x y)) \ s x\ if \x \ F\ for x by (metis B'B fin_B' finite_sum_le_has_sum has_sum_infsum norm_ge_zero s_def sum that) have \(\p\P. norm (\x\F. f x (p x))) \ (\p\Pi\<^sub>E F B'. norm (\x\F. f x (p x)))\ by (simp add: sum_mono2) also have \\ = (\p\Pi\<^sub>E F B'. \x\F. norm (f x (p x)))\ by (simp add: prod_norm) also have \\ = (\x\F. \y\B' x. norm (f x y))\ proof (use \finite F\ in induction) case empty then show ?case by simp next case (insert x F) have aux: \a = b \ c * a = c * b\ for a b c :: real by auto have inj: \inj_on (\(g, y). g(x := y)) (Pi\<^sub>E F B' \ B' x)\ by (simp add: inj_combinator' insert.hyps) then have \(\p\Pi\<^sub>E (insert x F) B'. \x\insert x F. norm (f x (p x))) = (\(p,y)\Pi\<^sub>E F B' \ B' x. \x'\insert x F. norm (f x' ((p(x := y)) x')))\ by (simp add: pi sum.reindex case_prod_unfold) also have \\ = (\(p,y)\Pi\<^sub>E F B' \ B' x. norm (f x y) * (\x'\F. norm (f x' ((p(x := y)) x'))))\ using insert.hyps by auto also have \\ = (\(p, y)\Pi\<^sub>E F B' \ B' x. norm (f x y) * (\x'\F. norm (f x' (p x'))))\ by (smt (verit) fun_upd_apply insert.hyps(2) prod.cong split_def sum.cong) also have \\ = (\y\B' x. norm (f x y)) * (\p\Pi\<^sub>E F B'. \x'\F. norm (f x' (p x')))\ by (simp add: sum_product sum.swap [of _ "Pi\<^sub>E F B'"] sum.cartesian_product) also have \\ = (\y\B' x. norm (f x y)) * (\x\F. \y\B' x. norm (f x y))\ by (simp add: insert.IH) also have \\ = (\x\insert x F. \y\B' x. norm (f x y))\ using insert.hyps(1) insert.hyps(2) by force finally show ?case . qed also have \\ = (\x\F. \\<^sub>\y\B' x. norm (f x y))\ by auto also have \\ \ (\x\F. s x)\ using s_bound by (simp add: prod_mono sum_nonneg) finally show ?thesis . qed have "bdd_above (sum (\g. norm (\x\insert x A. f x (g x))) ` {F. F \ Pi\<^sub>E (insert x A) B \ finite F})" apply (rule bdd_aboveI) using * insert.hyps insert.prems by blast then have \(\g. \x\insert x A. f x (g x)) abs_summable_on Pi\<^sub>E (insert x A) B\ using nonneg_bdd_above_summable_on by (metis (mono_tags, lifting) Collect_cong norm_ge_zero) also have \Pi\<^sub>E (insert x A) B = (\(g,y). g(x:=y)) ` (Pi\<^sub>E A B \ B x)\ by (simp only: pi) also have "(\g. \x\insert x A. f x (g x)) abs_summable_on \ \ ((\g. \x\insert x A. f x (g x)) \ (\(g,y). g(x:=y))) abs_summable_on (Pi\<^sub>E A B \ B x)" using inj by (subst summable_on_reindex) (auto simp: o_def) also have "(\z\A. f z ((g(x := y)) z)) = (\z\A. f z (g z))" for g y using insert.hyps by (intro prod.cong) auto hence "((\g. \x\insert x A. f x (g x)) \ (\(g,y). g(x:=y))) = (\(p, y). f x y * (\x'\A. f x' (p x')))" using insert.hyps by (auto simp: fun_eq_iff cong: prod.cong_simp) finally have summable2: \(\(p, y). f x y * (\x'\A. f x' (p x'))) abs_summable_on Pi\<^sub>E A B \ B x\ . have \(\\<^sub>\g\Pi\<^sub>E (insert x A) B. \x\insert x A. f x (g x)) = (\\<^sub>\(p,y)\Pi\<^sub>E A B \ B x. \x'\insert x A. f x' ((p(x:=y)) x'))\ using inj by (simp add: pi infsum_reindex o_def case_prod_unfold) also have \\ = (\\<^sub>\(p,y) \ Pi\<^sub>E A B \ B x. f x y * (\x'\A. f x' ((p(x:=y)) x')))\ using insert.hyps by auto also have \\ = (\\<^sub>\(p,y) \ Pi\<^sub>E A B \ B x. f x y * (\x'\A. f x' (p x')))\ using prod by presburger also have \\ = (\\<^sub>\p\Pi\<^sub>E A B. \\<^sub>\y\B x. f x y * (\x'\A. f x' (p x')))\ using abs_summable_summable infsum_Sigma'_banach summable2 by fastforce also have \\ = (\\<^sub>\y\B x. f x y) * (\\<^sub>\p\Pi\<^sub>E A B. \x'\A. f x' (p x'))\ by (smt (verit, best) infsum_cmult_left' infsum_cmult_right' infsum_cong) also have \\ = (\x\insert x A. infsum (f x) (B x))\ by (simp add: insert) finally show ?case by simp qed subsection \Absolute convergence\ lemma abs_summable_countable: assumes \f abs_summable_on A\ shows \countable {x\A. f x \ 0}\ proof - have fin: \finite {x\A. norm (f x) \ t}\ if \t > 0\ for t proof (rule ccontr) assume *: \infinite {x \ A. t \ norm (f x)}\ have \infsum (\x. norm (f x)) A \ b\ for b proof - obtain b' where b': \of_nat b' \ b / t\ by (meson real_arch_simple) from * obtain F where cardF: \card F \ b'\ and \finite F\ and F: \F \ {x \ A. t \ norm (f x)}\ by (meson finite_if_finite_subsets_card_bdd nle_le) have \b \ of_nat b' * t\ using b' \t > 0\ by (simp add: field_simps split: if_splits) also have \\ \ of_nat (card F) * t\ by (simp add: cardF that) also have \\ = sum (\x. t) F\ by simp also have \\ \ sum (\x. norm (f x)) F\ by (metis (mono_tags, lifting) F in_mono mem_Collect_eq sum_mono) also have \\ = infsum (\x. norm (f x)) F\ using \finite F\ by (rule infsum_finite[symmetric]) also have \\ \ infsum (\x. norm (f x)) A\ by (rule infsum_mono_neutral) (use \finite F\ assms F in auto) finally show ?thesis . qed then show False by (meson gt_ex linorder_not_less) qed have \countable (\i\{1..}. {x\A. norm (f x) \ 1/of_nat i})\ by (rule countable_UN) (use fin in \auto intro!: countable_finite\) also have \\ = {x\A. f x \ 0}\ proof safe fix x assume x: "x \ A" "f x \ 0" define i where "i = max 1 (nat (ceiling (1 / norm (f x))))" have "i \ 1" by (simp add: i_def) moreover have "real i \ 1 / norm (f x)" unfolding i_def by linarith hence "1 / real i \ norm (f x)" using \f x \ 0\ by (auto simp: divide_simps mult_ac) ultimately show "x \ (\i\{1..}. {x \ A. 1 / real i \ norm (f x)})" using \x \ A\ by auto qed auto finally show ?thesis . qed (* Logically belongs in the section about reals, but needed as a dependency here *) lemma summable_on_iff_abs_summable_on_real: fixes f :: \'a \ real\ shows \f summable_on A \ f abs_summable_on A\ proof (rule iffI) assume \f summable_on A\ define n A\<^sub>p A\<^sub>n where \n x = norm (f x)\ and \A\<^sub>p = {x\A. f x \ 0}\ and \A\<^sub>n = {x\A. f x < 0}\ for x have A: \A\<^sub>p \ A\<^sub>n = A\ \A\<^sub>p \ A\<^sub>n = {}\ by (auto simp: A\<^sub>p_def A\<^sub>n_def) from \f summable_on A\ have \f summable_on A\<^sub>p\ \f summable_on A\<^sub>n\ using A\<^sub>p_def A\<^sub>n_def summable_on_subset_banach by fastforce+ then have \n summable_on A\<^sub>p\ by (smt (verit) A\<^sub>p_def n_def mem_Collect_eq real_norm_def summable_on_cong) moreover have \n summable_on A\<^sub>n\ by (smt (verit, best) \f summable_on A\<^sub>n\ summable_on_uminus A\<^sub>n_def n_def summable_on_cong mem_Collect_eq real_norm_def) ultimately show \n summable_on A\ using A summable_on_Un_disjoint by blast next show \f abs_summable_on A \ f summable_on A\ using abs_summable_summable by blast qed lemma abs_summable_on_Sigma_iff: shows "f abs_summable_on Sigma A B \ (\x\A. (\y. f (x, y)) abs_summable_on B x) \ ((\x. infsum (\y. norm (f (x, y))) (B x)) abs_summable_on A)" proof (intro iffI conjI ballI) assume asm: \f abs_summable_on Sigma A B\ then have \(\x. infsum (\y. norm (f (x,y))) (B x)) summable_on A\ by (simp add: cond_case_prod_eta summable_on_Sigma_banach) then show \(\x. \\<^sub>\y\B x. norm (f (x, y))) abs_summable_on A\ using summable_on_iff_abs_summable_on_real by force show \(\y. f (x, y)) abs_summable_on B x\ if \x \ A\ for x proof - from asm have \f abs_summable_on Pair x ` B x\ by (simp add: image_subset_iff summable_on_subset_banach that) then show ?thesis by (metis (mono_tags, lifting) o_def inj_on_def summable_on_reindex prod.inject summable_on_cong) qed next assume asm: \(\x\A. (\xa. f (x, xa)) abs_summable_on B x) \ (\x. \\<^sub>\y\B x. norm (f (x, y))) abs_summable_on A\ have \(\xy\F. norm (f xy)) \ (\\<^sub>\x\A. \\<^sub>\y\B x. norm (f (x, y)))\ if \F \ Sigma A B\ and [simp]: \finite F\ for F proof - have [simp]: \(SIGMA x:fst ` F. {y. (x, y) \ F}) = F\ by (auto intro!: set_eqI simp add: Domain.DomainI fst_eq_Domain) have [simp]: \finite {y. (x, y) \ F}\ for x by (metis \finite F\ Range.intros finite_Range finite_subset mem_Collect_eq subsetI) have \(\xy\F. norm (f xy)) = (\x\fst ` F. \y\{y. (x,y)\F}. norm (f (x,y)))\ by (simp add: sum.Sigma) also have \\ = (\\<^sub>\x\fst ` F. \\<^sub>\y\{y. (x,y)\F}. norm (f (x,y)))\ by auto also have \\ \ (\\<^sub>\x\fst ` F. \\<^sub>\y\B x. norm (f (x,y)))\ using asm that(1) by (intro infsum_mono infsum_mono_neutral) auto also have \\ \ (\\<^sub>\x\A. \\<^sub>\y\B x. norm (f (x,y)))\ by (rule infsum_mono_neutral) (use asm that(1) in \auto simp add: infsum_nonneg\) finally show ?thesis . qed then show \f abs_summable_on Sigma A B\ by (intro nonneg_bdd_above_summable_on) (auto simp: bdd_above_def) qed lemma abs_summable_on_comparison_test: assumes "g abs_summable_on A" assumes "\x. x \ A \ norm (f x) \ norm (g x)" shows "f abs_summable_on A" proof (rule nonneg_bdd_above_summable_on) show "bdd_above (sum (\x. norm (f x)) ` {F. F \ A \ finite F})" proof (rule bdd_aboveI2) fix F assume F: "F \ {F. F \ A \ finite F}" have \sum (\x. norm (f x)) F \ sum (\x. norm (g x)) F\ using assms F by (intro sum_mono) auto also have \\ = infsum (\x. norm (g x)) F\ using F by simp also have \\ \ infsum (\x. norm (g x)) A\ proof (rule infsum_mono_neutral) show "g abs_summable_on F" by (rule summable_on_subset_banach[OF assms(1)]) (use F in auto) qed (use F assms in auto) finally show "(\x\F. norm (f x)) \ (\\<^sub>\x\A. norm (g x))" . qed qed auto lemma abs_summable_iff_bdd_above: fixes f :: \'a \ 'b::real_normed_vector\ shows \f abs_summable_on A \ bdd_above (sum (\x. norm (f x)) ` {F. F\A \ finite F})\ proof (rule iffI) assume \f abs_summable_on A\ show \bdd_above (sum (\x. norm (f x)) ` {F. F \ A \ finite F})\ proof (rule bdd_aboveI2) fix F assume F: "F \ {F. F \ A \ finite F}" show "(\x\F. norm (f x)) \ (\\<^sub>\x\A. norm (f x))" by (rule finite_sum_le_infsum) (use \f abs_summable_on A\ F in auto) qed next assume \bdd_above (sum (\x. norm (f x)) ` {F. F\A \ finite F})\ then show \f abs_summable_on A\ by (simp add: nonneg_bdd_above_summable_on) qed lemma abs_summable_product: fixes x :: "'a \ 'b::{real_normed_div_algebra,banach,second_countable_topology}" assumes x2_sum: "(\i. (x i) * (x i)) abs_summable_on A" and y2_sum: "(\i. (y i) * (y i)) abs_summable_on A" shows "(\i. x i * y i) abs_summable_on A" proof (rule nonneg_bdd_above_summable_on) show "bdd_above (sum (\xa. norm (x xa * y xa)) ` {F. F \ A \ finite F})" proof (rule bdd_aboveI2) fix F assume F: \F \ {F. F \ A \ finite F}\ then have r1: "finite F" and b4: "F \ A" by auto have a1: "(\\<^sub>\i\F. norm (x i * x i)) \ (\\<^sub>\i\A. norm (x i * x i))" by (metis (no_types, lifting) b4 infsum_mono2 norm_ge_zero summable_on_subset_banach x2_sum) have "norm (x i * y i) \ norm (x i * x i) + norm (y i * y i)" for i unfolding norm_mult by (smt mult_left_mono mult_nonneg_nonneg mult_right_mono norm_ge_zero) hence "(\i\F. norm (x i * y i)) \ (\i\F. norm (x i * x i) + norm (y i * y i))" by (simp add: sum_mono) also have "\ = (\i\F. norm (x i * x i)) + (\i\F. norm (y i * y i))" by (simp add: sum.distrib) also have "\ = (\\<^sub>\i\F. norm (x i * x i)) + (\\<^sub>\i\F. norm (y i * y i))" by (simp add: \finite F\) also have "\ \ (\\<^sub>\i\A. norm (x i * x i)) + (\\<^sub>\i\A. norm (y i * y i))" using F assms by (intro add_mono infsum_mono2) auto finally show \(\xa\F. norm (x xa * y xa)) \ (\\<^sub>\i\A. norm (x i * x i)) + (\\<^sub>\i\A. norm (y i * y i))\ by simp qed qed auto subsection \Extended reals and nats\ lemma summable_on_ennreal[simp]: \(f::_ \ ennreal) summable_on S\ by (rule nonneg_summable_on_complete) simp lemma summable_on_enat[simp]: \(f::_ \ enat) summable_on S\ by (rule nonneg_summable_on_complete) simp lemma has_sum_superconst_infinite_ennreal: fixes f :: \'a \ ennreal\ assumes geqb: \\x. x \ S \ f x \ b\ assumes b: \b > 0\ assumes \infinite S\ shows "has_sum f S \" proof - have \(sum f \ \) (finite_subsets_at_top S)\ proof (rule order_tendstoI[rotated], simp) fix y :: ennreal assume \y < \\ then have \y / b < \\ by (metis b ennreal_divide_eq_top_iff gr_implies_not_zero infinity_ennreal_def top.not_eq_extremum) then obtain F where \finite F\ and \F \ S\ and cardF: \card F > y / b\ using \infinite S\ by (metis ennreal_Ex_less_of_nat infinite_arbitrarily_large infinity_ennreal_def) moreover have \sum f Y > y\ if \finite Y\ and \F \ Y\ and \Y \ S\ for Y proof - have \y < b * card F\ by (metis \y < \\ b cardF divide_less_ennreal ennreal_mult_eq_top_iff gr_implies_not_zero infinity_ennreal_def mult.commute top.not_eq_extremum) also have \\ \ b * card Y\ by (meson b card_mono less_imp_le mult_left_mono of_nat_le_iff that(1) that(2)) also have \\ = sum (\_. b) Y\ by (simp add: mult.commute) also have \\ \ sum f Y\ using geqb by (meson subset_eq sum_mono that(3)) finally show ?thesis . qed ultimately show \\\<^sub>F x in finite_subsets_at_top S. y < sum f x\ unfolding eventually_finite_subsets_at_top by auto qed then show ?thesis by (simp add: has_sum_def) qed lemma infsum_superconst_infinite_ennreal: fixes f :: \'a \ ennreal\ assumes \\x. x \ S \ f x \ b\ assumes \b > 0\ assumes \infinite S\ shows "infsum f S = \" using assms infsumI has_sum_superconst_infinite_ennreal by blast lemma infsum_superconst_infinite_ereal: fixes f :: \'a \ ereal\ assumes geqb: \\x. x \ S \ f x \ b\ assumes b: \b > 0\ assumes \infinite S\ shows "infsum f S = \" proof - obtain b' where b': \e2ennreal b' = b\ and \b' > 0\ using b by blast have "0 < e2ennreal b" using b' b by (metis dual_order.refl enn2ereal_e2ennreal gr_zeroI order_less_le zero_ennreal.abs_eq) hence *: \infsum (e2ennreal \ f) S = \\ using assms b' by (intro infsum_superconst_infinite_ennreal[where b=b']) (auto intro!: e2ennreal_mono) have \infsum f S = infsum (enn2ereal \ (e2ennreal \ f)) S\ using geqb b by (intro infsum_cong) (fastforce simp: enn2ereal_e2ennreal) also have \\ = enn2ereal \\ using * by (simp add: infsum_comm_additive_general continuous_at_enn2ereal) also have \\ = \\ by simp finally show ?thesis . qed lemma has_sum_superconst_infinite_ereal: fixes f :: \'a \ ereal\ assumes \\x. x \ S \ f x \ b\ assumes \b > 0\ assumes \infinite S\ shows "has_sum f S \" by (metis Infty_neq_0(1) assms infsum_def has_sum_infsum infsum_superconst_infinite_ereal) lemma infsum_superconst_infinite_enat: fixes f :: \'a \ enat\ assumes geqb: \\x. x \ S \ f x \ b\ assumes b: \b > 0\ assumes \infinite S\ shows "infsum f S = \" proof - have \ennreal_of_enat (infsum f S) = infsum (ennreal_of_enat \ f) S\ by (simp flip: infsum_comm_additive_general) also have \\ = \\ by (metis assms(3) b comp_def ennreal_of_enat_0 ennreal_of_enat_le_iff geqb infsum_superconst_infinite_ennreal leD leI) also have \\ = ennreal_of_enat \\ by simp finally show ?thesis by (rule ennreal_of_enat_inj[THEN iffD1]) qed lemma has_sum_superconst_infinite_enat: fixes f :: \'a \ enat\ assumes \\x. x \ S \ f x \ b\ assumes \b > 0\ assumes \infinite S\ shows "has_sum f S \" by (metis assms i0_lb has_sum_infsum infsum_superconst_infinite_enat nonneg_summable_on_complete) text \This lemma helps to relate a real-valued infsum to a supremum over extended nonnegative reals.\ lemma infsum_nonneg_is_SUPREMUM_ennreal: fixes f :: "'a \ real" assumes summable: "f summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "ennreal (infsum f A) = (SUP F\{F. finite F \ F \ A}. (ennreal (sum f F)))" proof - have \
: "\F. \finite F; F \ A\ \ sum (ennreal \ f) F = ennreal (sum f F)" by (metis (mono_tags, lifting) comp_def fnn subsetD sum.cong sum_ennreal) then have \ennreal (infsum f A) = infsum (ennreal \ f) A\ by (simp add: infsum_comm_additive_general local.summable) also have \\ = (SUP F\{F. finite F \ F \ A}. (ennreal (sum f F)))\ by (metis (mono_tags, lifting) \
image_cong mem_Collect_eq nonneg_infsum_complete zero_le) finally show ?thesis . qed text \This lemma helps to related a real-valued infsum to a supremum over extended reals.\ lemma infsum_nonneg_is_SUPREMUM_ereal: fixes f :: "'a \ real" assumes summable: "f summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "ereal (infsum f A) = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))" proof - have "\F. \finite F; F \ A\ \ sum (ereal \ f) F = ereal (sum f F)" by auto then have \ereal (infsum f A) = infsum (ereal \ f) A\ by (simp add: infsum_comm_additive_general local.summable) also have \\ = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))\ by (subst nonneg_infsum_complete) (simp_all add: assms) finally show ?thesis . qed subsection \Real numbers\ text \Most lemmas in the general property section already apply to real numbers. A few ones that are specific to reals are given here.\ lemma infsum_nonneg_is_SUPREMUM_real: fixes f :: "'a \ real" assumes summable: "f summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "infsum f A = (SUP F\{F. finite F \ F \ A}. (sum f F))" proof - have "ereal (infsum f A) = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))" using assms by (rule infsum_nonneg_is_SUPREMUM_ereal) also have "\ = ereal (SUP F\{F. finite F \ F \ A}. (sum f F))" proof (subst ereal_SUP) show "\SUP a\{F. finite F \ F \ A}. ereal (sum f a)\ \ \" using calculation by fastforce show "(SUP F\{F. finite F \ F \ A}. ereal (sum f F)) = (SUP a\{F. finite F \ F \ A}. ereal (sum f a))" by simp qed finally show ?thesis by simp qed lemma has_sum_nonneg_SUPREMUM_real: fixes f :: "'a \ real" assumes "f summable_on A" and "\x. x\A \ f x \ 0" shows "has_sum f A (SUP F\{F. finite F \ F \ A}. (sum f F))" by (metis (mono_tags, lifting) assms has_sum_infsum infsum_nonneg_is_SUPREMUM_real) lemma summable_countable_real: fixes f :: \'a \ real\ assumes \f summable_on A\ shows \countable {x\A. f x \ 0}\ using abs_summable_countable assms summable_on_iff_abs_summable_on_real by blast subsection \Complex numbers\ lemma has_sum_cnj_iff[simp]: fixes f :: \'a \ complex\ shows \has_sum (\x. cnj (f x)) M (cnj a) \ has_sum f M a\ by (simp add: has_sum_def lim_cnj del: cnj_sum add: cnj_sum[symmetric, abs_def, of f]) lemma summable_on_cnj_iff[simp]: "(\i. cnj (f i)) summable_on A \ f summable_on A" by (metis complex_cnj_cnj summable_on_def has_sum_cnj_iff) lemma infsum_cnj[simp]: \infsum (\x. cnj (f x)) M = cnj (infsum f M)\ by (metis complex_cnj_zero infsumI has_sum_cnj_iff infsum_def summable_on_cnj_iff has_sum_infsum) lemma has_sum_Re: assumes "has_sum f M a" shows "has_sum (\x. Re (f x)) M (Re a)" using has_sum_comm_additive[where f=Re] using assms tendsto_Re by (fastforce simp add: o_def additive_def) lemma infsum_Re: assumes "f summable_on M" shows "infsum (\x. Re (f x)) M = Re (infsum f M)" by (simp add: assms has_sum_Re infsumI) lemma summable_on_Re: assumes "f summable_on M" shows "(\x. Re (f x)) summable_on M" by (metis assms has_sum_Re summable_on_def) lemma has_sum_Im: assumes "has_sum f M a" shows "has_sum (\x. Im (f x)) M (Im a)" using has_sum_comm_additive[where f=Im] using assms tendsto_Im by (fastforce simp add: o_def additive_def) lemma infsum_Im: assumes "f summable_on M" shows "infsum (\x. Im (f x)) M = Im (infsum f M)" by (simp add: assms has_sum_Im infsumI) lemma summable_on_Im: assumes "f summable_on M" shows "(\x. Im (f x)) summable_on M" by (metis assms has_sum_Im summable_on_def) lemma nonneg_infsum_le_0D_complex: fixes f :: "'a \ complex" assumes "infsum f A \ 0" and abs_sum: "f summable_on A" and nneg: "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" proof - have \Im (f x) = 0\ using assms(4) less_eq_complex_def nneg by auto moreover have \Re (f x) = 0\ using assms by (auto simp add: summable_on_Re infsum_Re less_eq_complex_def intro: nonneg_infsum_le_0D[where A=A]) ultimately show ?thesis by (simp add: complex_eqI) qed lemma nonneg_has_sum_le_0D_complex: fixes f :: "'a \ complex" assumes "has_sum f A a" and \a \ 0\ and "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" by (metis assms infsumI nonneg_infsum_le_0D_complex summable_on_def) text \The lemma @{thm [source] infsum_mono_neutral} above applies to various linear ordered monoids such as the reals but not to the complex numbers. Thus we have a separate corollary for those:\ lemma infsum_mono_neutral_complex: fixes f :: "'a \ complex" assumes [simp]: "f summable_on A" and [simp]: "g summable_on B" assumes \\x. x \ A\B \ f x \ g x\ assumes \\x. x \ A-B \ f x \ 0\ assumes \\x. x \ B-A \ g x \ 0\ shows \infsum f A \ infsum g B\ proof - have \infsum (\x. Re (f x)) A \ infsum (\x. Re (g x)) B\ by (smt (verit) assms infsum_cong infsum_mono_neutral less_eq_complex_def summable_on_Re zero_complex.simps(1)) then have Re: \Re (infsum f A) \ Re (infsum g B)\ by (metis assms(1-2) infsum_Re) have \infsum (\x. Im (f x)) A = infsum (\x. Im (g x)) B\ by (smt (verit, best) assms(3-5) infsum_cong_neutral less_eq_complex_def zero_complex.simps(2)) then have Im: \Im (infsum f A) = Im (infsum g B)\ by (metis assms(1-2) infsum_Im) from Re Im show ?thesis by (auto simp: less_eq_complex_def) qed lemma infsum_mono_complex: \ \For \<^typ>\real\, @{thm [source] infsum_mono} can be used. But \<^typ>\complex\ does not have the right typeclass.\ fixes f g :: "'a \ complex" assumes f_sum: "f summable_on A" and g_sum: "g summable_on A" assumes leq: "\x. x \ A \ f x \ g x" shows "infsum f A \ infsum g A" by (metis DiffE IntD1 f_sum g_sum infsum_mono_neutral_complex leq) lemma infsum_nonneg_complex: fixes f :: "'a \ complex" assumes "f summable_on M" and "\x. x \ M \ 0 \ f x" shows "infsum f M \ 0" (is "?lhs \ _") by (metis assms infsum_0_simp summable_on_0_simp infsum_mono_complex) lemma infsum_cmod: assumes "f summable_on M" and fnn: "\x. x \ M \ 0 \ f x" shows "infsum (\x. cmod (f x)) M = cmod (infsum f M)" proof - have \complex_of_real (infsum (\x. cmod (f x)) M) = infsum (\x. complex_of_real (cmod (f x))) M\ proof (rule infsum_comm_additive[symmetric, unfolded o_def]) have "(\z. Re (f z)) summable_on M" using assms summable_on_Re by blast also have "?this \ f abs_summable_on M" using fnn by (intro summable_on_cong) (auto simp: less_eq_complex_def cmod_def) finally show \ . qed (auto simp: additive_def) also have \\ = infsum f M\ using fnn cmod_eq_Re complex_is_Real_iff less_eq_complex_def by (force cong: infsum_cong) finally show ?thesis by (metis abs_of_nonneg infsum_def le_less_trans norm_ge_zero norm_infsum_bound norm_of_real not_le order_refl) qed lemma summable_on_iff_abs_summable_on_complex: fixes f :: \'a \ complex\ shows \f summable_on A \ f abs_summable_on A\ proof (rule iffI) assume \f summable_on A\ define i r ni nr n where \i x = Im (f x)\ and \r x = Re (f x)\ and \ni x = norm (i x)\ and \nr x = norm (r x)\ and \n x = norm (f x)\ for x from \f summable_on A\ have \i summable_on A\ by (simp add: i_def[abs_def] summable_on_Im) then have [simp]: \ni summable_on A\ using ni_def[abs_def] summable_on_iff_abs_summable_on_real by force from \f summable_on A\ have \r summable_on A\ by (simp add: r_def[abs_def] summable_on_Re) then have [simp]: \nr summable_on A\ by (metis nr_def summable_on_cong summable_on_iff_abs_summable_on_real) have n_sum: \n x \ nr x + ni x\ for x by (simp add: n_def nr_def ni_def r_def i_def cmod_le) have *: \(\x. nr x + ni x) summable_on A\ by (simp add: summable_on_add) have "bdd_above (sum n ` {F. F \ A \ finite F})" apply (rule bdd_aboveI[where M=\infsum (\x. nr x + ni x) A\]) using * n_sum by (auto simp flip: infsum_finite simp: ni_def nr_def intro!: infsum_mono_neutral) then show \n summable_on A\ by (simp add: n_def nonneg_bdd_above_summable_on) next show \f abs_summable_on A \ f summable_on A\ using abs_summable_summable by blast qed lemma summable_countable_complex: fixes f :: \'a \ complex\ assumes \f summable_on A\ shows \countable {x\A. f x \ 0}\ using abs_summable_countable assms summable_on_iff_abs_summable_on_complex by blast end diff --git a/src/HOL/Analysis/Metric_Arith.thy b/src/HOL/Analysis/Metric_Arith.thy --- a/src/HOL/Analysis/Metric_Arith.thy +++ b/src/HOL/Analysis/Metric_Arith.thy @@ -1,114 +1,114 @@ (* Title: HOL/Analysis/Metric_Arith.thy Author: Maximilian Schäffeler (port from HOL Light) *) chapter \Functional Analysis\ section\<^marker>\tag unimportant\ \A decision procedure for metric spaces\ theory Metric_Arith imports HOL.Real_Vector_Spaces begin text\<^marker>\tag unimportant\ \ A port of the decision procedure ``Formalization of metric spaces in HOL Light'' -@{cite "DBLP:journals/jar/Maggesi18"} for the type class @{class metric_space}, +\<^cite>\"DBLP:journals/jar/Maggesi18"\ for the type class @{class metric_space}, with the \Argo\ solver as backend. \ named_theorems metric_prenex named_theorems metric_nnf named_theorems metric_unfold named_theorems metric_pre_arith lemmas pre_arith_simps = max.bounded_iff max_less_iff_conj le_max_iff_disj less_max_iff_disj simp_thms HOL.eq_commute declare pre_arith_simps [metric_pre_arith] lemmas unfold_simps = Un_iff subset_iff disjoint_iff_not_equal Ball_def Bex_def declare unfold_simps [metric_unfold] declare HOL.nnf_simps(4) [metric_prenex] lemma imp_prenex [metric_prenex]: "\P Q. (\x. P x) \ Q \ \x. (P x \ Q)" "\P Q. P \ (\x. Q x) \ \x. (P \ Q x)" "\P Q. (\x. P x) \ Q \ \x. (P x \ Q)" "\P Q. P \ (\x. Q x) \ \x. (P \ Q x)" by simp_all lemma ex_prenex [metric_prenex]: "\P Q. (\x. P x) \ Q \ \x. (P x \ Q)" "\P Q. P \ (\x. Q x) \ \x. (P \ Q x)" "\P Q. (\x. P x) \ Q \ \x. (P x \ Q)" "\P Q. P \ (\x. Q x) \ \x. (P \ Q x)" "\P. \(\x. P x) \ \x. \P x" by simp_all lemma all_prenex [metric_prenex]: "\P Q. (\x. P x) \ Q \ \x. (P x \ Q)" "\P Q. P \ (\x. Q x) \ \x. (P \ Q x)" "\P Q. (\x. P x) \ Q \ \x. (P x \ Q)" "\P Q. P \ (\x. Q x) \ \x. (P \ Q x)" "\P. \(\x. P x) \ \x. \P x" by simp_all lemma nnf_thms [metric_nnf]: "(\ (P \ Q)) = (\ P \ \ Q)" "(\ (P \ Q)) = (\ P \ \ Q)" "(P \ Q) = (\ P \ Q)" "(P = Q) \ (P \ \ Q) \ (\ P \ Q)" "(\ (P = Q)) \ (\ P \ \ Q) \ (P \ Q)" "(\ \ P) = P" by blast+ lemmas nnf_simps = nnf_thms linorder_not_less linorder_not_le declare nnf_simps[metric_nnf] lemma ball_insert: "(\x\insert a B. P x) = (P a \ (\x\B. P x))" by blast lemma Sup_insert_insert: fixes a::real shows "Sup (insert a (insert b s)) = Sup (insert (max a b) s)" by (simp add: Sup_real_def) lemma real_abs_dist: "\dist x y\ = dist x y" by simp lemma maxdist_thm [THEN HOL.eq_reflection]: assumes "finite s" "x \ s" "y \ s" defines "\a. f a \ \dist x a - dist a y\" shows "dist x y = Sup (f ` s)" proof - have "dist x y \ Sup (f ` s)" proof - have "finite (f ` s)" by (simp add: \finite s\) moreover have "\dist x y - dist y y\ \ f ` s" by (metis \y \ s\ f_def imageI) ultimately show ?thesis using le_cSup_finite by simp qed also have "Sup (f ` s) \ dist x y" using \x \ s\ cSUP_least[of s f] abs_dist_diff_le unfolding f_def by blast finally show ?thesis . qed theorem metric_eq_thm [THEN HOL.eq_reflection]: "x \ s \ y \ s \ x = y \ (\a\s. dist x a = dist y a)" by auto ML_file \metric_arith.ML\ method_setup metric = \ Scan.succeed (SIMPLE_METHOD' o Metric_Arith.metric_arith_tac) \ "prove simple linear statements in metric spaces (\\\<^sub>p fragment)" end diff --git a/src/HOL/Data_Structures/Braun_Tree.thy b/src/HOL/Data_Structures/Braun_Tree.thy --- a/src/HOL/Data_Structures/Braun_Tree.thy +++ b/src/HOL/Data_Structures/Braun_Tree.thy @@ -1,333 +1,333 @@ (* Author: Tobias Nipkow *) section \Braun Trees\ theory Braun_Tree imports "HOL-Library.Tree_Real" begin -text \Braun Trees were studied by Braun and Rem~\cite{BraunRem} -and later Hoogerwoord~\cite{Hoogerwoord}.\ +text \Braun Trees were studied by Braun and Rem~\<^cite>\"BraunRem"\ +and later Hoogerwoord~\<^cite>\"Hoogerwoord"\.\ fun braun :: "'a tree \ bool" where "braun Leaf = True" | "braun (Node l x r) = ((size l = size r \ size l = size r + 1) \ braun l \ braun r)" lemma braun_Node': "braun (Node l x r) = (size r \ size l \ size l \ size r + 1 \ braun l \ braun r)" by auto text \The shape of a Braun-tree is uniquely determined by its size:\ lemma braun_unique: "\ braun (t1::unit tree); braun t2; size t1 = size t2 \ \ t1 = t2" proof (induction t1 arbitrary: t2) case Leaf thus ?case by simp next case (Node l1 _ r1) from Node.prems(3) have "t2 \ Leaf" by auto then obtain l2 x2 r2 where [simp]: "t2 = Node l2 x2 r2" by (meson neq_Leaf_iff) with Node.prems have "size l1 = size l2 \ size r1 = size r2" by auto thus ?case using Node.prems(1,2) Node.IH by auto qed text \Braun trees are almost complete:\ lemma acomplete_if_braun: "braun t \ acomplete t" proof(induction t) case Leaf show ?case by (simp add: acomplete_def) next case (Node l x r) thus ?case using acomplete_Node_if_wbal2 by force qed subsection \Numbering Nodes\ text \We show that a tree is a Braun tree iff a parity-based numbering (\braun_indices\) of nodes yields an interval of numbers.\ fun braun_indices :: "'a tree \ nat set" where "braun_indices Leaf = {}" | "braun_indices (Node l _ r) = {1} \ (*) 2 ` braun_indices l \ Suc ` (*) 2 ` braun_indices r" lemma braun_indices1: "0 \ braun_indices t" by (induction t) auto lemma finite_braun_indices: "finite(braun_indices t)" by (induction t) auto text "One direction:" lemma braun_indices_if_braun: "braun t \ braun_indices t = {1..size t}" proof(induction t) case Leaf thus ?case by simp next have *: "(*) 2 ` {a..b} \ Suc ` (*) 2 ` {a..b} = {2*a..2*b+1}" (is "?l = ?r") for a b proof show "?l \ ?r" by auto next have "\x2\{a..b}. x \ {Suc (2*x2), 2*x2}" if *: "x \ {2*a .. 2*b+1}" for x proof - have "x div 2 \ {a..b}" using * by auto moreover have "x \ {2 * (x div 2), Suc(2 * (x div 2))}" by auto ultimately show ?thesis by blast qed thus "?r \ ?l" by fastforce qed case (Node l x r) hence "size l = size r \ size l = size r + 1" (is "?A \ ?B") by auto thus ?case proof assume ?A with Node show ?thesis by (auto simp: *) next assume ?B with Node show ?thesis by (auto simp: * atLeastAtMostSuc_conv) qed qed text "The other direction is more complicated. The following proof is due to Thomas Sewell." lemma disj_evens_odds: "(*) 2 ` A \ Suc ` (*) 2 ` B = {}" using double_not_eq_Suc_double by auto lemma card_braun_indices: "card (braun_indices t) = size t" proof (induction t) case Leaf thus ?case by simp next case Node thus ?case by(auto simp: UNION_singleton_eq_range finite_braun_indices card_Un_disjoint card_insert_if disj_evens_odds card_image inj_on_def braun_indices1) qed lemma braun_indices_intvl_base_1: assumes bi: "braun_indices t = {m..n}" shows "{m..n} = {1..size t}" proof (cases "t = Leaf") case True then show ?thesis using bi by simp next case False note eqs = eqset_imp_iff[OF bi] from eqs[of 0] have 0: "0 < m" by (simp add: braun_indices1) from eqs[of 1] have 1: "m \ 1" by (cases t; simp add: False) from 0 1 have eq1: "m = 1" by simp from card_braun_indices[of t] show ?thesis by (simp add: bi eq1) qed lemma even_of_intvl_intvl: fixes S :: "nat set" assumes "S = {m..n} \ {i. even i}" shows "\m' n'. S = (\i. i * 2) ` {m'..n'}" apply (rule exI[where x="Suc m div 2"], rule exI[where x="n div 2"]) apply (fastforce simp add: assms mult.commute) done lemma odd_of_intvl_intvl: fixes S :: "nat set" assumes "S = {m..n} \ {i. odd i}" shows "\m' n'. S = Suc ` (\i. i * 2) ` {m'..n'}" proof - have step1: "\m'. S = Suc ` ({m'..n - 1} \ {i. even i})" apply (rule_tac x="if n = 0 then 1 else m - 1" in exI) apply (auto simp: assms image_def elim!: oddE) done thus ?thesis by (metis even_of_intvl_intvl) qed lemma image_int_eq_image: "(\i \ S. f i \ T) \ (f ` S) \ T = f ` S" "(\i \ S. f i \ T) \ (f ` S) \ T = {}" by auto lemma braun_indices1_le: "i \ braun_indices t \ Suc 0 \ i" using braun_indices1 not_less_eq_eq by blast lemma braun_if_braun_indices: "braun_indices t = {1..size t} \ braun t" proof(induction t) case Leaf then show ?case by simp next case (Node l x r) obtain t where t: "t = Node l x r" by simp from Node.prems have eq: "{2 .. size t} = (\i. i * 2) ` braun_indices l \ Suc ` (\i. i * 2) ` braun_indices r" (is "?R = ?S \ ?T") apply clarsimp apply (drule_tac f="\S. S \ {2..}" in arg_cong) apply (simp add: t mult.commute Int_Un_distrib2 image_int_eq_image braun_indices1_le) done then have ST: "?S = ?R \ {i. even i}" "?T = ?R \ {i. odd i}" by (simp_all add: Int_Un_distrib2 image_int_eq_image) from ST have l: "braun_indices l = {1 .. size l}" by (fastforce dest: braun_indices_intvl_base_1 dest!: even_of_intvl_intvl simp: mult.commute inj_image_eq_iff[OF inj_onI]) from ST have r: "braun_indices r = {1 .. size r}" by (fastforce dest: braun_indices_intvl_base_1 dest!: odd_of_intvl_intvl simp: mult.commute inj_image_eq_iff[OF inj_onI]) note STa = ST[THEN eqset_imp_iff, THEN iffD2] note STb = STa[of "size t"] STa[of "size t - 1"] then have sizes: "size l = size r \ size l = size r + 1" apply (clarsimp simp: t l r inj_image_mem_iff[OF inj_onI]) apply (cases "even (size l)"; cases "even (size r)"; clarsimp elim!: oddE; fastforce) done from l r sizes show ?case by (clarsimp simp: Node.IH) qed lemma braun_iff_braun_indices: "braun t \ braun_indices t = {1..size t}" using braun_if_braun_indices braun_indices_if_braun by blast (* An older less appealing proof: lemma Suc0_notin_double: "Suc 0 \ ( * ) 2 ` A" by(auto) lemma zero_in_double_iff: "(0::nat) \ ( * ) 2 ` A \ 0 \ A" by(auto) lemma Suc_in_Suc_image_iff: "Suc n \ Suc ` A \ n \ A" by(auto) lemmas nat_in_image = Suc0_notin_double zero_in_double_iff Suc_in_Suc_image_iff lemma disj_union_eq_iff: "\ L1 \ R2 = {}; L2 \ R1 = {} \ \ L1 \ R1 = L2 \ R2 \ L1 = L2 \ R1 = R2" by blast lemma inj_braun_indices: "braun_indices t1 = braun_indices t2 \ t1 = (t2::unit tree)" proof(induction t1 arbitrary: t2) case Leaf thus ?case using braun_indices.elims by blast next case (Node l1 x1 r1) have "t2 \ Leaf" proof assume "t2 = Leaf" with Node.prems show False by simp qed thus ?case using Node by (auto simp: neq_Leaf_iff insert_ident nat_in_image braun_indices1 disj_union_eq_iff disj_evens_odds inj_image_eq_iff inj_def) qed text \How many even/odd natural numbers are there between m and n?\ lemma card_Icc_even_nat: "card {i \ {m..n::nat}. even i} = (n+1-m + (m+1) mod 2) div 2" (is "?l m n = ?r m n") proof(induction "n+1 - m" arbitrary: n m) case 0 thus ?case by simp next case Suc have "m \ n" using Suc(2) by arith hence "{m..n} = insert m {m+1..n}" by auto hence "?l m n = card {i \ insert m {m+1..n}. even i}" by simp also have "\ = ?r m n" (is "?l = ?r") proof (cases) assume "even m" hence "{i \ insert m {m+1..n}. even i} = insert m {i \ {m+1..n}. even i}" by auto hence "?l = card {i \ {m+1..n}. even i} + 1" by simp also have "\ = (n-m + (m+2) mod 2) div 2 + 1" using Suc(1)[of n "m+1"] Suc(2) by simp also have "\ = ?r" using \even m\ \m \ n\ by auto finally show ?thesis . next assume "odd m" hence "{i \ insert m {m+1..n}. even i} = {i \ {m+1..n}. even i}" by auto hence "?l = card ..." by simp also have "\ = (n-m + (m+2) mod 2) div 2" using Suc(1)[of n "m+1"] Suc(2) by simp also have "\ = ?r" using \odd m\ \m \ n\ even_iff_mod_2_eq_zero[of m] by simp finally show ?thesis . qed finally show ?case . qed lemma card_Icc_odd_nat: "card {i \ {m..n::nat}. odd i} = (n+1-m + m mod 2) div 2" proof - let ?A = "{i \ {m..n}. odd i}" let ?B = "{i \ {m+1..n+1}. even i}" have "card ?A = card (Suc ` ?A)" by (simp add: card_image) also have "Suc ` ?A = ?B" using Suc_le_D by(force simp: image_iff) also have "card ?B = (n+1-m + (m) mod 2) div 2" using card_Icc_even_nat[of "m+1" "n+1"] by simp finally show ?thesis . qed lemma compact_Icc_even: assumes "A = {i \ {m..n}. even i}" shows "A = (\j. 2*(j-1) + m + m mod 2) ` {1..card A}" (is "_ = ?A") proof let ?a = "(n+1-m + (m+1) mod 2) div 2" have "\j \ {1..?a}. i = 2*(j-1) + m + m mod 2" if *: "i \ {m..n}" "even i" for i proof - let ?j = "(i - (m + m mod 2)) div 2 + 1" have "?j \ {1..?a} \ i = 2*(?j-1) + m + m mod 2" using * by(auto simp: mod2_eq_if) presburger+ thus ?thesis by blast qed thus "A \ ?A" using assms by(auto simp: image_iff card_Icc_even_nat simp del: atLeastAtMost_iff) next let ?a = "(n+1-m + (m+1) mod 2) div 2" have 1: "2 * (j - 1) + m + m mod 2 \ {m..n}" if *: "j \ {1..?a}" for j using * by(auto simp: mod2_eq_if) have 2: "even (2 * (j - 1) + m + m mod 2)" for j by presburger show "?A \ A" apply(simp add: assms card_Icc_even_nat del: atLeastAtMost_iff One_nat_def) using 1 2 by blast qed lemma compact_Icc_odd: assumes "B = {i \ {m..n}. odd i}" shows "B = (\i. 2*(i-1) + m + (m+1) mod 2) ` {1..card B}" proof - define A :: " nat set" where "A = Suc ` B" have "A = {i \ {m+1..n+1}. even i}" using Suc_le_D by(force simp add: A_def assms image_iff) from compact_Icc_even[OF this] have "A = Suc ` (\i. 2 * (i - 1) + m + (m + 1) mod 2) ` {1..card A}" by (simp add: image_comp o_def) hence B: "B = (\i. 2 * (i - 1) + m + (m + 1) mod 2) ` {1..card A}" using A_def by (simp add: inj_image_eq_iff) have "card A = card B" by (metis A_def bij_betw_Suc bij_betw_same_card) with B show ?thesis by simp qed lemma even_odd_decomp: assumes "\x \ A. even x" "\x \ B. odd x" "A \ B = {m..n}" shows "(let a = card A; b = card B in a + b = n+1-m \ A = (\i. 2*(i-1) + m + m mod 2) ` {1..a} \ B = (\i. 2*(i-1) + m + (m+1) mod 2) ` {1..b} \ (a = b \ a = b+1 \ even m \ a+1 = b \ odd m))" proof - let ?a = "card A" let ?b = "card B" have "finite A \ finite B" by (metis \A \ B = {m..n}\ finite_Un finite_atLeastAtMost) hence ab: "?a + ?b = Suc n - m" by (metis Int_emptyI assms card_Un_disjoint card_atLeastAtMost) have A: "A = {i \ {m..n}. even i}" using assms by auto hence A': "A = (\i. 2*(i-1) + m + m mod 2) ` {1..?a}" by(rule compact_Icc_even) have B: "B = {i \ {m..n}. odd i}" using assms by auto hence B': "B = (\i. 2*(i-1) + m + (m+1) mod 2) ` {1..?b}" by(rule compact_Icc_odd) have "?a = ?b \ ?a = ?b+1 \ even m \ ?a+1 = ?b \ odd m" apply(simp add: Let_def mod2_eq_if card_Icc_even_nat[of m n, simplified A[symmetric]] card_Icc_odd_nat[of m n, simplified B[symmetric]] split!: if_splits) by linarith with ab A' B' show ?thesis by simp qed lemma braun_if_braun_indices: "braun_indices t = {1..size t} \ braun t" proof(induction t) case Leaf then show ?case by simp next case (Node t1 x2 t2) have 1: "i > 0 \ Suc(Suc(2 * (i - Suc 0))) = 2*i" for i::nat by(simp add: algebra_simps) have 2: "i > 0 \ 2 * (i - Suc 0) + 3 = 2*i + 1" for i::nat by(simp add: algebra_simps) have 3: "( * ) 2 ` braun_indices t1 \ Suc ` ( * ) 2 ` braun_indices t2 = {2..size t1 + size t2 + 1}" using Node.prems by (simp add: insert_ident Icc_eq_insert_lb_nat nat_in_image braun_indices1) thus ?case using Node.IH even_odd_decomp[OF _ _ 3] by(simp add: card_image inj_on_def card_braun_indices Let_def 1 2 inj_image_eq_iff image_comp cong: image_cong_simp) qed *) end \ No newline at end of file diff --git a/src/HOL/Examples/Cantor.thy b/src/HOL/Examples/Cantor.thy --- a/src/HOL/Examples/Cantor.thy +++ b/src/HOL/Examples/Cantor.thy @@ -1,136 +1,136 @@ (* Title: HOL/Examples/Cantor.thy Author: Makarius *) section \Cantor's Theorem\ theory Cantor imports Main begin subsection \Mathematical statement and proof\ text \ Cantor's Theorem states that there is no surjection from a set to its powerset. The proof works by diagonalization. E.g.\ see \<^item> \<^url>\http://mathworld.wolfram.com/CantorDiagonalMethod.html\ \<^item> \<^url>\https://en.wikipedia.org/wiki/Cantor's_diagonal_argument\ \ theorem Cantor: "\f :: 'a \ 'a set. \A. \x. A = f x" proof assume "\f :: 'a \ 'a set. \A. \x. A = f x" then obtain f :: "'a \ 'a set" where *: "\A. \x. A = f x" .. let ?D = "{x. x \ f x}" from * obtain a where "?D = f a" by blast moreover have "a \ ?D \ a \ f a" by blast ultimately show False by blast qed subsection \Automated proofs\ text \ These automated proofs are much shorter, but lack information why and how it works. \ theorem "\f :: 'a \ 'a set. \A. \x. f x = A" by best theorem "\f :: 'a \ 'a set. \A. \x. f x = A" by force subsection \Elementary version in higher-order predicate logic\ text \ The subsequent formulation bypasses set notation of HOL; it uses elementary \\\-calculus and predicate logic, with standard introduction and elimination rules. This also shows that the proof does not require classical reasoning. \ lemma iff_contradiction: assumes *: "\ A \ A" shows False proof (rule notE) show "\ A" proof assume A with * have "\ A" .. from this and \A\ show False .. qed with * show A .. qed theorem Cantor': "\f :: 'a \ 'a \ bool. \A. \x. A = f x" proof assume "\f :: 'a \ 'a \ bool. \A. \x. A = f x" then obtain f :: "'a \ 'a \ bool" where *: "\A. \x. A = f x" .. let ?D = "\x. \ f x x" from * have "\x. ?D = f x" .. then obtain a where "?D = f a" .. then have "?D a \ f a a" by (rule arg_cong) then have "\ f a a \ f a a" . then show False by (rule iff_contradiction) qed subsection \Classic Isabelle/HOL example\ text \ The following treatment of Cantor's Theorem follows the classic example from the early 1990s, e.g.\ see the file \<^verbatim>\92/HOL/ex/set.ML\ in - Isabelle92 or @{cite \\S18.7\ "paulson-isa-book"}. The old tactic scripts + Isabelle92 or \<^cite>\\\S18.7\ in "paulson-isa-book"\. The old tactic scripts synthesize key information of the proof by refinement of schematic goal states. In contrast, the Isar proof needs to say explicitly what is proven. \<^bigskip> Cantor's Theorem states that every set has more subsets than it has elements. It has become a favourite basic example in pure higher-order logic since it is so easily expressed: @{text [display] \\f::\ \ \ \ bool. \S::\ \ bool. \x::\. f x \ S\} Viewing types as sets, \\ \ bool\ represents the powerset of \\\. This version of the theorem states that for every function from \\\ to its powerset, some subset is outside its range. The Isabelle/Isar proofs below uses HOL's set theory, with the type \\ set\ and the operator \range :: (\ \ \) \ \ set\. \ theorem "\S. S \ range (f :: 'a \ 'a set)" proof let ?S = "{x. x \ f x}" show "?S \ range f" proof assume "?S \ range f" then obtain y where "?S = f y" .. then show False proof (rule equalityCE) assume "y \ f y" assume "y \ ?S" then have "y \ f y" .. with \y \ f y\ show ?thesis by contradiction next assume "y \ ?S" assume "y \ f y" then have "y \ ?S" .. with \y \ ?S\ show ?thesis by contradiction qed qed qed text \ How much creativity is required? As it happens, Isabelle can prove this theorem automatically using best-first search. Depth-first search would diverge, but best-first search successfully navigates through the large search space. The context of Isabelle's classical prover contains rules for the relevant constructs of HOL's set theory. \ theorem "\S. S \ range (f :: 'a \ 'a set)" by best end diff --git a/src/HOL/Examples/Knaster_Tarski.thy b/src/HOL/Examples/Knaster_Tarski.thy --- a/src/HOL/Examples/Knaster_Tarski.thy +++ b/src/HOL/Examples/Knaster_Tarski.thy @@ -1,109 +1,109 @@ (* Title: HOL/Examples/Knaster_Tarski.thy Author: Makarius Typical textbook proof example. *) section \Textbook-style reasoning: the Knaster-Tarski Theorem\ theory Knaster_Tarski imports Main begin unbundle lattice_syntax subsection \Prose version\ text \ - According to the textbook @{cite \pages 93--94\ "davey-priestley"}, the + According to the textbook \<^cite>\\pages 93--94\ in "davey-priestley"\, the Knaster-Tarski fixpoint theorem is as follows.\<^footnote>\We have dualized the argument, and tuned the notation a little bit.\ \<^bold>\The Knaster-Tarski Fixpoint Theorem.\ Let \L\ be a complete lattice and \f: L \ L\ an order-preserving map. Then \\{x \ L | f(x) \ x}\ is a fixpoint of \f\. \<^bold>\Proof.\ Let \H = {x \ L | f(x) \ x}\ and \a = \H\. For all \x \ H\ we have \a \ x\, so \f(a) \ f(x) \ x\. Thus \f(a)\ is a lower bound of \H\, whence \f(a) \ a\. We now use this inequality to prove the reverse one (!) and thereby complete the proof that \a\ is a fixpoint. Since \f\ is order-preserving, \f(f(a)) \ f(a)\. This says \f(a) \ H\, so \a \ f(a)\.\ subsection \Formal versions\ text \ The Isar proof below closely follows the original presentation. Virtually all of the prose narration has been rephrased in terms of formal Isar language elements. Just as many textbook-style proofs, there is a strong bias towards forward proof, and several bends in the course of reasoning. \ theorem Knaster_Tarski: fixes f :: "'a::complete_lattice \ 'a" assumes "mono f" shows "\a. f a = a" proof let ?H = "{u. f u \ u}" let ?a = "\?H" show "f ?a = ?a" proof - { fix x assume "x \ ?H" then have "?a \ x" by (rule Inf_lower) with \mono f\ have "f ?a \ f x" .. also from \x \ ?H\ have "\ \ x" .. finally have "f ?a \ x" . } then have "f ?a \ ?a" by (rule Inf_greatest) { also presume "\ \ f ?a" finally (order_antisym) show ?thesis . } from \mono f\ and \f ?a \ ?a\ have "f (f ?a) \ f ?a" .. then have "f ?a \ ?H" .. then show "?a \ f ?a" by (rule Inf_lower) qed qed text \ Above we have used several advanced Isar language elements, such as explicit block structure and weak assumptions. Thus we have mimicked the particular way of reasoning of the original text. In the subsequent version the order of reasoning is changed to achieve structured top-down decomposition of the problem at the outer level, while only the inner steps of reasoning are done in a forward manner. We are certainly more at ease here, requiring only the most basic features of the Isar language. \ theorem Knaster_Tarski': fixes f :: "'a::complete_lattice \ 'a" assumes "mono f" shows "\a. f a = a" proof let ?H = "{u. f u \ u}" let ?a = "\?H" show "f ?a = ?a" proof (rule order_antisym) show "f ?a \ ?a" proof (rule Inf_greatest) fix x assume "x \ ?H" then have "?a \ x" by (rule Inf_lower) with \mono f\ have "f ?a \ f x" .. also from \x \ ?H\ have "\ \ x" .. finally show "f ?a \ x" . qed show "?a \ f ?a" proof (rule Inf_lower) from \mono f\ and \f ?a \ ?a\ have "f (f ?a) \ f ?a" .. then show "f ?a \ ?H" .. qed qed qed end diff --git a/src/HOL/HOLCF/IMP/HoareEx.thy b/src/HOL/HOLCF/IMP/HoareEx.thy --- a/src/HOL/HOLCF/IMP/HoareEx.thy +++ b/src/HOL/HOLCF/IMP/HoareEx.thy @@ -1,33 +1,33 @@ (* Title: HOL/HOLCF/IMP/HoareEx.thy Author: Tobias Nipkow, TUM Copyright 1997 TUM *) section "Correctness of Hoare by Fixpoint Reasoning" theory HoareEx imports Denotational begin text \ An example from the HOLCF paper by Müller, Nipkow, Oheimb, Slotosch - @{cite MuellerNvOS99}. It demonstrates fixpoint reasoning by showing + \<^cite>\MuellerNvOS99\. It demonstrates fixpoint reasoning by showing the correctness of the Hoare rule for while-loops. \ type_synonym assn = "state \ bool" definition hoare_valid :: "[assn, com, assn] \ bool" ("|= {(1_)}/ (_)/ {(1_)}" 50) where "|= {P} c {Q} = (\s t. P s \ D c\(Discr s) = Def t \ Q t)" lemma WHILE_rule_sound: "|= {A} c {A} \ |= {A} WHILE b DO c {\s. A s \ \ bval b s}" apply (unfold hoare_valid_def) apply (simp (no_asm)) apply (rule fix_ind) apply (simp (no_asm)) \ \simplifier with enhanced \adm\-tactic\ apply (simp (no_asm)) apply (simp (no_asm)) apply blast done end diff --git a/src/HOL/Hahn_Banach/Hahn_Banach.thy b/src/HOL/Hahn_Banach/Hahn_Banach.thy --- a/src/HOL/Hahn_Banach/Hahn_Banach.thy +++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy @@ -1,493 +1,493 @@ (* Title: HOL/Hahn_Banach/Hahn_Banach.thy Author: Gertrud Bauer, TU Munich *) section \The Hahn-Banach Theorem\ theory Hahn_Banach imports Hahn_Banach_Lemmas begin text \ We present the proof of two different versions of the Hahn-Banach Theorem, - closely following @{cite \\S36\ "Heuser:1986"}. + closely following \<^cite>\\\S36\ in "Heuser:1986"\. \ subsection \The Hahn-Banach Theorem for vector spaces\ paragraph \Hahn-Banach Theorem.\ text \ Let \F\ be a subspace of a real vector space \E\, let \p\ be a semi-norm on \E\, and \f\ be a linear form defined on \F\ such that \f\ is bounded by \p\, i.e. \\x \ F. f x \ p x\. Then \f\ can be extended to a linear form \h\ on \E\ such that \h\ is norm-preserving, i.e. \h\ is also bounded by \p\. \ paragraph \Proof Sketch.\ text \ \<^enum> Define \M\ as the set of norm-preserving extensions of \f\ to subspaces of \E\. The linear forms in \M\ are ordered by domain extension. \<^enum> We show that every non-empty chain in \M\ has an upper bound in \M\. \<^enum> With Zorn's Lemma we conclude that there is a maximal function \g\ in \M\. \<^enum> The domain \H\ of \g\ is the whole space \E\, as shown by classical contradiction: \<^item> Assuming \g\ is not defined on whole \E\, it can still be extended in a norm-preserving way to a super-space \H'\ of \H\. \<^item> Thus \g\ can not be maximal. Contradiction! \ theorem Hahn_Banach: assumes E: "vectorspace E" and "subspace F E" and "seminorm E p" and "linearform F f" assumes fp: "\x \ F. f x \ p x" shows "\h. linearform E h \ (\x \ F. h x = f x) \ (\x \ E. h x \ p x)" \ \Let \E\ be a vector space, \F\ a subspace of \E\, \p\ a seminorm on \E\,\ \ \and \f\ a linear form on \F\ such that \f\ is bounded by \p\,\ \ \then \f\ can be extended to a linear form \h\ on \E\ in a norm-preserving way. \<^smallskip>\ proof - interpret vectorspace E by fact interpret subspace F E by fact interpret seminorm E p by fact interpret linearform F f by fact define M where "M = norm_pres_extensions E p F f" then have M: "M = \" by (simp only:) from E have F: "vectorspace F" .. note FE = \F \ E\ { fix c assume cM: "c \ chains M" and ex: "\x. x \ c" have "\c \ M" \ \Show that every non-empty chain \c\ of \M\ has an upper bound in \M\:\ \ \\\c\ is greater than any element of the chain \c\, so it suffices to show \\c \ M\.\ unfolding M_def proof (rule norm_pres_extensionI) let ?H = "domain (\c)" let ?h = "funct (\c)" have a: "graph ?H ?h = \c" proof (rule graph_domain_funct) fix x y z assume "(x, y) \ \c" and "(x, z) \ \c" with M_def cM show "z = y" by (rule sup_definite) qed moreover from M cM a have "linearform ?H ?h" by (rule sup_lf) moreover from a M cM ex FE E have "?H \ E" by (rule sup_subE) moreover from a M cM ex FE have "F \ ?H" by (rule sup_supF) moreover from a M cM ex have "graph F f \ graph ?H ?h" by (rule sup_ext) moreover from a M cM have "\x \ ?H. ?h x \ p x" by (rule sup_norm_pres) ultimately show "\H h. \c = graph H h \ linearform H h \ H \ E \ F \ H \ graph F f \ graph H h \ (\x \ H. h x \ p x)" by blast qed } then have "\g \ M. \x \ M. g \ x \ x = g" \ \With Zorn's Lemma we can conclude that there is a maximal element in \M\. \<^smallskip>\ proof (rule Zorn's_Lemma) \ \We show that \M\ is non-empty:\ show "graph F f \ M" unfolding M_def proof (rule norm_pres_extensionI2) show "linearform F f" by fact show "F \ E" by fact from F show "F \ F" by (rule vectorspace.subspace_refl) show "graph F f \ graph F f" .. show "\x\F. f x \ p x" by fact qed qed then obtain g where gM: "g \ M" and gx: "\x \ M. g \ x \ g = x" by blast from gM obtain H h where g_rep: "g = graph H h" and linearform: "linearform H h" and HE: "H \ E" and FH: "F \ H" and graphs: "graph F f \ graph H h" and hp: "\x \ H. h x \ p x" unfolding M_def .. \ \\g\ is a norm-preserving extension of \f\, in other words:\ \ \\g\ is the graph of some linear form \h\ defined on a subspace \H\ of \E\,\ \ \and \h\ is an extension of \f\ that is again bounded by \p\. \<^smallskip>\ from HE E have H: "vectorspace H" by (rule subspace.vectorspace) have HE_eq: "H = E" \ \We show that \h\ is defined on whole \E\ by classical contradiction. \<^smallskip>\ proof (rule classical) assume neq: "H \ E" \ \Assume \h\ is not defined on whole \E\. Then show that \h\ can be extended\ \ \in a norm-preserving way to a function \h'\ with the graph \g'\. \<^smallskip>\ have "\g' \ M. g \ g' \ g \ g'" proof - from HE have "H \ E" .. with neq obtain x' where x'E: "x' \ E" and "x' \ H" by blast obtain x': "x' \ 0" proof show "x' \ 0" proof assume "x' = 0" with H have "x' \ H" by (simp only: vectorspace.zero) with \x' \ H\ show False by contradiction qed qed define H' where "H' = H + lin x'" \ \Define \H'\ as the direct sum of \H\ and the linear closure of \x'\. \<^smallskip>\ have HH': "H \ H'" proof (unfold H'_def) from x'E have "vectorspace (lin x')" .. with H show "H \ H + lin x'" .. qed obtain xi where xi: "\y \ H. - p (y + x') - h y \ xi \ xi \ p (y + x') - h y" \ \Pick a real number \\\ that fulfills certain inequality; this will\ \ \be used to establish that \h'\ is a norm-preserving extension of \h\. \label{ex-xi-use}\<^smallskip>\ proof - from H have "\xi. \y \ H. - p (y + x') - h y \ xi \ xi \ p (y + x') - h y" proof (rule ex_xi) fix u v assume u: "u \ H" and v: "v \ H" with HE have uE: "u \ E" and vE: "v \ E" by auto from H u v linearform have "h v - h u = h (v - u)" by (simp add: linearform.diff) also from hp and H u v have "\ \ p (v - u)" by (simp only: vectorspace.diff_closed) also from x'E uE vE have "v - u = x' + - x' + v + - u" by (simp add: diff_eq1) also from x'E uE vE have "\ = v + x' + - (u + x')" by (simp add: add_ac) also from x'E uE vE have "\ = (v + x') - (u + x')" by (simp add: diff_eq1) also from x'E uE vE E have "p \ \ p (v + x') + p (u + x')" by (simp add: diff_subadditive) finally have "h v - h u \ p (v + x') + p (u + x')" . then show "- p (u + x') - h u \ p (v + x') - h v" by simp qed then show thesis by (blast intro: that) qed define h' where "h' x = (let (y, a) = SOME (y, a). x = y + a \ x' \ y \ H in h y + a * xi)" for x \ \Define the extension \h'\ of \h\ to \H'\ using \\\. \<^smallskip>\ have "g \ graph H' h' \ g \ graph H' h'" \ \\h'\ is an extension of \h\ \dots \<^smallskip>\ proof show "g \ graph H' h'" proof - have "graph H h \ graph H' h'" proof (rule graph_extI) fix t assume t: "t \ H" from E HE t have "(SOME (y, a). t = y + a \ x' \ y \ H) = (t, 0)" using \x' \ H\ \x' \ E\ \x' \ 0\ by (rule decomp_H'_H) with h'_def show "h t = h' t" by (simp add: Let_def) next from HH' show "H \ H'" .. qed with g_rep show ?thesis by (simp only:) qed show "g \ graph H' h'" proof - have "graph H h \ graph H' h'" proof assume eq: "graph H h = graph H' h'" have "x' \ H'" unfolding H'_def proof from H show "0 \ H" by (rule vectorspace.zero) from x'E show "x' \ lin x'" by (rule x_lin_x) from x'E show "x' = 0 + x'" by simp qed then have "(x', h' x') \ graph H' h'" .. with eq have "(x', h' x') \ graph H h" by (simp only:) then have "x' \ H" .. with \x' \ H\ show False by contradiction qed with g_rep show ?thesis by simp qed qed moreover have "graph H' h' \ M" \ \and \h'\ is norm-preserving. \<^smallskip>\ proof (unfold M_def) show "graph H' h' \ norm_pres_extensions E p F f" proof (rule norm_pres_extensionI2) show "linearform H' h'" using h'_def H'_def HE linearform \x' \ H\ \x' \ E\ \x' \ 0\ E by (rule h'_lf) show "H' \ E" unfolding H'_def proof show "H \ E" by fact show "vectorspace E" by fact from x'E show "lin x' \ E" .. qed from H \F \ H\ HH' show FH': "F \ H'" by (rule vectorspace.subspace_trans) show "graph F f \ graph H' h'" proof (rule graph_extI) fix x assume x: "x \ F" with graphs have "f x = h x" .. also have "\ = h x + 0 * xi" by simp also have "\ = (let (y, a) = (x, 0) in h y + a * xi)" by (simp add: Let_def) also have "(x, 0) = (SOME (y, a). x = y + a \ x' \ y \ H)" using E HE proof (rule decomp_H'_H [symmetric]) from FH x show "x \ H" .. from x' show "x' \ 0" . show "x' \ H" by fact show "x' \ E" by fact qed also have "(let (y, a) = (SOME (y, a). x = y + a \ x' \ y \ H) in h y + a * xi) = h' x" by (simp only: h'_def) finally show "f x = h' x" . next from FH' show "F \ H'" .. qed show "\x \ H'. h' x \ p x" using h'_def H'_def \x' \ H\ \x' \ E\ \x' \ 0\ E HE \seminorm E p\ linearform and hp xi by (rule h'_norm_pres) qed qed ultimately show ?thesis .. qed then have "\ (\x \ M. g \ x \ g = x)" by simp \ \So the graph \g\ of \h\ cannot be maximal. Contradiction! \<^smallskip>\ with gx show "H = E" by contradiction qed from HE_eq and linearform have "linearform E h" by (simp only:) moreover have "\x \ F. h x = f x" proof fix x assume "x \ F" with graphs have "f x = h x" .. then show "h x = f x" .. qed moreover from HE_eq and hp have "\x \ E. h x \ p x" by (simp only:) ultimately show ?thesis by blast qed subsection \Alternative formulation\ text \ The following alternative formulation of the Hahn-Banach Theorem\label{abs-Hahn-Banach} uses the fact that for a real linear form \f\ and a seminorm \p\ the following inequality are equivalent:\footnote{This was shown in lemma @{thm [source] abs_ineq_iff} (see page \pageref{abs-ineq-iff}).} \begin{center} \begin{tabular}{lll} \\x \ H. \h x\ \ p x\ & and & \\x \ H. h x \ p x\ \\ \end{tabular} \end{center} \ theorem abs_Hahn_Banach: assumes E: "vectorspace E" and FE: "subspace F E" and lf: "linearform F f" and sn: "seminorm E p" assumes fp: "\x \ F. \f x\ \ p x" shows "\g. linearform E g \ (\x \ F. g x = f x) \ (\x \ E. \g x\ \ p x)" proof - interpret vectorspace E by fact interpret subspace F E by fact interpret linearform F f by fact interpret seminorm E p by fact have "\g. linearform E g \ (\x \ F. g x = f x) \ (\x \ E. g x \ p x)" using E FE sn lf proof (rule Hahn_Banach) show "\x \ F. f x \ p x" using FE E sn lf and fp by (rule abs_ineq_iff [THEN iffD1]) qed then obtain g where lg: "linearform E g" and *: "\x \ F. g x = f x" and **: "\x \ E. g x \ p x" by blast have "\x \ E. \g x\ \ p x" using _ E sn lg ** proof (rule abs_ineq_iff [THEN iffD2]) show "E \ E" .. qed with lg * show ?thesis by blast qed subsection \The Hahn-Banach Theorem for normed spaces\ text \ Every continuous linear form \f\ on a subspace \F\ of a norm space \E\, can be extended to a continuous linear form \g\ on \E\ such that \\f\ = \g\\. \ theorem norm_Hahn_Banach: fixes V and norm ("\_\") fixes B defines "\V f. B V f \ {0} \ {\f x\ / \x\ | x. x \ 0 \ x \ V}" fixes fn_norm ("\_\\_" [0, 1000] 999) defines "\V f. \f\\V \ \(B V f)" assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E" and linearform: "linearform F f" and "continuous F f norm" shows "\g. linearform E g \ continuous E g norm \ (\x \ F. g x = f x) \ \g\\E = \f\\F" proof - interpret normed_vectorspace E norm by fact interpret normed_vectorspace_with_fn_norm E norm B fn_norm by (auto simp: B_def fn_norm_def) intro_locales interpret subspace F E by fact interpret linearform F f by fact interpret continuous F f norm by fact have E: "vectorspace E" by intro_locales have F: "vectorspace F" by rule intro_locales have F_norm: "normed_vectorspace F norm" using FE E_norm by (rule subspace_normed_vs) have ge_zero: "0 \ \f\\F" by (rule normed_vectorspace_with_fn_norm.fn_norm_ge_zero [OF normed_vectorspace_with_fn_norm.intro, OF F_norm \continuous F f norm\ , folded B_def fn_norm_def]) txt \We define a function \p\ on \E\ as follows: \p x = \f\ \ \x\\\ define p where "p x = \f\\F * \x\" for x txt \\p\ is a seminorm on \E\:\ have q: "seminorm E p" proof fix x y a assume x: "x \ E" and y: "y \ E" txt \\p\ is positive definite:\ have "0 \ \f\\F" by (rule ge_zero) moreover from x have "0 \ \x\" .. ultimately show "0 \ p x" by (simp add: p_def zero_le_mult_iff) txt \\p\ is absolutely homogeneous:\ show "p (a \ x) = \a\ * p x" proof - have "p (a \ x) = \f\\F * \a \ x\" by (simp only: p_def) also from x have "\a \ x\ = \a\ * \x\" by (rule abs_homogenous) also have "\f\\F * (\a\ * \x\) = \a\ * (\f\\F * \x\)" by simp also have "\ = \a\ * p x" by (simp only: p_def) finally show ?thesis . qed txt \Furthermore, \p\ is subadditive:\ show "p (x + y) \ p x + p y" proof - have "p (x + y) = \f\\F * \x + y\" by (simp only: p_def) also have a: "0 \ \f\\F" by (rule ge_zero) from x y have "\x + y\ \ \x\ + \y\" .. with a have " \f\\F * \x + y\ \ \f\\F * (\x\ + \y\)" by (simp add: mult_left_mono) also have "\ = \f\\F * \x\ + \f\\F * \y\" by (simp only: distrib_left) also have "\ = p x + p y" by (simp only: p_def) finally show ?thesis . qed qed txt \\f\ is bounded by \p\.\ have "\x \ F. \f x\ \ p x" proof fix x assume "x \ F" with \continuous F f norm\ and linearform show "\f x\ \ p x" unfolding p_def by (rule normed_vectorspace_with_fn_norm.fn_norm_le_cong [OF normed_vectorspace_with_fn_norm.intro, OF F_norm, folded B_def fn_norm_def]) qed txt \Using the fact that \p\ is a seminorm and \f\ is bounded by \p\ we can apply the Hahn-Banach Theorem for real vector spaces. So \f\ can be extended in a norm-preserving way to some function \g\ on the whole vector space \E\.\ with E FE linearform q obtain g where linearformE: "linearform E g" and a: "\x \ F. g x = f x" and b: "\x \ E. \g x\ \ p x" by (rule abs_Hahn_Banach [elim_format]) iprover txt \We furthermore have to show that \g\ is also continuous:\ have g_cont: "continuous E g norm" using linearformE proof fix x assume "x \ E" with b show "\g x\ \ \f\\F * \x\" by (simp only: p_def) qed txt \To complete the proof, we show that \\g\ = \f\\.\ have "\g\\E = \f\\F" proof (rule order_antisym) txt \ First we show \\g\ \ \f\\. The function norm \\g\\ is defined as the smallest \c \ \\ such that \begin{center} \begin{tabular}{l} \\x \ E. \g x\ \ c \ \x\\ \end{tabular} \end{center} \<^noindent> Furthermore holds \begin{center} \begin{tabular}{l} \\x \ E. \g x\ \ \f\ \ \x\\ \end{tabular} \end{center} \ from g_cont _ ge_zero show "\g\\E \ \f\\F" proof fix x assume "x \ E" with b show "\g x\ \ \f\\F * \x\" by (simp only: p_def) qed txt \The other direction is achieved by a similar argument.\ show "\f\\F \ \g\\E" proof (rule normed_vectorspace_with_fn_norm.fn_norm_least [OF normed_vectorspace_with_fn_norm.intro, OF F_norm, folded B_def fn_norm_def]) fix x assume x: "x \ F" show "\f x\ \ \g\\E * \x\" proof - from a x have "g x = f x" .. then have "\f x\ = \g x\" by (simp only:) also from g_cont have "\ \ \g\\E * \x\" proof (rule fn_norm_le_cong [OF _ linearformE, folded B_def fn_norm_def]) from FE x show "x \ E" .. qed finally show ?thesis . qed next show "0 \ \g\\E" using g_cont by (rule fn_norm_ge_zero [of g, folded B_def fn_norm_def]) show "continuous F f norm" by fact qed qed with linearformE a g_cont show ?thesis by blast qed end diff --git a/src/HOL/Imperative_HOL/Overview.thy b/src/HOL/Imperative_HOL/Overview.thy --- a/src/HOL/Imperative_HOL/Overview.thy +++ b/src/HOL/Imperative_HOL/Overview.thy @@ -1,240 +1,240 @@ (* Title: HOL/Imperative_HOL/Overview.thy Author: Florian Haftmann, TU Muenchen *) (*<*) theory Overview imports Imperative_HOL "HOL-Library.LaTeXsugar" begin (* type constraints with spacing *) no_syntax (output) "_constrain" :: "logic => type => logic" ("_::_" [4, 0] 3) "_constrain" :: "prop' => type => prop'" ("_::_" [4, 0] 3) syntax (output) "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3) "_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3) (*>*) text \ \Imperative HOL\ is a lightweight framework for reasoning about imperative data structures in \Isabelle/HOL\ - @{cite "Nipkow-et-al:2002:tutorial"}. Its basic ideas are described in - @{cite "Bulwahn-et-al:2008:imp_HOL"}. However their concrete + \<^cite>\"Nipkow-et-al:2002:tutorial"\. Its basic ideas are described in + \<^cite>\"Bulwahn-et-al:2008:imp_HOL"\. However their concrete realisation has changed since, due to both extensions and refinements. Therefore this overview wants to present the framework \qt{as it is} by now. It focusses on the user-view, less on matters of construction. For details study of the theory sources is encouraged. \ section \A polymorphic heap inside a monad\ text \ Heaps (\<^type>\heap\) can be populated by values of class \<^class>\heap\; HOL's default types are already instantiated to class \<^class>\heap\. Class \<^class>\heap\ is a subclass of \<^class>\countable\; see theory \Countable\ for ways to instantiate types as \<^class>\countable\. The heap is wrapped up in a monad \<^typ>\'a Heap\ by means of the following specification: \begin{quote} \<^datatype>\Heap\ \end{quote} Unwrapping of this monad type happens through \begin{quote} \<^term_type>\execute\ \\ @{thm execute.simps [no_vars]} \end{quote} This allows for equational reasoning about monadic expressions; the fact collection \execute_simps\ contains appropriate rewrites for all fundamental operations. Primitive fine-granular control over heaps is available through rule \Heap_cases\: \begin{quote} @{thm [break] Heap_cases [no_vars]} \end{quote} Monadic expression involve the usual combinators: \begin{quote} \<^term_type>\return\ \\ \<^term_type>\bind\ \\ \<^term_type>\raise\ \end{quote} This is also associated with nice monad do-syntax. The \<^typ>\string\ argument to \<^const>\raise\ is just a codified comment. Among a couple of generic combinators the following is helpful for establishing invariants: \begin{quote} \<^term_type>\assert\ \\ @{thm assert_def [no_vars]} \end{quote} \ section \Relational reasoning about \<^type>\Heap\ expressions\ text \ To establish correctness of imperative programs, predicate \begin{quote} \<^term_type>\effect\ \end{quote} provides a simple relational calculus. Primitive rules are \effectI\ and \effectE\, rules appropriate for reasoning about imperative operations are available in the \effect_intros\ and \effect_elims\ fact collections. Often non-failure of imperative computations does not depend on the heap at all; reasoning then can be easier using predicate \begin{quote} \<^term_type>\success\ \end{quote} Introduction rules for \<^const>\success\ are available in the \success_intro\ fact collection. \<^const>\execute\, \<^const>\effect\, \<^const>\success\ and \<^const>\bind\ are related by rules \execute_bind_success\, \success_bind_executeI\, \success_bind_effectI\, \effect_bindI\, \effect_bindE\ and \execute_bind_eq_SomeI\. \ section \Monadic data structures\ text \ The operations for monadic data structures (arrays and references) come in two flavours: \begin{itemize} \item Operations on the bare heap; their number is kept minimal to facilitate proving. \item Operations on the heap wrapped up in a monad; these are designed for executing. \end{itemize} Provided proof rules are such that they reduce monad operations to operations on bare heaps. Note that HOL equality coincides with reference equality and may be used as primitive executable operation. \ subsection \Arrays\ text \ Heap operations: \begin{quote} \<^term_type>\Array.alloc\ \\ \<^term_type>\Array.present\ \\ \<^term_type>\Array.get\ \\ \<^term_type>\Array.set\ \\ \<^term_type>\Array.length\ \\ \<^term_type>\Array.update\ \\ \<^term_type>\Array.noteq\ \end{quote} Monad operations: \begin{quote} \<^term_type>\Array.new\ \\ \<^term_type>\Array.of_list\ \\ \<^term_type>\Array.make\ \\ \<^term_type>\Array.len\ \\ \<^term_type>\Array.nth\ \\ \<^term_type>\Array.upd\ \\ \<^term_type>\Array.map_entry\ \\ \<^term_type>\Array.swap\ \\ \<^term_type>\Array.freeze\ \end{quote} \ subsection \References\ text \ Heap operations: \begin{quote} \<^term_type>\Ref.alloc\ \\ \<^term_type>\Ref.present\ \\ \<^term_type>\Ref.get\ \\ \<^term_type>\Ref.set\ \\ \<^term_type>\Ref.noteq\ \end{quote} Monad operations: \begin{quote} \<^term_type>\Ref.ref\ \\ \<^term_type>\Ref.lookup\ \\ \<^term_type>\Ref.update\ \\ \<^term_type>\Ref.change\ \end{quote} \ section \Code generation\ text \ Imperative HOL sets up the code generator in a way that imperative operations are mapped to suitable counterparts in the target language. For \Haskell\, a suitable \ST\ monad is used; for \SML\, \Ocaml\ and \Scala\ unit values ensure that the evaluation order is the same as you would expect from the original monadic expressions. These units may look cumbersome; the target language variants \SML_imp\, \Ocaml_imp\ and \Scala_imp\ make some effort to optimize some of them away. \ section \Some hints for using the framework\ text \ Of course a framework itself does not by itself indicate how to make best use of it. Here some hints drawn from prior experiences with Imperative HOL: \begin{itemize} \item Proofs on bare heaps should be strictly separated from those for monadic expressions. The first capture the essence, while the latter just describe a certain wrapping-up. \item A good methodology is to gradually improve an imperative program from a functional one. In the extreme case this means that an original functional program is decomposed into suitable operations with exactly one corresponding imperative operation. Having shown suitable correspondence lemmas between those, the correctness prove of the whole imperative program simply consists of composing those. \item Whether one should prefer equational reasoning (fact collection \execute_simps\ or relational reasoning (fact collections \effect_intros\ and \effect_elims\) depends on the problems to solve. For complex expressions or expressions involving binders, the relation style is usually superior but requires more proof text. \item Note that you can extend the fact collections of Imperative HOL yourself whenever appropriate. \end{itemize} \ end diff --git a/src/HOL/Induct/Comb.thy b/src/HOL/Induct/Comb.thy --- a/src/HOL/Induct/Comb.thy +++ b/src/HOL/Induct/Comb.thy @@ -1,201 +1,201 @@ (* Title: HOL/Induct/Comb.thy Author: Lawrence C Paulson Copyright 1996 University of Cambridge *) section \Combinatory Logic example: the Church-Rosser Theorem\ theory Comb imports Main begin text \ Combinator terms do not have free variables. - Example taken from @{cite camilleri92}. + Example taken from \<^cite>\camilleri92\. \ subsection \Definitions\ text \Datatype definition of combinators \S\ and \K\.\ datatype comb = K | S | Ap comb comb (infixl "\" 90) text \ Inductive definition of contractions, \\\<^sup>1\ and (multi-step) reductions, \\\. \ inductive contract1 :: "[comb,comb] \ bool" (infixl "\\<^sup>1" 50) where K: "K\x\y \\<^sup>1 x" | S: "S\x\y\z \\<^sup>1 (x\z)\(y\z)" | Ap1: "x \\<^sup>1 y \ x\z \\<^sup>1 y\z" | Ap2: "x \\<^sup>1 y \ z\x \\<^sup>1 z\y" abbreviation contract :: "[comb,comb] \ bool" (infixl "\" 50) where "contract \ contract1\<^sup>*\<^sup>*" text \ Inductive definition of parallel contractions, \\\<^sup>1\ and (multi-step) parallel reductions, \\\. \ inductive parcontract1 :: "[comb,comb] \ bool" (infixl "\\<^sup>1" 50) where refl: "x \\<^sup>1 x" | K: "K\x\y \\<^sup>1 x" | S: "S\x\y\z \\<^sup>1 (x\z)\(y\z)" | Ap: "\x \\<^sup>1 y; z \\<^sup>1 w\ \ x\z \\<^sup>1 y\w" abbreviation parcontract :: "[comb,comb] \ bool" (infixl "\" 50) where "parcontract \ parcontract1\<^sup>*\<^sup>*" text \ Misc definitions. \ definition I :: comb where "I \ S\K\K" definition diamond :: "([comb,comb] \ bool) \ bool" where \ \confluence; Lambda/Commutation treats this more abstractly\ "diamond r \ \x y. r x y \ (\y'. r x y' \ (\z. r y z \ r y' z))" subsection \Reflexive/Transitive closure preserves Church-Rosser property\ text\Remark: So does the Transitive closure, with a similar proof\ text\Strip lemma. The induction hypothesis covers all but the last diamond of the strip.\ lemma strip_lemma [rule_format]: assumes "diamond r" and r: "r\<^sup>*\<^sup>* x y" "r x y'" shows "\z. r\<^sup>*\<^sup>* y' z \ r y z" using r proof (induction rule: rtranclp_induct) case base then show ?case by blast next case (step y z) then show ?case using \diamond r\ unfolding diamond_def by (metis rtranclp.rtrancl_into_rtrancl) qed proposition diamond_rtrancl: assumes "diamond r" shows "diamond(r\<^sup>*\<^sup>*)" unfolding diamond_def proof (intro strip) fix x y y' assume "r\<^sup>*\<^sup>* x y" "r\<^sup>*\<^sup>* x y'" then show "\z. r\<^sup>*\<^sup>* y z \ r\<^sup>*\<^sup>* y' z" proof (induction rule: rtranclp_induct) case base then show ?case by blast next case (step y z) then show ?case by (meson assms strip_lemma rtranclp.rtrancl_into_rtrancl) qed qed subsection \Non-contraction results\ text \Derive a case for each combinator constructor.\ inductive_cases K_contractE [elim!]: "K \\<^sup>1 r" and S_contractE [elim!]: "S \\<^sup>1 r" and Ap_contractE [elim!]: "p\q \\<^sup>1 r" declare contract1.K [intro!] contract1.S [intro!] declare contract1.Ap1 [intro] contract1.Ap2 [intro] lemma I_contract_E [iff]: "\ I \\<^sup>1 z" unfolding I_def by blast lemma K1_contractD [elim!]: "K\x \\<^sup>1 z \ (\x'. z = K\x' \ x \\<^sup>1 x')" by blast lemma Ap_reduce1 [intro]: "x \ y \ x\z \ y\z" by (induction rule: rtranclp_induct; blast intro: rtranclp_trans) lemma Ap_reduce2 [intro]: "x \ y \ z\x \ z\y" by (induction rule: rtranclp_induct; blast intro: rtranclp_trans) text \Counterexample to the diamond property for \<^term>\x \\<^sup>1 y\\ lemma not_diamond_contract: "\ diamond(contract1)" unfolding diamond_def by (metis S_contractE contract1.K) subsection \Results about Parallel Contraction\ text \Derive a case for each combinator constructor.\ inductive_cases K_parcontractE [elim!]: "K \\<^sup>1 r" and S_parcontractE [elim!]: "S \\<^sup>1 r" and Ap_parcontractE [elim!]: "p\q \\<^sup>1 r" declare parcontract1.intros [intro] subsection \Basic properties of parallel contraction\ text\The rules below are not essential but make proofs much faster\ lemma K1_parcontractD [dest!]: "K\x \\<^sup>1 z \ (\x'. z = K\x' \ x \\<^sup>1 x')" by blast lemma S1_parcontractD [dest!]: "S\x \\<^sup>1 z \ (\x'. z = S\x' \ x \\<^sup>1 x')" by blast lemma S2_parcontractD [dest!]: "S\x\y \\<^sup>1 z \ (\x' y'. z = S\x'\y' \ x \\<^sup>1 x' \ y \\<^sup>1 y')" by blast text\Church-Rosser property for parallel contraction\ proposition diamond_parcontract: "diamond parcontract1" proof - have "(\z. w \\<^sup>1 z \ y' \\<^sup>1 z)" if "y \\<^sup>1 w" "y \\<^sup>1 y'" for w y y' using that by (induction arbitrary: y' rule: parcontract1.induct) fast+ then show ?thesis by (auto simp: diamond_def) qed subsection \Equivalence of \<^prop>\p \ q\ and \<^prop>\p \ q\.\ lemma contract_imp_parcontract: "x \\<^sup>1 y \ x \\<^sup>1 y" by (induction rule: contract1.induct; blast) text\Reductions: simply throw together reflexivity, transitivity and the one-step reductions\ proposition reduce_I: "I\x \ x" unfolding I_def by (meson contract1.K contract1.S r_into_rtranclp rtranclp.rtrancl_into_rtrancl) lemma parcontract_imp_reduce: "x \\<^sup>1 y \ x \ y" proof (induction rule: parcontract1.induct) case (Ap x y z w) then show ?case by (meson Ap_reduce1 Ap_reduce2 rtranclp_trans) qed auto lemma reduce_eq_parreduce: "x \ y \ x \ y" by (metis contract_imp_parcontract parcontract_imp_reduce predicate2I rtranclp_subset) theorem diamond_reduce: "diamond(contract)" using diamond_parcontract diamond_rtrancl reduce_eq_parreduce by presburger end diff --git a/src/HOL/Isar_Examples/Basic_Logic.thy b/src/HOL/Isar_Examples/Basic_Logic.thy --- a/src/HOL/Isar_Examples/Basic_Logic.thy +++ b/src/HOL/Isar_Examples/Basic_Logic.thy @@ -1,408 +1,407 @@ (* Title: HOL/Isar_Examples/Basic_Logic.thy Author: Makarius Basic propositional and quantifier reasoning. *) section \Basic logical reasoning\ theory Basic_Logic imports Main begin subsection \Pure backward reasoning\ text \ In order to get a first idea of how Isabelle/Isar proof documents may look like, we consider the propositions \I\, \K\, and \S\. The following (rather explicit) proofs should require little extra explanations. \ lemma I: "A \ A" proof assume A show A by fact qed lemma K: "A \ B \ A" proof assume A show "B \ A" proof show A by fact qed qed lemma S: "(A \ B \ C) \ (A \ B) \ A \ C" proof assume "A \ B \ C" show "(A \ B) \ A \ C" proof assume "A \ B" show "A \ C" proof assume A show C proof (rule mp) show "B \ C" by (rule mp) fact+ show B by (rule mp) fact+ qed qed qed qed text \ Isar provides several ways to fine-tune the reasoning, avoiding excessive detail. Several abbreviated language elements are available, enabling the writer to express proofs in a more concise way, even without referring to any automated proof tools yet. Concluding any (sub-)proof already involves solving any remaining goals by assumption\<^footnote>\This is not a completely trivial operation, as proof by assumption may involve full higher-order unification.\. Thus we may skip the rather vacuous body of the above proof. \ lemma "A \ A" proof qed text \ Note that the \<^theory_text>\proof\ command refers to the \rule\ method (without arguments) by default. Thus it implicitly applies a single rule, as determined from the syntactic form of the statements involved. The \<^theory_text>\by\ command abbreviates any proof with empty body, so the proof may be further pruned. \ lemma "A \ A" by rule text \ Proof by a single rule may be abbreviated as double-dot. \ lemma "A \ A" .. text \ Thus we have arrived at an adequate representation of the proof of a tautology that holds by a single standard rule.\<^footnote>\Apparently, the rule here is implication introduction.\ \<^medskip> Let us also reconsider \K\. Its statement is composed of iterated connectives. Basic decomposition is by a single rule at a time, which is why our first version above was by nesting two proofs. The \intro\ proof method repeatedly decomposes a goal's conclusion.\<^footnote>\The dual method is \elim\, acting on a goal's premises.\ \ lemma "A \ B \ A" proof (intro impI) assume A show A by fact qed text \Again, the body may be collapsed.\ lemma "A \ B \ A" by (intro impI) text \ Just like \rule\, the \intro\ and \elim\ proof methods pick standard structural rules, in case no explicit arguments are given. While implicit rules are usually just fine for single rule application, this may go too far with iteration. Thus in practice, \intro\ and \elim\ would be typically restricted to certain structures by giving a few rules only, e.g.\ \<^theory_text>\proof (intro impI allI)\ to strip implications and universal quantifiers. Such well-tuned iterated decomposition of certain structures is the prime application of \intro\ and \elim\. In contrast, terminal steps that solve a goal completely are usually performed by actual automated proof methods (such as \<^theory_text>\by blast\. \ subsection \Variations of backward vs.\ forward reasoning\ text \ Certainly, any proof may be performed in backward-style only. On the other hand, small steps of reasoning are often more naturally expressed in forward-style. Isar supports both backward and forward reasoning as a first-class concept. In order to demonstrate the difference, we consider several proofs of \A \ B \ B \ A\. The first version is purely backward. \ lemma "A \ B \ B \ A" proof assume "A \ B" show "B \ A" proof show B by (rule conjunct2) fact show A by (rule conjunct1) fact qed qed text \ Above, the projection rules \conjunct1\ / \conjunct2\ had to be named explicitly, since the goals \B\ and \A\ did not provide any structural clue. This may be avoided using \<^theory_text>\from\ to focus on the \A \ B\ assumption as the current facts, enabling the use of double-dot proofs. Note that \<^theory_text>\from\ already does forward-chaining, involving the \conjE\ rule here. \ lemma "A \ B \ B \ A" proof assume "A \ B" show "B \ A" proof from \A \ B\ show B .. from \A \ B\ show A .. qed qed text \ In the next version, we move the forward step one level upwards. Forward-chaining from the most recent facts is indicated by the \<^theory_text>\then\ command. Thus the proof of \B \ A\ from \A \ B\ actually becomes an elimination, rather than an introduction. The resulting proof structure directly corresponds to that of the \conjE\ rule, including the repeated goal proposition that is abbreviated as \?thesis\ below. \ lemma "A \ B \ B \ A" proof assume "A \ B" then show "B \ A" proof \ \rule \conjE\ of \A \ B\\ assume B A then show ?thesis .. \ \rule \conjI\ of \B \ A\\ qed qed text \ In the subsequent version we flatten the structure of the main body by doing forward reasoning all the time. Only the outermost decomposition step is left as backward. \ lemma "A \ B \ B \ A" proof assume "A \ B" from \A \ B\ have A .. from \A \ B\ have B .. from \B\ \A\ show "B \ A" .. qed text \ We can still push forward-reasoning a bit further, even at the risk of getting ridiculous. Note that we force the initial proof step to do nothing here, by referring to the \-\ proof method. \ lemma "A \ B \ B \ A" proof - { assume "A \ B" from \A \ B\ have A .. from \A \ B\ have B .. from \B\ \A\ have "B \ A" .. } then show ?thesis .. \ \rule \impI\\ qed text \ \<^medskip> With these examples we have shifted through a whole range from purely backward to purely forward reasoning. Apparently, in the extreme ends we get slightly ill-structured proofs, which also require much explicit naming of either rules (backward) or local facts (forward). The general lesson learned here is that good proof style would achieve just the \<^emph>\right\ balance of top-down backward decomposition, and bottom-up forward composition. In general, there is no single best way to arrange some pieces of formal reasoning, of course. Depending on the actual applications, the intended audience etc., rules (and methods) on the one hand vs.\ facts on the other hand have to be emphasized in an appropriate way. This requires the proof writer to develop good taste, and some practice, of course. \<^medskip> For our example the most appropriate way of reasoning is probably the middle one, with conjunction introduction done after elimination. \ lemma "A \ B \ B \ A" proof assume "A \ B" then show "B \ A" proof assume B A then show ?thesis .. qed qed subsection \A few examples from ``Introduction to Isabelle''\ text \ - We rephrase some of the basic reasoning examples of @{cite - "isabelle-intro"}, using HOL rather than FOL. + We rephrase some of the basic reasoning examples of \<^cite>\"isabelle-intro"\, using HOL rather than FOL. \ subsubsection \A propositional proof\ text \ We consider the proposition \P \ P \ P\. The proof below involves forward-chaining from \P \ P\, followed by an explicit case-analysis on the two \<^emph>\identical\ cases. \ lemma "P \ P \ P" proof assume "P \ P" then show P proof \ \rule \disjE\: \smash{$\infer{\C\}{\A \ B\ & \infer*{\C\}{[\A\]} & \infer*{\C\}{[\B\]}}$}\ assume P show P by fact next assume P show P by fact qed qed text \ Case splits are \<^emph>\not\ hardwired into the Isar language as a special feature. The \<^theory_text>\next\ command used to separate the cases above is just a short form of managing block structure. \<^medskip> In general, applying proof methods may split up a goal into separate ``cases'', i.e.\ new subgoals with individual local assumptions. The corresponding proof text typically mimics this by establishing results in appropriate contexts, separated by blocks. In order to avoid too much explicit parentheses, the Isar system implicitly opens an additional block for any new goal, the \<^theory_text>\next\ statement then closes one block level, opening a new one. The resulting behaviour is what one would expect from separating cases, only that it is more flexible. E.g.\ an induction base case (which does not introduce local assumptions) would \<^emph>\not\ require \<^theory_text>\next\ to separate the subsequent step case. \<^medskip> In our example the situation is even simpler, since the two cases actually coincide. Consequently the proof may be rephrased as follows. \ lemma "P \ P \ P" proof assume "P \ P" then show P proof assume P show P by fact show P by fact qed qed text \Again, the rather vacuous body of the proof may be collapsed. Thus the case analysis degenerates into two assumption steps, which are implicitly performed when concluding the single rule step of the double-dot proof as follows.\ lemma "P \ P \ P" proof assume "P \ P" then show P .. qed subsubsection \A quantifier proof\ text \ To illustrate quantifier reasoning, let us prove \(\x. P (f x)) \ (\y. P y)\. Informally, this holds because any \a\ with \P (f a)\ may be taken as a witness for the second existential statement. The first proof is rather verbose, exhibiting quite a lot of (redundant) detail. It gives explicit rules, even with some instantiation. Furthermore, we encounter two new language elements: the \<^theory_text>\fix\ command augments the context by some new ``arbitrary, but fixed'' element; the \<^theory_text>\is\ annotation binds term abbreviations by higher-order pattern matching. \ lemma "(\x. P (f x)) \ (\y. P y)" proof assume "\x. P (f x)" then show "\y. P y" proof (rule exE) \ \rule \exE\: \smash{$\infer{\B\}{\\x. A(x)\ & \infer*{\B\}{[\A(x)\]_x}}$}\ fix a assume "P (f a)" (is "P ?witness") then show ?thesis by (rule exI [of P ?witness]) qed qed text \ While explicit rule instantiation may occasionally improve readability of certain aspects of reasoning, it is usually quite redundant. Above, the basic proof outline gives already enough structural clues for the system to infer both the rules and their instances (by higher-order unification). Thus we may as well prune the text as follows. \ lemma "(\x. P (f x)) \ (\y. P y)" proof assume "\x. P (f x)" then show "\y. P y" proof fix a assume "P (f a)" then show ?thesis .. qed qed text \ Explicit \\\-elimination as seen above can become quite cumbersome in practice. The derived Isar language element ``\<^theory_text>\obtain\'' provides a more handsome way to do generalized existence reasoning. \ lemma "(\x. P (f x)) \ (\y. P y)" proof assume "\x. P (f x)" then obtain a where "P (f a)" .. then show "\y. P y" .. qed text \ Technically, \<^theory_text>\obtain\ is similar to \<^theory_text>\fix\ and \<^theory_text>\assume\ together with a soundness proof of the elimination involved. Thus it behaves similar to any other forward proof element. Also note that due to the nature of general existence reasoning involved here, any result exported from the context of an \<^theory_text>\obtain\ statement may \<^emph>\not\ refer to the parameters introduced there. \ subsubsection \Deriving rules in Isabelle\ text \ We derive the conjunction elimination rule from the corresponding projections. The proof is quite straight-forward, since Isabelle/Isar supports non-atomic goals and assumptions fully transparently. \ theorem conjE: "A \ B \ (A \ B \ C) \ C" proof - assume "A \ B" assume r: "A \ B \ C" show C proof (rule r) show A by (rule conjunct1) fact show B by (rule conjunct2) fact qed qed end diff --git a/src/HOL/Isar_Examples/Fibonacci.thy b/src/HOL/Isar_Examples/Fibonacci.thy --- a/src/HOL/Isar_Examples/Fibonacci.thy +++ b/src/HOL/Isar_Examples/Fibonacci.thy @@ -1,169 +1,169 @@ (* Title: HOL/Isar_Examples/Fibonacci.thy Author: Gertrud Bauer Copyright 1999 Technische Universitaet Muenchen The Fibonacci function. Original tactic script by Lawrence C Paulson. Fibonacci numbers: proofs of laws taken from R. L. Graham, D. E. Knuth, O. Patashnik. Concrete Mathematics. (Addison-Wesley, 1989) *) section \Fib and Gcd commute\ theory Fibonacci imports "HOL-Computational_Algebra.Primes" begin text_raw \\<^footnote>\Isar version by Gertrud Bauer. Original tactic script by Larry - Paulson. A few proofs of laws taken from @{cite "Concrete-Math"}.\\ + Paulson. A few proofs of laws taken from \<^cite>\"Concrete-Math"\.\\ subsection \Fibonacci numbers\ fun fib :: "nat \ nat" where "fib 0 = 0" | "fib (Suc 0) = 1" | "fib (Suc (Suc x)) = fib x + fib (Suc x)" lemma [simp]: "fib (Suc n) > 0" by (induct n rule: fib.induct) simp_all text \Alternative induction rule.\ theorem fib_induct: "P 0 \ P 1 \ (\n. P (n + 1) \ P n \ P (n + 2)) \ P n" for n :: nat by (induct rule: fib.induct) simp_all subsection \Fib and gcd commute\ -text \A few laws taken from @{cite "Concrete-Math"}.\ +text \A few laws taken from \<^cite>\"Concrete-Math"\.\ lemma fib_add: "fib (n + k + 1) = fib (k + 1) * fib (n + 1) + fib k * fib n" (is "?P n") - \ \see @{cite \page 280\ "Concrete-Math"}\ + \ \see \<^cite>\\page 280\ in "Concrete-Math"\\ proof (induct n rule: fib_induct) show "?P 0" by simp show "?P 1" by simp fix n have "fib (n + 2 + k + 1) = fib (n + k + 1) + fib (n + 1 + k + 1)" by simp also assume "fib (n + k + 1) = fib (k + 1) * fib (n + 1) + fib k * fib n" (is " _ = ?R1") also assume "fib (n + 1 + k + 1) = fib (k + 1) * fib (n + 1 + 1) + fib k * fib (n + 1)" (is " _ = ?R2") also have "?R1 + ?R2 = fib (k + 1) * fib (n + 2 + 1) + fib k * fib (n + 2)" by (simp add: add_mult_distrib2) finally show "?P (n + 2)" . qed lemma coprime_fib_Suc: "coprime (fib n) (fib (n + 1))" (is "?P n") proof (induct n rule: fib_induct) show "?P 0" by simp show "?P 1" by simp fix n assume P: "coprime (fib (n + 1)) (fib (n + 1 + 1))" have "fib (n + 2 + 1) = fib (n + 1) + fib (n + 2)" by simp also have "\ = fib (n + 2) + fib (n + 1)" by simp also have "gcd (fib (n + 2)) \ = gcd (fib (n + 2)) (fib (n + 1))" by (rule gcd_add2) also have "\ = gcd (fib (n + 1)) (fib (n + 1 + 1))" by (simp add: gcd.commute) also have "\ = 1" using P by simp finally show "?P (n + 2)" by (simp add: coprime_iff_gcd_eq_1) qed lemma gcd_mult_add: "(0::nat) < n \ gcd (n * k + m) n = gcd m n" proof - assume "0 < n" then have "gcd (n * k + m) n = gcd n (m mod n)" by (simp add: gcd_non_0_nat add.commute) also from \0 < n\ have "\ = gcd m n" by (simp add: gcd_non_0_nat) finally show ?thesis . qed lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)" proof (cases m) case 0 then show ?thesis by simp next case (Suc k) then have "gcd (fib m) (fib (n + m)) = gcd (fib (n + k + 1)) (fib (k + 1))" by (simp add: gcd.commute) also have "fib (n + k + 1) = fib (k + 1) * fib (n + 1) + fib k * fib n" by (rule fib_add) also have "gcd \ (fib (k + 1)) = gcd (fib k * fib n) (fib (k + 1))" by (simp add: gcd_mult_add) also have "\ = gcd (fib n) (fib (k + 1))" using coprime_fib_Suc [of k] gcd_mult_left_right_cancel [of "fib (k + 1)" "fib k" "fib n"] by (simp add: ac_simps) also have "\ = gcd (fib m) (fib n)" using Suc by (simp add: gcd.commute) finally show ?thesis . qed lemma gcd_fib_diff: "gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)" if "m \ n" proof - have "gcd (fib m) (fib (n - m)) = gcd (fib m) (fib (n - m + m))" by (simp add: gcd_fib_add) also from \m \ n\ have "n - m + m = n" by simp finally show ?thesis . qed lemma gcd_fib_mod: "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" if "0 < m" proof (induct n rule: nat_less_induct) case hyp: (1 n) show ?case proof - have "n mod m = (if n < m then n else (n - m) mod m)" by (rule mod_if) also have "gcd (fib m) (fib \) = gcd (fib m) (fib n)" proof (cases "n < m") case True then show ?thesis by simp next case False then have "m \ n" by simp from \0 < m\ and False have "n - m < n" by simp with hyp have "gcd (fib m) (fib ((n - m) mod m)) = gcd (fib m) (fib (n - m))" by simp also have "\ = gcd (fib m) (fib n)" using \m \ n\ by (rule gcd_fib_diff) finally have "gcd (fib m) (fib ((n - m) mod m)) = gcd (fib m) (fib n)" . with False show ?thesis by simp qed finally show ?thesis . qed qed theorem fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)" (is "?P m n") proof (induct m n rule: gcd_nat_induct) fix m n :: nat show "fib (gcd m 0) = gcd (fib m) (fib 0)" by simp assume n: "0 < n" then have "gcd m n = gcd n (m mod n)" by (simp add: gcd_non_0_nat) also assume hyp: "fib \ = gcd (fib n) (fib (m mod n))" also from n have "\ = gcd (fib n) (fib m)" by (rule gcd_fib_mod) also have "\ = gcd (fib m) (fib n)" by (rule gcd.commute) finally show "fib (gcd m n) = gcd (fib m) (fib n)" . qed end diff --git a/src/HOL/Isar_Examples/Hoare.thy b/src/HOL/Isar_Examples/Hoare.thy --- a/src/HOL/Isar_Examples/Hoare.thy +++ b/src/HOL/Isar_Examples/Hoare.thy @@ -1,413 +1,412 @@ (* Title: HOL/Isar_Examples/Hoare.thy Author: Makarius A formulation of Hoare logic suitable for Isar. *) section \Hoare Logic\ theory Hoare imports "HOL-Hoare.Hoare_Tac" begin subsection \Abstract syntax and semantics\ text \ The following abstract syntax and semantics of Hoare Logic over \<^verbatim>\WHILE\ programs closely follows the existing tradition in Isabelle/HOL of - formalizing the presentation given in @{cite \\S6\ "Winskel:1993"}. See also - \<^dir>\~~/src/HOL/Hoare\ and @{cite "Nipkow:1998:Winskel"}. + formalizing the presentation given in \<^cite>\\\S6\ in "Winskel:1993"\. See also + \<^dir>\~~/src/HOL/Hoare\ and \<^cite>\"Nipkow:1998:Winskel"\. \ type_synonym 'a bexp = "'a set" type_synonym 'a assn = "'a set" type_synonym 'a var = "'a \ nat" datatype 'a com = Basic "'a \ 'a" | Seq "'a com" "'a com" ("(_;/ _)" [60, 61] 60) | Cond "'a bexp" "'a com" "'a com" | While "'a bexp" "'a assn" "'a var" "'a com" abbreviation Skip ("SKIP") where "SKIP \ Basic id" type_synonym 'a sem = "'a \ 'a \ bool" primrec iter :: "nat \ 'a bexp \ 'a sem \ 'a sem" where "iter 0 b S s s' \ s \ b \ s = s'" | "iter (Suc n) b S s s' \ s \ b \ (\s''. S s s'' \ iter n b S s'' s')" primrec Sem :: "'a com \ 'a sem" where "Sem (Basic f) s s' \ s' = f s" | "Sem (c1; c2) s s' \ (\s''. Sem c1 s s'' \ Sem c2 s'' s')" | "Sem (Cond b c1 c2) s s' \ (if s \ b then Sem c1 s s' else Sem c2 s s')" | "Sem (While b x y c) s s' \ (\n. iter n b (Sem c) s s')" definition Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" ("(3\ _/ (2_)/ _)" [100, 55, 100] 50) where "\ P c Q \ (\s s'. Sem c s s' \ s \ P \ s' \ Q)" lemma ValidI [intro?]: "(\s s'. Sem c s s' \ s \ P \ s' \ Q) \ \ P c Q" by (simp add: Valid_def) lemma ValidD [dest?]: "\ P c Q \ Sem c s s' \ s \ P \ s' \ Q" by (simp add: Valid_def) subsection \Primitive Hoare rules\ text \ From the semantics defined above, we derive the standard set of primitive - Hoare rules; e.g.\ see @{cite \\S6\ "Winskel:1993"}. Usually, variant forms + Hoare rules; e.g.\ see \<^cite>\\\S6\ in "Winskel:1993"\. Usually, variant forms of these rules are applied in actual proof, see also \S\ref{sec:hoare-isar} and \S\ref{sec:hoare-vcg}. \<^medskip> The \basic\ rule represents any kind of atomic access to the state space. This subsumes the common rules of \skip\ and \assign\, as formulated in \S\ref{sec:hoare-isar}. \ theorem basic: "\ {s. f s \ P} (Basic f) P" proof fix s s' assume s: "s \ {s. f s \ P}" assume "Sem (Basic f) s s'" then have "s' = f s" by simp with s show "s' \ P" by simp qed text \ The rules for sequential commands and semantic consequences are established in a straight forward manner as follows. \ theorem seq: "\ P c1 Q \ \ Q c2 R \ \ P (c1; c2) R" proof assume cmd1: "\ P c1 Q" and cmd2: "\ Q c2 R" fix s s' assume s: "s \ P" assume "Sem (c1; c2) s s'" then obtain s'' where sem1: "Sem c1 s s''" and sem2: "Sem c2 s'' s'" by auto from cmd1 sem1 s have "s'' \ Q" .. with cmd2 sem2 show "s' \ R" .. qed theorem conseq: "P' \ P \ \ P c Q \ Q \ Q' \ \ P' c Q'" proof assume P'P: "P' \ P" and QQ': "Q \ Q'" assume cmd: "\ P c Q" fix s s' :: 'a assume sem: "Sem c s s'" assume "s \ P'" with P'P have "s \ P" .. with cmd sem have "s' \ Q" .. with QQ' show "s' \ Q'" .. qed text \ The rule for conditional commands is directly reflected by the corresponding semantics; in the proof we just have to look closely which cases apply. \ theorem cond: assumes case_b: "\ (P \ b) c1 Q" and case_nb: "\ (P \ -b) c2 Q" shows "\ P (Cond b c1 c2) Q" proof fix s s' assume s: "s \ P" assume sem: "Sem (Cond b c1 c2) s s'" show "s' \ Q" proof cases assume b: "s \ b" from case_b show ?thesis proof from sem b show "Sem c1 s s'" by simp from s b show "s \ P \ b" by simp qed next assume nb: "s \ b" from case_nb show ?thesis proof from sem nb show "Sem c2 s s'" by simp from s nb show "s \ P \ -b" by simp qed qed qed text \ The \while\ rule is slightly less trivial --- it is the only one based on recursion, which is expressed in the semantics by a Kleene-style least fixed-point construction. The auxiliary statement below, which is by induction on the number of iterations is the main point to be proven; the rest is by routine application of the semantics of \<^verbatim>\WHILE\. \ theorem while: assumes body: "\ (P \ b) c P" shows "\ P (While b X Y c) (P \ -b)" proof fix s s' assume s: "s \ P" assume "Sem (While b X Y c) s s'" then obtain n where "iter n b (Sem c) s s'" by auto from this and s show "s' \ P \ -b" proof (induct n arbitrary: s) case 0 then show ?case by auto next case (Suc n) then obtain s'' where b: "s \ b" and sem: "Sem c s s''" and iter: "iter n b (Sem c) s'' s'" by auto from Suc and b have "s \ P \ b" by simp with body sem have "s'' \ P" .. with iter show ?case by (rule Suc) qed qed subsection \Concrete syntax for assertions\ text \ We now introduce concrete syntax for describing commands (with embedded expressions) and assertions. The basic technique is that of semantic ``quote-antiquote''. A \<^emph>\quotation\ is a syntactic entity delimited by an implicit abstraction, say over the state space. An \<^emph>\antiquotation\ is a marked expression within a quotation that refers the implicit argument; a typical antiquotation would select (or even update) components from the state. We will see some examples later in the concrete rules and applications. \<^medskip> The following specification of syntax and translations is for Isabelle experts only; feel free to ignore it. While the first part is still a somewhat intelligible specification of the concrete syntactic representation of our Hoare language, the actual ``ML drivers'' is quite involved. Just note that the we re-use the basic quote/antiquote translations as already defined in Isabelle/Pure (see \<^ML>\Syntax_Trans.quote_tr\, and \<^ML>\Syntax_Trans.quote_tr'\,). \ syntax "_quote" :: "'b \ ('a \ 'b)" "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) "_Subst" :: "'a bexp \ 'b \ idt \ 'a bexp" ("_[_'/\_]" [1000] 999) "_Assert" :: "'a \ 'a set" ("(\_\)" [0] 1000) "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61) "_While_inv" :: "'a bexp \ 'a assn \ 'a com \ 'a com" ("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61) "_While" :: "'a bexp \ 'a com \ 'a com" ("(0WHILE _ //DO _ /OD)" [0, 0] 61) translations "\b\" \ "CONST Collect (_quote b)" "B [a/\x]" \ "\\(_update_name x (\_. a)) \ B\" "\x := a" \ "CONST Basic (_quote (\(_update_name x (\_. a))))" "IF b THEN c1 ELSE c2 FI" \ "CONST Cond \b\ c1 c2" "WHILE b INV i DO c OD" \ "CONST While \b\ i (\_. 0) c" "WHILE b DO c OD" \ "WHILE b INV CONST undefined DO c OD" parse_translation \ let fun quote_tr [t] = Syntax_Trans.quote_tr \<^syntax_const>\_antiquote\ t | quote_tr ts = raise TERM ("quote_tr", ts); in [(\<^syntax_const>\_quote\, K quote_tr)] end \ text \ As usual in Isabelle syntax translations, the part for printing is more complicated --- we cannot express parts as macro rules as above. Don't look here, unless you have to do similar things for yourself. \ print_translation \ let fun quote_tr' f (t :: ts) = Term.list_comb (f $ Syntax_Trans.quote_tr' \<^syntax_const>\_antiquote\ t, ts) | quote_tr' _ _ = raise Match; val assert_tr' = quote_tr' (Syntax.const \<^syntax_const>\_Assert\); fun bexp_tr' name ((Const (\<^const_syntax>\Collect\, _) $ t) :: ts) = quote_tr' (Syntax.const name) (t :: ts) | bexp_tr' _ _ = raise Match; fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = quote_tr' (Syntax.const \<^syntax_const>\_Assign\ $ Syntax_Trans.update_name_tr' f) (Abs (x, dummyT, Syntax_Trans.const_abs_tr' k) :: ts) | assign_tr' _ = raise Match; in [(\<^const_syntax>\Collect\, K assert_tr'), (\<^const_syntax>\Basic\, K assign_tr'), (\<^const_syntax>\Cond\, K (bexp_tr' \<^syntax_const>\_Cond\)), (\<^const_syntax>\While\, K (bexp_tr' \<^syntax_const>\_While_inv\))] end \ subsection \Rules for single-step proof \label{sec:hoare-isar}\ text \ We are now ready to introduce a set of Hoare rules to be used in single-step structured proofs in Isabelle/Isar. We refer to the concrete syntax introduce above. \<^medskip> Assertions of Hoare Logic may be manipulated in calculational proofs, with the inclusion expressed in terms of sets or predicates. Reversed order is supported as well. \ lemma [trans]: "\ P c Q \ P' \ P \ \ P' c Q" by (unfold Valid_def) blast lemma [trans] : "P' \ P \ \ P c Q \ \ P' c Q" by (unfold Valid_def) blast lemma [trans]: "Q \ Q' \ \ P c Q \ \ P c Q'" by (unfold Valid_def) blast lemma [trans]: "\ P c Q \ Q \ Q' \ \ P c Q'" by (unfold Valid_def) blast lemma [trans]: "\ \\P\ c Q \ (\s. P' s \ P s) \ \ \\P'\ c Q" by (simp add: Valid_def) lemma [trans]: "(\s. P' s \ P s) \ \ \\P\ c Q \ \ \\P'\ c Q" by (simp add: Valid_def) lemma [trans]: "\ P c \\Q\ \ (\s. Q s \ Q' s) \ \ P c \\Q'\" by (simp add: Valid_def) lemma [trans]: "(\s. Q s \ Q' s) \ \ P c \\Q\ \ \ P c \\Q'\" by (simp add: Valid_def) text \ Identity and basic assignments.\<^footnote>\The \hoare\ method introduced in \S\ref{sec:hoare-vcg} is able to provide proper instances for any number of basic assignments, without producing additional verification conditions.\ \ lemma skip [intro?]: "\ P SKIP P" proof - have "\ {s. id s \ P} SKIP P" by (rule basic) then show ?thesis by simp qed lemma assign: "\ P [\a/\x::'a] \x := \a P" by (rule basic) text \ Note that above formulation of assignment corresponds to our preferred way - to model state spaces, using (extensible) record types in HOL @{cite - "Naraschewski-Wenzel:1998:HOOL"}. For any record field \x\, Isabelle/HOL + to model state spaces, using (extensible) record types in HOL \<^cite>\"Naraschewski-Wenzel:1998:HOOL"\. For any record field \x\, Isabelle/HOL provides a functions \x\ (selector) and \x_update\ (update). Above, there is only a place-holder appearing for the latter kind of function: due to concrete syntax \\x := \a\ also contains \x_update\.\<^footnote>\Note that due to the external nature of HOL record fields, we could not even state a general theorem relating selector and update functions (if this were required here); this would only work for any particular instance of record fields introduced so far.\ \<^medskip> Sequential composition --- normalizing with associativity achieves proper of chunks of code verified separately. \ lemmas [trans, intro?] = seq lemma seq_assoc [simp]: "\ P c1;(c2;c3) Q \ \ P (c1;c2);c3 Q" by (auto simp add: Valid_def) text \Conditional statements.\ lemmas [trans, intro?] = cond lemma [trans, intro?]: "\ \\P \ \b\ c1 Q \ \ \\P \ \ \b\ c2 Q \ \ \\P\ IF \b THEN c1 ELSE c2 FI Q" by (rule cond) (simp_all add: Valid_def) text \While statements --- with optional invariant.\ lemma [intro?]: "\ (P \ b) c P \ \ P (While b P V c) (P \ -b)" by (rule while) lemma [intro?]: "\ (P \ b) c P \ \ P (While b undefined V c) (P \ -b)" by (rule while) lemma [intro?]: "\ \\P \ \b\ c \\P\ \ \ \\P\ WHILE \b INV \\P\ DO c OD \\P \ \ \b\" by (simp add: while Collect_conj_eq Collect_neg_eq) lemma [intro?]: "\ \\P \ \b\ c \\P\ \ \ \\P\ WHILE \b DO c OD \\P \ \ \b\" by (simp add: while Collect_conj_eq Collect_neg_eq) subsection \Verification conditions \label{sec:hoare-vcg}\ text \ We now load the \<^emph>\original\ ML file for proof scripts and tactic definition for the Hoare Verification Condition Generator (see \<^dir>\~~/src/HOL/Hoare\). As far as we are concerned here, the result is a proof method \hoare\, which may be applied to a Hoare Logic assertion to extract purely logical verification conditions. It is important to note that the method requires \<^verbatim>\WHILE\ loops to be fully annotated with invariants beforehand. Furthermore, only \<^emph>\concrete\ pieces of code are handled --- the underlying tactic fails ungracefully if supplied with meta-variables or parameters, for example. \ lemma SkipRule: "p \ q \ Valid p (Basic id) q" by (auto simp add: Valid_def) lemma BasicRule: "p \ {s. f s \ q} \ Valid p (Basic f) q" by (auto simp: Valid_def) lemma SeqRule: "Valid P c1 Q \ Valid Q c2 R \ Valid P (c1;c2) R" by (auto simp: Valid_def) lemma CondRule: "p \ {s. (s \ b \ s \ w) \ (s \ b \ s \ w')} \ Valid w c1 q \ Valid w' c2 q \ Valid p (Cond b c1 c2) q" by (auto simp: Valid_def) lemma iter_aux: "\s s'. Sem c s s' \ s \ I \ s \ b \ s' \ I \ (\s s'. s \ I \ iter n b (Sem c) s s' \ s' \ I \ s' \ b)" by (induct n) auto lemma WhileRule: "p \ i \ Valid (i \ b) c i \ i \ (-b) \ q \ Valid p (While b i v c) q" apply (clarsimp simp: Valid_def) apply (drule iter_aux) prefer 2 apply assumption apply blast apply blast done declare BasicRule [Hoare_Tac.BasicRule] and SkipRule [Hoare_Tac.SkipRule] and SeqRule [Hoare_Tac.SeqRule] and CondRule [Hoare_Tac.CondRule] and WhileRule [Hoare_Tac.WhileRule] method_setup hoare = \Scan.succeed (fn ctxt => (SIMPLE_METHOD' (Hoare_Tac.hoare_tac ctxt (simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm "Record.K_record_comp"}] )))))\ "verification condition generator for Hoare logic" end diff --git a/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy b/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy --- a/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy +++ b/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy @@ -1,291 +1,290 @@ (* Title: HOL/Isar_Examples/Mutilated_Checkerboard.thy Author: Markus Wenzel, TU Muenchen (Isar document) Author: Lawrence C Paulson, Cambridge University Computer Laboratory (original scripts) *) section \The Mutilated Checker Board Problem\ theory Mutilated_Checkerboard imports Main begin text \ - The Mutilated Checker Board Problem, formalized inductively. See @{cite - "paulson-mutilated-board"} for the original tactic script version. + The Mutilated Checker Board Problem, formalized inductively. See \<^cite>\"paulson-mutilated-board"\ for the original tactic script version. \ subsection \Tilings\ inductive_set tiling :: "'a set set \ 'a set set" for A :: "'a set set" where empty: "{} \ tiling A" | Un: "a \ t \ tiling A" if "a \ A" and "t \ tiling A" and "a \ - t" text \The union of two disjoint tilings is a tiling.\ lemma tiling_Un: assumes "t \ tiling A" and "u \ tiling A" and "t \ u = {}" shows "t \ u \ tiling A" proof - let ?T = "tiling A" from \t \ ?T\ and \t \ u = {}\ show "t \ u \ ?T" proof (induct t) case empty with \u \ ?T\ show "{} \ u \ ?T" by simp next case (Un a t) show "(a \ t) \ u \ ?T" proof - have "a \ (t \ u) \ ?T" using \a \ A\ proof (rule tiling.Un) from \(a \ t) \ u = {}\ have "t \ u = {}" by blast then show "t \ u \ ?T" by (rule Un) from \a \ - t\ and \(a \ t) \ u = {}\ show "a \ - (t \ u)" by blast qed also have "a \ (t \ u) = (a \ t) \ u" by (simp only: Un_assoc) finally show ?thesis . qed qed qed subsection \Basic properties of ``below''\ definition below :: "nat \ nat set" where "below n = {i. i < n}" lemma below_less_iff [iff]: "i \ below k \ i < k" by (simp add: below_def) lemma below_0: "below 0 = {}" by (simp add: below_def) lemma Sigma_Suc1: "m = n + 1 \ below m \ B = ({n} \ B) \ (below n \ B)" by (simp add: below_def less_Suc_eq) blast lemma Sigma_Suc2: "m = n + 2 \ A \ below m = (A \ {n}) \ (A \ {n + 1}) \ (A \ below n)" by (auto simp add: below_def) lemmas Sigma_Suc = Sigma_Suc1 Sigma_Suc2 subsection \Basic properties of ``evnodd''\ definition evnodd :: "(nat \ nat) set \ nat \ (nat \ nat) set" where "evnodd A b = A \ {(i, j). (i + j) mod 2 = b}" lemma evnodd_iff: "(i, j) \ evnodd A b \ (i, j) \ A \ (i + j) mod 2 = b" by (simp add: evnodd_def) lemma evnodd_subset: "evnodd A b \ A" unfolding evnodd_def by (rule Int_lower1) lemma evnoddD: "x \ evnodd A b \ x \ A" by (rule subsetD) (rule evnodd_subset) lemma evnodd_finite: "finite A \ finite (evnodd A b)" by (rule finite_subset) (rule evnodd_subset) lemma evnodd_Un: "evnodd (A \ B) b = evnodd A b \ evnodd B b" unfolding evnodd_def by blast lemma evnodd_Diff: "evnodd (A - B) b = evnodd A b - evnodd B b" unfolding evnodd_def by blast lemma evnodd_empty: "evnodd {} b = {}" by (simp add: evnodd_def) lemma evnodd_insert: "evnodd (insert (i, j) C) b = (if (i + j) mod 2 = b then insert (i, j) (evnodd C b) else evnodd C b)" by (simp add: evnodd_def) subsection \Dominoes\ inductive_set domino :: "(nat \ nat) set set" where horiz: "{(i, j), (i, j + 1)} \ domino" | vertl: "{(i, j), (i + 1, j)} \ domino" lemma dominoes_tile_row: "{i} \ below (2 * n) \ tiling domino" (is "?B n \ ?T") proof (induct n) case 0 show ?case by (simp add: below_0 tiling.empty) next case (Suc n) let ?a = "{i} \ {2 * n + 1} \ {i} \ {2 * n}" have "?B (Suc n) = ?a \ ?B n" by (auto simp add: Sigma_Suc Un_assoc) also have "\ \ ?T" proof (rule tiling.Un) have "{(i, 2 * n), (i, 2 * n + 1)} \ domino" by (rule domino.horiz) also have "{(i, 2 * n), (i, 2 * n + 1)} = ?a" by blast finally show "\ \ domino" . show "?B n \ ?T" by (rule Suc) show "?a \ - ?B n" by blast qed finally show ?case . qed lemma dominoes_tile_matrix: "below m \ below (2 * n) \ tiling domino" (is "?B m \ ?T") proof (induct m) case 0 show ?case by (simp add: below_0 tiling.empty) next case (Suc m) let ?t = "{m} \ below (2 * n)" have "?B (Suc m) = ?t \ ?B m" by (simp add: Sigma_Suc) also have "\ \ ?T" proof (rule tiling_Un) show "?t \ ?T" by (rule dominoes_tile_row) show "?B m \ ?T" by (rule Suc) show "?t \ ?B m = {}" by blast qed finally show ?case . qed lemma domino_singleton: assumes "d \ domino" and "b < 2" shows "\i j. evnodd d b = {(i, j)}" (is "?P d") using assms proof induct from \b < 2\ have b_cases: "b = 0 \ b = 1" by arith fix i j note [simp] = evnodd_empty evnodd_insert mod_Suc from b_cases show "?P {(i, j), (i, j + 1)}" by rule auto from b_cases show "?P {(i, j), (i + 1, j)}" by rule auto qed lemma domino_finite: assumes "d \ domino" shows "finite d" using assms proof induct fix i j :: nat show "finite {(i, j), (i, j + 1)}" by (intro finite.intros) show "finite {(i, j), (i + 1, j)}" by (intro finite.intros) qed subsection \Tilings of dominoes\ lemma tiling_domino_finite: assumes t: "t \ tiling domino" (is "t \ ?T") shows "finite t" (is "?F t") using t proof induct show "?F {}" by (rule finite.emptyI) fix a t assume "?F t" assume "a \ domino" then have "?F a" by (rule domino_finite) from this and \?F t\ show "?F (a \ t)" by (rule finite_UnI) qed lemma tiling_domino_01: assumes t: "t \ tiling domino" (is "t \ ?T") shows "card (evnodd t 0) = card (evnodd t 1)" using t proof induct case empty show ?case by (simp add: evnodd_def) next case (Un a t) let ?e = evnodd note hyp = \card (?e t 0) = card (?e t 1)\ and at = \a \ - t\ have card_suc: "card (?e (a \ t) b) = Suc (card (?e t b))" if "b < 2" for b :: nat proof - have "?e (a \ t) b = ?e a b \ ?e t b" by (rule evnodd_Un) also obtain i j where e: "?e a b = {(i, j)}" proof - from \a \ domino\ and \b < 2\ have "\i j. ?e a b = {(i, j)}" by (rule domino_singleton) then show ?thesis by (blast intro: that) qed also have "\ \ ?e t b = insert (i, j) (?e t b)" by simp also have "card \ = Suc (card (?e t b))" proof (rule card_insert_disjoint) from \t \ tiling domino\ have "finite t" by (rule tiling_domino_finite) then show "finite (?e t b)" by (rule evnodd_finite) from e have "(i, j) \ ?e a b" by simp with at show "(i, j) \ ?e t b" by (blast dest: evnoddD) qed finally show ?thesis . qed then have "card (?e (a \ t) 0) = Suc (card (?e t 0))" by simp also from hyp have "card (?e t 0) = card (?e t 1)" . also from card_suc have "Suc \ = card (?e (a \ t) 1)" by simp finally show ?case . qed subsection \Main theorem\ definition mutilated_board :: "nat \ nat \ (nat \ nat) set" where "mutilated_board m n = below (2 * (m + 1)) \ below (2 * (n + 1)) - {(0, 0)} - {(2 * m + 1, 2 * n + 1)}" theorem mutil_not_tiling: "mutilated_board m n \ tiling domino" proof (unfold mutilated_board_def) let ?T = "tiling domino" let ?t = "below (2 * (m + 1)) \ below (2 * (n + 1))" let ?t' = "?t - {(0, 0)}" let ?t'' = "?t' - {(2 * m + 1, 2 * n + 1)}" show "?t'' \ ?T" proof have t: "?t \ ?T" by (rule dominoes_tile_matrix) assume t'': "?t'' \ ?T" let ?e = evnodd have fin: "finite (?e ?t 0)" by (rule evnodd_finite, rule tiling_domino_finite, rule t) note [simp] = evnodd_iff evnodd_empty evnodd_insert evnodd_Diff have "card (?e ?t'' 0) < card (?e ?t' 0)" proof - have "card (?e ?t' 0 - {(2 * m + 1, 2 * n + 1)}) < card (?e ?t' 0)" proof (rule card_Diff1_less) from _ fin show "finite (?e ?t' 0)" by (rule finite_subset) auto show "(2 * m + 1, 2 * n + 1) \ ?e ?t' 0" by simp qed then show ?thesis by simp qed also have "\ < card (?e ?t 0)" proof - have "(0, 0) \ ?e ?t 0" by simp with fin have "card (?e ?t 0 - {(0, 0)}) < card (?e ?t 0)" by (rule card_Diff1_less) then show ?thesis by simp qed also from t have "\ = card (?e ?t 1)" by (rule tiling_domino_01) also have "?e ?t 1 = ?e ?t'' 1" by simp also from t'' have "card \ = card (?e ?t'' 0)" by (rule tiling_domino_01 [symmetric]) finally have "\ < \" . then show False .. qed qed end diff --git a/src/HOL/Lattice/CompleteLattice.thy b/src/HOL/Lattice/CompleteLattice.thy --- a/src/HOL/Lattice/CompleteLattice.thy +++ b/src/HOL/Lattice/CompleteLattice.thy @@ -1,452 +1,452 @@ (* Title: HOL/Lattice/CompleteLattice.thy Author: Markus Wenzel, TU Muenchen *) section \Complete lattices\ theory CompleteLattice imports Lattice begin subsection \Complete lattice operations\ text \ A \emph{complete lattice} is a partial order with general (infinitary) infimum of any set of elements. General supremum exists as well, as a consequence of the connection of infinitary bounds (see \S\ref{sec:connect-bounds}). \ class complete_lattice = assumes ex_Inf: "\inf. is_Inf A inf" theorem ex_Sup: "\sup::'a::complete_lattice. is_Sup A sup" proof - from ex_Inf obtain sup where "is_Inf {b. \a\A. a \ b} sup" by blast then have "is_Sup A sup" by (rule Inf_Sup) then show ?thesis .. qed text \ The general \\\ (meet) and \\\ (join) operations select such infimum and supremum elements. \ definition Meet :: "'a::complete_lattice set \ 'a" ("\_" [90] 90) where "\A = (THE inf. is_Inf A inf)" definition Join :: "'a::complete_lattice set \ 'a" ("\_" [90] 90) where "\A = (THE sup. is_Sup A sup)" text \ Due to unique existence of bounds, the complete lattice operations may be exhibited as follows. \ lemma Meet_equality [elim?]: "is_Inf A inf \ \A = inf" proof (unfold Meet_def) assume "is_Inf A inf" then show "(THE inf. is_Inf A inf) = inf" by (rule the_equality) (rule is_Inf_uniq [OF _ \is_Inf A inf\]) qed lemma MeetI [intro?]: "(\a. a \ A \ inf \ a) \ (\b. \a \ A. b \ a \ b \ inf) \ \A = inf" by (rule Meet_equality, rule is_InfI) blast+ lemma Join_equality [elim?]: "is_Sup A sup \ \A = sup" proof (unfold Join_def) assume "is_Sup A sup" then show "(THE sup. is_Sup A sup) = sup" by (rule the_equality) (rule is_Sup_uniq [OF _ \is_Sup A sup\]) qed lemma JoinI [intro?]: "(\a. a \ A \ a \ sup) \ (\b. \a \ A. a \ b \ sup \ b) \ \A = sup" by (rule Join_equality, rule is_SupI) blast+ text \ \medskip The \\\ and \\\ operations indeed determine bounds on a complete lattice structure. \ lemma is_Inf_Meet [intro?]: "is_Inf A (\A)" proof (unfold Meet_def) from ex_Inf obtain inf where "is_Inf A inf" .. then show "is_Inf A (THE inf. is_Inf A inf)" by (rule theI) (rule is_Inf_uniq [OF _ \is_Inf A inf\]) qed lemma Meet_greatest [intro?]: "(\a. a \ A \ x \ a) \ x \ \A" by (rule is_Inf_greatest, rule is_Inf_Meet) blast lemma Meet_lower [intro?]: "a \ A \ \A \ a" by (rule is_Inf_lower) (rule is_Inf_Meet) lemma is_Sup_Join [intro?]: "is_Sup A (\A)" proof (unfold Join_def) from ex_Sup obtain sup where "is_Sup A sup" .. then show "is_Sup A (THE sup. is_Sup A sup)" by (rule theI) (rule is_Sup_uniq [OF _ \is_Sup A sup\]) qed lemma Join_least [intro?]: "(\a. a \ A \ a \ x) \ \A \ x" by (rule is_Sup_least, rule is_Sup_Join) blast lemma Join_lower [intro?]: "a \ A \ a \ \A" by (rule is_Sup_upper) (rule is_Sup_Join) subsection \The Knaster-Tarski Theorem\ text \ The Knaster-Tarski Theorem (in its simplest formulation) states that any monotone function on a complete lattice has a least fixed-point - (see @{cite \pages 93--94\ "Davey-Priestley:1990"} for example). This + (see \<^cite>\\pages 93--94\ in "Davey-Priestley:1990"\ for example). This is a consequence of the basic boundary properties of the complete lattice operations. \ theorem Knaster_Tarski: assumes mono: "\x y. x \ y \ f x \ f y" obtains a :: "'a::complete_lattice" where "f a = a" and "\a'. f a' = a' \ a \ a'" proof let ?H = "{u. f u \ u}" let ?a = "\?H" show "f ?a = ?a" proof - have ge: "f ?a \ ?a" proof fix x assume x: "x \ ?H" then have "?a \ x" .. then have "f ?a \ f x" by (rule mono) also from x have "... \ x" .. finally show "f ?a \ x" . qed also have "?a \ f ?a" proof from ge have "f (f ?a) \ f ?a" by (rule mono) then show "f ?a \ ?H" .. qed finally show ?thesis . qed fix a' assume "f a' = a'" then have "f a' \ a'" by (simp only: leq_refl) then have "a' \ ?H" .. then show "?a \ a'" .. qed theorem Knaster_Tarski_dual: assumes mono: "\x y. x \ y \ f x \ f y" obtains a :: "'a::complete_lattice" where "f a = a" and "\a'. f a' = a' \ a' \ a" proof let ?H = "{u. u \ f u}" let ?a = "\?H" show "f ?a = ?a" proof - have le: "?a \ f ?a" proof fix x assume x: "x \ ?H" then have "x \ f x" .. also from x have "x \ ?a" .. then have "f x \ f ?a" by (rule mono) finally show "x \ f ?a" . qed have "f ?a \ ?a" proof from le have "f ?a \ f (f ?a)" by (rule mono) then show "f ?a \ ?H" .. qed from this and le show ?thesis by (rule leq_antisym) qed fix a' assume "f a' = a'" then have "a' \ f a'" by (simp only: leq_refl) then have "a' \ ?H" .. then show "a' \ ?a" .. qed subsection \Bottom and top elements\ text \ With general bounds available, complete lattices also have least and greatest elements. \ definition bottom :: "'a::complete_lattice" ("\") where "\ = \UNIV" definition top :: "'a::complete_lattice" ("\") where "\ = \UNIV" lemma bottom_least [intro?]: "\ \ x" proof (unfold bottom_def) have "x \ UNIV" .. then show "\UNIV \ x" .. qed lemma bottomI [intro?]: "(\a. x \ a) \ \ = x" proof (unfold bottom_def) assume "\a. x \ a" show "\UNIV = x" proof fix a show "x \ a" by fact next fix b :: "'a::complete_lattice" assume b: "\a \ UNIV. b \ a" have "x \ UNIV" .. with b show "b \ x" .. qed qed lemma top_greatest [intro?]: "x \ \" proof (unfold top_def) have "x \ UNIV" .. then show "x \ \UNIV" .. qed lemma topI [intro?]: "(\a. a \ x) \ \ = x" proof (unfold top_def) assume "\a. a \ x" show "\UNIV = x" proof fix a show "a \ x" by fact next fix b :: "'a::complete_lattice" assume b: "\a \ UNIV. a \ b" have "x \ UNIV" .. with b show "x \ b" .. qed qed subsection \Duality\ text \ The class of complete lattices is closed under formation of dual structures. \ instance dual :: (complete_lattice) complete_lattice proof fix A' :: "'a::complete_lattice dual set" show "\inf'. is_Inf A' inf'" proof - have "\sup. is_Sup (undual ` A') sup" by (rule ex_Sup) then have "\sup. is_Inf (dual ` undual ` A') (dual sup)" by (simp only: dual_Inf) then show ?thesis by (simp add: dual_ex [symmetric] image_comp) qed qed text \ Apparently, the \\\ and \\\ operations are dual to each other. \ theorem dual_Meet [intro?]: "dual (\A) = \(dual ` A)" proof - from is_Inf_Meet have "is_Sup (dual ` A) (dual (\A))" .. then have "\(dual ` A) = dual (\A)" .. then show ?thesis .. qed theorem dual_Join [intro?]: "dual (\A) = \(dual ` A)" proof - from is_Sup_Join have "is_Inf (dual ` A) (dual (\A))" .. then have "\(dual ` A) = dual (\A)" .. then show ?thesis .. qed text \ Likewise are \\\ and \\\ duals of each other. \ theorem dual_bottom [intro?]: "dual \ = \" proof - have "\ = dual \" proof fix a' have "\ \ undual a'" .. then have "dual (undual a') \ dual \" .. then show "a' \ dual \" by simp qed then show ?thesis .. qed theorem dual_top [intro?]: "dual \ = \" proof - have "\ = dual \" proof fix a' have "undual a' \ \" .. then have "dual \ \ dual (undual a')" .. then show "dual \ \ a'" by simp qed then show ?thesis .. qed subsection \Complete lattices are lattices\ text \ Complete lattices (with general bounds available) are indeed plain lattices as well. This holds due to the connection of general versus binary bounds that has been formally established in \S\ref{sec:gen-bin-bounds}. \ lemma is_inf_binary: "is_inf x y (\{x, y})" proof - have "is_Inf {x, y} (\{x, y})" .. then show ?thesis by (simp only: is_Inf_binary) qed lemma is_sup_binary: "is_sup x y (\{x, y})" proof - have "is_Sup {x, y} (\{x, y})" .. then show ?thesis by (simp only: is_Sup_binary) qed instance complete_lattice \ lattice proof fix x y :: "'a::complete_lattice" from is_inf_binary show "\inf. is_inf x y inf" .. from is_sup_binary show "\sup. is_sup x y sup" .. qed theorem meet_binary: "x \ y = \{x, y}" by (rule meet_equality) (rule is_inf_binary) theorem join_binary: "x \ y = \{x, y}" by (rule join_equality) (rule is_sup_binary) subsection \Complete lattices and set-theory operations\ text \ The complete lattice operations are (anti) monotone wrt.\ set inclusion. \ theorem Meet_subset_antimono: "A \ B \ \B \ \A" proof (rule Meet_greatest) fix a assume "a \ A" also assume "A \ B" finally have "a \ B" . then show "\B \ a" .. qed theorem Join_subset_mono: "A \ B \ \A \ \B" proof - assume "A \ B" then have "dual ` A \ dual ` B" by blast then have "\(dual ` B) \ \(dual ` A)" by (rule Meet_subset_antimono) then have "dual (\B) \ dual (\A)" by (simp only: dual_Join) then show ?thesis by (simp only: dual_leq) qed text \ Bounds over unions of sets may be obtained separately. \ theorem Meet_Un: "\(A \ B) = \A \ \B" proof fix a assume "a \ A \ B" then show "\A \ \B \ a" proof assume a: "a \ A" have "\A \ \B \ \A" .. also from a have "\ \ a" .. finally show ?thesis . next assume a: "a \ B" have "\A \ \B \ \B" .. also from a have "\ \ a" .. finally show ?thesis . qed next fix b assume b: "\a \ A \ B. b \ a" show "b \ \A \ \B" proof show "b \ \A" proof fix a assume "a \ A" then have "a \ A \ B" .. with b show "b \ a" .. qed show "b \ \B" proof fix a assume "a \ B" then have "a \ A \ B" .. with b show "b \ a" .. qed qed qed theorem Join_Un: "\(A \ B) = \A \ \B" proof - have "dual (\(A \ B)) = \(dual ` A \ dual ` B)" by (simp only: dual_Join image_Un) also have "\ = \(dual ` A) \ \(dual ` B)" by (rule Meet_Un) also have "\ = dual (\A \ \B)" by (simp only: dual_join dual_Join) finally show ?thesis .. qed text \ Bounds over singleton sets are trivial. \ theorem Meet_singleton: "\{x} = x" proof fix a assume "a \ {x}" then have "a = x" by simp then show "x \ a" by (simp only: leq_refl) next fix b assume "\a \ {x}. b \ a" then show "b \ x" by simp qed theorem Join_singleton: "\{x} = x" proof - have "dual (\{x}) = \{dual x}" by (simp add: dual_Join) also have "\ = dual x" by (rule Meet_singleton) finally show ?thesis .. qed text \ Bounds over the empty and universal set correspond to each other. \ theorem Meet_empty: "\{} = \UNIV" proof fix a :: "'a::complete_lattice" assume "a \ {}" then have False by simp then show "\UNIV \ a" .. next fix b :: "'a::complete_lattice" have "b \ UNIV" .. then show "b \ \UNIV" .. qed theorem Join_empty: "\{} = \UNIV" proof - have "dual (\{}) = \{}" by (simp add: dual_Join) also have "\ = \UNIV" by (rule Meet_empty) also have "\ = dual (\UNIV)" by (simp add: dual_Meet) finally show ?thesis .. qed end diff --git a/src/HOL/Library/BigO.thy b/src/HOL/Library/BigO.thy --- a/src/HOL/Library/BigO.thy +++ b/src/HOL/Library/BigO.thy @@ -1,797 +1,796 @@ (* Title: HOL/Library/BigO.thy Authors: Jeremy Avigad and Kevin Donnelly *) section \Big O notation\ theory BigO imports Complex_Main Function_Algebras Set_Algebras begin text \ This library is designed to support asymptotic ``big O'' calculations, i.e.~reasoning with expressions of the form \f = O(g)\ and \f = g + O(h)\. - An earlier version of this library is described in detail in @{cite - "Avigad-Donnelly"}. + An earlier version of this library is described in detail in \<^cite>\"Avigad-Donnelly"\. The main changes in this version are as follows: \<^item> We have eliminated the \O\ operator on sets. (Most uses of this seem to be inessential.) \<^item> We no longer use \+\ as output syntax for \+o\ \<^item> Lemmas involving \sumr\ have been replaced by more general lemmas involving `\sum\. \<^item> The library has been expanded, with e.g.~support for expressions of the form \f < g + O(h)\. Note also since the Big O library includes rules that demonstrate set inclusion, to use the automated reasoners effectively with the library one should redeclare the theorem \subsetI\ as an intro rule, rather than as an \intro!\ rule, for example, using \<^theory_text>\declare subsetI [del, intro]\. \ subsection \Definitions\ definition bigo :: "('a \ 'b::linordered_idom) \ ('a \ 'b) set" ("(1O'(_'))") where "O(f:: 'a \ 'b) = {h. \c. \x. \h x\ \ c * \f x\}" lemma bigo_pos_const: "(\c::'a::linordered_idom. \x. \h x\ \ c * \f x\) \ (\c. 0 < c \ (\x. \h x\ \ c * \f x\))" by (metis (no_types, opaque_lifting) abs_ge_zero abs_not_less_zero abs_of_nonneg dual_order.trans mult_1 zero_less_abs_iff zero_less_mult_iff zero_less_one) lemma bigo_alt_def: "O(f) = {h. \c. 0 < c \ (\x. \h x\ \ c * \f x\)}" by (auto simp add: bigo_def bigo_pos_const) lemma bigo_elt_subset [intro]: "f \ O(g) \ O(f) \ O(g)" apply (auto simp add: bigo_alt_def) by (metis (no_types, opaque_lifting) mult.assoc mult_le_cancel_iff2 order.trans zero_less_mult_iff) lemma bigo_refl [intro]: "f \ O(f)" using bigo_def comm_monoid_mult_class.mult_1 dual_order.eq_iff by blast lemma bigo_zero: "0 \ O(g)" using bigo_def mult_le_cancel_left1 by fastforce lemma bigo_zero2: "O(\x. 0) = {\x. 0}" by (auto simp add: bigo_def) lemma bigo_plus_self_subset [intro]: "O(f) + O(f) \ O(f)" apply (auto simp add: bigo_alt_def set_plus_def) apply (rule_tac x = "c + ca" in exI) by (smt (verit, best) abs_triangle_ineq add_mono add_pos_pos comm_semiring_class.distrib dual_order.trans) lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)" by (simp add: antisym bigo_plus_self_subset bigo_zero set_zero_plus2) lemma bigo_plus_subset [intro]: "O(f + g) \ O(f) + O(g)" apply (rule subsetI) apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def) apply (subst bigo_pos_const [symmetric])+ apply (rule_tac x = "\n. if \g n\ \ \f n\ then x n else 0" in exI) apply (rule conjI) apply (rule_tac x = "c + c" in exI) apply (clarsimp) apply (subgoal_tac "c * \f xa + g xa\ \ (c + c) * \f xa\") apply (metis mult_2 order_trans) apply (subgoal_tac "c * \f xa + g xa\ \ c * (\f xa\ + \g xa\)") apply auto[1] using abs_triangle_ineq mult_le_cancel_iff2 apply blast apply (simp add: order_less_le) apply (rule_tac x = "\n. if \f n\ < \g n\ then x n else 0" in exI) apply (rule conjI) apply (rule_tac x = "c + c" in exI) apply auto apply (subgoal_tac "c * \f xa + g xa\ \ (c + c) * \g xa\") apply (metis mult_2 order.trans) apply simp done lemma bigo_plus_subset2 [intro]: "A \ O(f) \ B \ O(f) \ A + B \ O(f)" using bigo_plus_idemp set_plus_mono2 by blast lemma bigo_plus_eq: "\x. 0 \ f x \ \x. 0 \ g x \ O(f + g) = O(f) + O(g)" apply (rule equalityI) apply (rule bigo_plus_subset) apply (simp add: bigo_alt_def set_plus_def func_plus) apply clarify apply (rule_tac x = "max c ca" in exI) apply (rule conjI) apply (subgoal_tac "c \ max c ca") apply linarith apply (rule max.cobounded1) apply clarify apply (drule_tac x = "xa" in spec)+ apply (subgoal_tac "0 \ f xa + g xa") apply (simp add: ring_distribs) apply (subgoal_tac "\a xa + b xa\ \ \a xa\ + \b xa\") apply (subgoal_tac "\a xa\ + \b xa\ \ max c ca * f xa + max c ca * g xa") apply force apply (metis add_mono le_max_iff_disj max_mult_distrib_right) using abs_triangle_ineq apply blast using add_nonneg_nonneg by blast lemma bigo_bounded_alt: "\x. 0 \ f x \ \x. f x \ c * g x \ f \ O(g)" apply (auto simp add: bigo_def) apply (rule_tac x = "\c\" in exI) apply auto apply (drule_tac x = x in spec)+ apply (simp flip: abs_mult) done lemma bigo_bounded: "\x. 0 \ f x \ \x. f x \ g x \ f \ O(g)" apply (erule bigo_bounded_alt [of f 1 g]) apply simp done lemma bigo_bounded2: "\x. lb x \ f x \ \x. f x \ lb x + g x \ f \ lb +o O(g)" apply (rule set_minus_imp_plus) apply (rule bigo_bounded) apply (auto simp add: fun_Compl_def func_plus) apply (drule_tac x = x in spec)+ apply force done lemma bigo_abs: "(\x. \f x\) =o O(f)" apply (unfold bigo_def) apply auto apply (rule_tac x = 1 in exI) apply auto done lemma bigo_abs2: "f =o O(\x. \f x\)" apply (unfold bigo_def) apply auto apply (rule_tac x = 1 in exI) apply auto done lemma bigo_abs3: "O(f) = O(\x. \f x\)" apply (rule equalityI) apply (rule bigo_elt_subset) apply (rule bigo_abs2) apply (rule bigo_elt_subset) apply (rule bigo_abs) done lemma bigo_abs4: "f =o g +o O(h) \ (\x. \f x\) =o (\x. \g x\) +o O(h)" apply (drule set_plus_imp_minus) apply (rule set_minus_imp_plus) apply (subst fun_diff_def) proof - assume *: "f - g \ O(h)" have "(\x. \f x\ - \g x\) =o O(\x. \\f x\ - \g x\\)" by (rule bigo_abs2) also have "\ \ O(\x. \f x - g x\)" apply (rule bigo_elt_subset) apply (rule bigo_bounded) apply force apply (rule allI) apply (rule abs_triangle_ineq3) done also have "\ \ O(f - g)" apply (rule bigo_elt_subset) apply (subst fun_diff_def) apply (rule bigo_abs) done also from * have "\ \ O(h)" by (rule bigo_elt_subset) finally show "(\x. \f x\ - \g x\) \ O(h)". qed lemma bigo_abs5: "f =o O(g) \ (\x. \f x\) =o O(g)" by (auto simp: bigo_def) lemma bigo_elt_subset2 [intro]: assumes *: "f \ g +o O(h)" shows "O(f) \ O(g) + O(h)" proof - note * also have "g +o O(h) \ O(g) + O(h)" by (auto del: subsetI) also have "\ = O(\x. \g x\) + O(\x. \h x\)" by (subst bigo_abs3 [symmetric])+ (rule refl) also have "\ = O((\x. \g x\) + (\x. \h x\))" by (rule bigo_plus_eq [symmetric]) auto finally have "f \ \" . then have "O(f) \ \" by (elim bigo_elt_subset) also have "\ = O(\x. \g x\) + O(\x. \h x\)" by (rule bigo_plus_eq, auto) finally show ?thesis by (simp flip: bigo_abs3) qed lemma bigo_mult [intro]: "O(f)*O(g) \ O(f * g)" apply (rule subsetI) apply (subst bigo_def) apply (auto simp add: bigo_alt_def set_times_def func_times) apply (rule_tac x = "c * ca" in exI) apply (rule allI) apply (erule_tac x = x in allE)+ apply (subgoal_tac "c * ca * \f x * g x\ = (c * \f x\) * (ca * \g x\)") apply (erule ssubst) apply (subst abs_mult) apply (rule mult_mono) apply assumption+ apply auto apply (simp add: ac_simps abs_mult) done lemma bigo_mult2 [intro]: "f *o O(g) \ O(f * g)" apply (auto simp add: bigo_def elt_set_times_def func_times abs_mult) apply (rule_tac x = c in exI) apply auto apply (drule_tac x = x in spec) apply (subgoal_tac "\f x\ * \b x\ \ \f x\ * (c * \g x\)") apply (force simp add: ac_simps) apply (rule mult_left_mono, assumption) apply (rule abs_ge_zero) done lemma bigo_mult3: "f \ O(h) \ g \ O(j) \ f * g \ O(h * j)" apply (rule subsetD) apply (rule bigo_mult) apply (erule set_times_intro, assumption) done lemma bigo_mult4 [intro]: "f \ k +o O(h) \ g * f \ (g * k) +o O(g * h)" apply (drule set_plus_imp_minus) apply (rule set_minus_imp_plus) apply (drule bigo_mult3 [where g = g and j = g]) apply (auto simp add: algebra_simps) done lemma bigo_mult5: fixes f :: "'a \ 'b::linordered_field" assumes "\x. f x \ 0" shows "O(f * g) \ f *o O(g)" proof fix h assume "h \ O(f * g)" then have "(\x. 1 / (f x)) * h \ (\x. 1 / f x) *o O(f * g)" by auto also have "\ \ O((\x. 1 / f x) * (f * g))" by (rule bigo_mult2) also have "(\x. 1 / f x) * (f * g) = g" apply (simp add: func_times) apply (rule ext) apply (simp add: assms nonzero_divide_eq_eq ac_simps) done finally have "(\x. (1::'b) / f x) * h \ O(g)" . then have "f * ((\x. (1::'b) / f x) * h) \ f *o O(g)" by auto also have "f * ((\x. (1::'b) / f x) * h) = h" apply (simp add: func_times) apply (rule ext) apply (simp add: assms nonzero_divide_eq_eq ac_simps) done finally show "h \ f *o O(g)" . qed lemma bigo_mult6: "\x. f x \ 0 \ O(f * g) = f *o O(g)" for f :: "'a \ 'b::linordered_field" apply (rule equalityI) apply (erule bigo_mult5) apply (rule bigo_mult2) done lemma bigo_mult7: "\x. f x \ 0 \ O(f * g) \ O(f) * O(g)" for f :: "'a \ 'b::linordered_field" apply (subst bigo_mult6) apply assumption apply (rule set_times_mono3) apply (rule bigo_refl) done lemma bigo_mult8: "\x. f x \ 0 \ O(f * g) = O(f) * O(g)" for f :: "'a \ 'b::linordered_field" apply (rule equalityI) apply (erule bigo_mult7) apply (rule bigo_mult) done lemma bigo_minus [intro]: "f \ O(g) \ - f \ O(g)" by (auto simp add: bigo_def fun_Compl_def) lemma bigo_minus2: "f \ g +o O(h) \ - f \ -g +o O(h)" apply (rule set_minus_imp_plus) apply (drule set_plus_imp_minus) apply (drule bigo_minus) apply simp done lemma bigo_minus3: "O(- f) = O(f)" by (auto simp add: bigo_def fun_Compl_def) lemma bigo_plus_absorb_lemma1: assumes *: "f \ O(g)" shows "f +o O(g) \ O(g)" proof - have "f \ O(f)" by auto then have "f +o O(g) \ O(f) + O(g)" by (auto del: subsetI) also have "\ \ O(g) + O(g)" proof - from * have "O(f) \ O(g)" by (auto del: subsetI) then show ?thesis by (auto del: subsetI) qed also have "\ \ O(g)" by simp finally show ?thesis . qed lemma bigo_plus_absorb_lemma2: assumes *: "f \ O(g)" shows "O(g) \ f +o O(g)" proof - from * have "- f \ O(g)" by auto then have "- f +o O(g) \ O(g)" by (elim bigo_plus_absorb_lemma1) then have "f +o (- f +o O(g)) \ f +o O(g)" by auto also have "f +o (- f +o O(g)) = O(g)" by (simp add: set_plus_rearranges) finally show ?thesis . qed lemma bigo_plus_absorb [simp]: "f \ O(g) \ f +o O(g) = O(g)" apply (rule equalityI) apply (erule bigo_plus_absorb_lemma1) apply (erule bigo_plus_absorb_lemma2) done lemma bigo_plus_absorb2 [intro]: "f \ O(g) \ A \ O(g) \ f +o A \ O(g)" apply (subgoal_tac "f +o A \ f +o O(g)") apply force+ done lemma bigo_add_commute_imp: "f \ g +o O(h) \ g \ f +o O(h)" apply (subst set_minus_plus [symmetric]) apply (subgoal_tac "g - f = - (f - g)") apply (erule ssubst) apply (rule bigo_minus) apply (subst set_minus_plus) apply assumption apply (simp add: ac_simps) done lemma bigo_add_commute: "f \ g +o O(h) \ g \ f +o O(h)" apply (rule iffI) apply (erule bigo_add_commute_imp)+ done lemma bigo_const1: "(\x. c) \ O(\x. 1)" by (auto simp add: bigo_def ac_simps) lemma bigo_const2 [intro]: "O(\x. c) \ O(\x. 1)" apply (rule bigo_elt_subset) apply (rule bigo_const1) done lemma bigo_const3: "c \ 0 \ (\x. 1) \ O(\x. c)" for c :: "'a::linordered_field" apply (simp add: bigo_def) apply (rule_tac x = "\inverse c\" in exI) apply (simp flip: abs_mult) done lemma bigo_const4: "c \ 0 \ O(\x. 1) \ O(\x. c)" for c :: "'a::linordered_field" apply (rule bigo_elt_subset) apply (rule bigo_const3) apply assumption done lemma bigo_const [simp]: "c \ 0 \ O(\x. c) = O(\x. 1)" for c :: "'a::linordered_field" apply (rule equalityI) apply (rule bigo_const2) apply (rule bigo_const4) apply assumption done lemma bigo_const_mult1: "(\x. c * f x) \ O(f)" apply (simp add: bigo_def) apply (rule_tac x = "\c\" in exI) apply (auto simp flip: abs_mult) done lemma bigo_const_mult2: "O(\x. c * f x) \ O(f)" apply (rule bigo_elt_subset) apply (rule bigo_const_mult1) done lemma bigo_const_mult3: "c \ 0 \ f \ O(\x. c * f x)" for c :: "'a::linordered_field" apply (simp add: bigo_def) apply (rule_tac x = "\inverse c\" in exI) apply (simp add: abs_mult mult.assoc [symmetric]) done lemma bigo_const_mult4: "c \ 0 \ O(f) \ O(\x. c * f x)" for c :: "'a::linordered_field" apply (rule bigo_elt_subset) apply (rule bigo_const_mult3) apply assumption done lemma bigo_const_mult [simp]: "c \ 0 \ O(\x. c * f x) = O(f)" for c :: "'a::linordered_field" apply (rule equalityI) apply (rule bigo_const_mult2) apply (erule bigo_const_mult4) done lemma bigo_const_mult5 [simp]: "c \ 0 \ (\x. c) *o O(f) = O(f)" for c :: "'a::linordered_field" apply (auto del: subsetI) apply (rule order_trans) apply (rule bigo_mult2) apply (simp add: func_times) apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) apply (rule_tac x = "\y. inverse c * x y" in exI) apply (simp add: mult.assoc [symmetric] abs_mult) apply (rule_tac x = "\inverse c\ * ca" in exI) apply auto done lemma bigo_const_mult6 [intro]: "(\x. c) *o O(f) \ O(f)" apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) apply (rule_tac x = "ca * \c\" in exI) apply (rule allI) apply (subgoal_tac "ca * \c\ * \f x\ = \c\ * (ca * \f x\)") apply (erule ssubst) apply (subst abs_mult) apply (rule mult_left_mono) apply (erule spec) apply simp apply (simp add: ac_simps) done lemma bigo_const_mult7 [intro]: assumes *: "f =o O(g)" shows "(\x. c * f x) =o O(g)" proof - from * have "(\x. c) * f =o (\x. c) *o O(g)" by auto also have "(\x. c) * f = (\x. c * f x)" by (simp add: func_times) also have "(\x. c) *o O(g) \ O(g)" by (auto del: subsetI) finally show ?thesis . qed lemma bigo_compose1: "f =o O(g) \ (\x. f (k x)) =o O(\x. g (k x))" by (auto simp: bigo_def) lemma bigo_compose2: "f =o g +o O(h) \ (\x. f (k x)) =o (\x. g (k x)) +o O(\x. h(k x))" apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus) apply (drule bigo_compose1) apply (simp add: fun_diff_def) done subsection \Sum\ lemma bigo_sum_main: "\x. \y \ A x. 0 \ h x y \ \c. \x. \y \ A x. \f x y\ \ c * h x y \ (\x. \y \ A x. f x y) =o O(\x. \y \ A x. h x y)" apply (auto simp add: bigo_def) apply (rule_tac x = "\c\" in exI) apply (subst abs_of_nonneg) back back apply (rule sum_nonneg) apply force apply (subst sum_distrib_left) apply (rule allI) apply (rule order_trans) apply (rule sum_abs) apply (rule sum_mono) apply (rule order_trans) apply (drule spec)+ apply (drule bspec)+ apply assumption+ apply (drule bspec) apply assumption+ apply (rule mult_right_mono) apply (rule abs_ge_self) apply force done lemma bigo_sum1: "\x y. 0 \ h x y \ \c. \x y. \f x y\ \ c * h x y \ (\x. \y \ A x. f x y) =o O(\x. \y \ A x. h x y)" apply (rule bigo_sum_main) apply force apply clarsimp apply (rule_tac x = c in exI) apply force done lemma bigo_sum2: "\y. 0 \ h y \ \c. \y. \f y\ \ c * (h y) \ (\x. \y \ A x. f y) =o O(\x. \y \ A x. h y)" by (rule bigo_sum1) auto lemma bigo_sum3: "f =o O(h) \ (\x. \y \ A x. l x y * f (k x y)) =o O(\x. \y \ A x. \l x y * h (k x y)\)" apply (rule bigo_sum1) apply (rule allI)+ apply (rule abs_ge_zero) apply (unfold bigo_def) apply auto apply (rule_tac x = c in exI) apply (rule allI)+ apply (subst abs_mult)+ apply (subst mult.left_commute) apply (rule mult_left_mono) apply (erule spec) apply (rule abs_ge_zero) done lemma bigo_sum4: "f =o g +o O(h) \ (\x. \y \ A x. l x y * f (k x y)) =o (\x. \y \ A x. l x y * g (k x y)) +o O(\x. \y \ A x. \l x y * h (k x y)\)" apply (rule set_minus_imp_plus) apply (subst fun_diff_def) apply (subst sum_subtractf [symmetric]) apply (subst right_diff_distrib [symmetric]) apply (rule bigo_sum3) apply (subst fun_diff_def [symmetric]) apply (erule set_plus_imp_minus) done lemma bigo_sum5: "f =o O(h) \ \x y. 0 \ l x y \ \x. 0 \ h x \ (\x. \y \ A x. l x y * f (k x y)) =o O(\x. \y \ A x. l x y * h (k x y))" apply (subgoal_tac "(\x. \y \ A x. l x y * h (k x y)) = (\x. \y \ A x. \l x y * h (k x y)\)") apply (erule ssubst) apply (erule bigo_sum3) apply (rule ext) apply (rule sum.cong) apply (rule refl) apply (subst abs_of_nonneg) apply auto done lemma bigo_sum6: "f =o g +o O(h) \ \x y. 0 \ l x y \ \x. 0 \ h x \ (\x. \y \ A x. l x y * f (k x y)) =o (\x. \y \ A x. l x y * g (k x y)) +o O(\x. \y \ A x. l x y * h (k x y))" apply (rule set_minus_imp_plus) apply (subst fun_diff_def) apply (subst sum_subtractf [symmetric]) apply (subst right_diff_distrib [symmetric]) apply (rule bigo_sum5) apply (subst fun_diff_def [symmetric]) apply (drule set_plus_imp_minus) apply auto done subsection \Misc useful stuff\ lemma bigo_useful_intro: "A \ O(f) \ B \ O(f) \ A + B \ O(f)" apply (subst bigo_plus_idemp [symmetric]) apply (rule set_plus_mono2) apply assumption+ done lemma bigo_useful_add: "f =o O(h) \ g =o O(h) \ f + g =o O(h)" apply (subst bigo_plus_idemp [symmetric]) apply (rule set_plus_intro) apply assumption+ done lemma bigo_useful_const_mult: "c \ 0 \ (\x. c) * f =o O(h) \ f =o O(h)" for c :: "'a::linordered_field" apply (rule subsetD) apply (subgoal_tac "(\x. 1 / c) *o O(h) \ O(h)") apply assumption apply (rule bigo_const_mult6) apply (subgoal_tac "f = (\x. 1 / c) * ((\x. c) * f)") apply (erule ssubst) apply (erule set_times_intro2) apply (simp add: func_times) done lemma bigo_fix: "(\x::nat. f (x + 1)) =o O(\x. h (x + 1)) \ f 0 = 0 \ f =o O(h)" apply (simp add: bigo_alt_def) apply auto apply (rule_tac x = c in exI) apply auto apply (case_tac "x = 0") apply simp apply (subgoal_tac "x = Suc (x - 1)") apply (erule ssubst) back apply (erule spec) apply simp done lemma bigo_fix2: "(\x. f ((x::nat) + 1)) =o (\x. g(x + 1)) +o O(\x. h(x + 1)) \ f 0 = g 0 \ f =o g +o O(h)" apply (rule set_minus_imp_plus) apply (rule bigo_fix) apply (subst fun_diff_def) apply (subst fun_diff_def [symmetric]) apply (rule set_plus_imp_minus) apply simp apply (simp add: fun_diff_def) done subsection \Less than or equal to\ definition lesso :: "('a \ 'b::linordered_idom) \ ('a \ 'b) \ 'a \ 'b" (infixl "x. max (f x - g x) 0)" lemma bigo_lesseq1: "f =o O(h) \ \x. \g x\ \ \f x\ \ g =o O(h)" apply (unfold bigo_def) apply clarsimp apply (rule_tac x = c in exI) apply (rule allI) apply (rule order_trans) apply (erule spec)+ done lemma bigo_lesseq2: "f =o O(h) \ \x. \g x\ \ f x \ g =o O(h)" apply (erule bigo_lesseq1) apply (rule allI) apply (drule_tac x = x in spec) apply (rule order_trans) apply assumption apply (rule abs_ge_self) done lemma bigo_lesseq3: "f =o O(h) \ \x. 0 \ g x \ \x. g x \ f x \ g =o O(h)" apply (erule bigo_lesseq2) apply (rule allI) apply (subst abs_of_nonneg) apply (erule spec)+ done lemma bigo_lesseq4: "f =o O(h) \ \x. 0 \ g x \ \x. g x \ \f x\ \ g =o O(h)" apply (erule bigo_lesseq1) apply (rule allI) apply (subst abs_of_nonneg) apply (erule spec)+ done lemma bigo_lesso1: "\x. f x \ g x \ f x. max (f x - g x) 0) = 0") apply (erule ssubst) apply (rule bigo_zero) apply (unfold func_zero) apply (rule ext) apply (simp split: split_max) done lemma bigo_lesso2: "f =o g +o O(h) \ \x. 0 \ k x \ \x. k x \ f x \ k k x - g x") apply simp apply (subst abs_of_nonneg) apply (drule_tac x = x in spec) back apply (simp add: algebra_simps) apply (subst diff_conv_add_uminus)+ apply (rule add_right_mono) apply (erule spec) apply (rule order_trans) prefer 2 apply (rule abs_ge_zero) apply (simp add: algebra_simps) done lemma bigo_lesso3: "f =o g +o O(h) \ \x. 0 \ k x \ \x. g x \ k x \ f f x - k x") apply simp apply (subst abs_of_nonneg) apply (drule_tac x = x in spec) back apply (simp add: algebra_simps) apply (subst diff_conv_add_uminus)+ apply (rule add_left_mono) apply (rule le_imp_neg_le) apply (erule spec) apply (rule order_trans) prefer 2 apply (rule abs_ge_zero) apply (simp add: algebra_simps) done lemma bigo_lesso4: "f g =o h +o O(k) \ f 'b::linordered_field" apply (unfold lesso_def) apply (drule set_plus_imp_minus) apply (drule bigo_abs5) back apply (simp add: fun_diff_def) apply (drule bigo_useful_add) apply assumption apply (erule bigo_lesseq2) back apply (rule allI) apply (auto simp add: func_plus fun_diff_def algebra_simps split: split_max abs_split) done lemma bigo_lesso5: "f \C. \x. f x \ g x + C * \h x\" apply (simp only: lesso_def bigo_alt_def) apply clarsimp apply (rule_tac x = c in exI) apply (rule allI) apply (drule_tac x = x in spec) apply (subgoal_tac "\max (f x - g x) 0\ = max (f x - g x) 0") apply (clarsimp simp add: algebra_simps) apply (rule abs_of_nonneg) apply (rule max.cobounded2) done lemma lesso_add: "f k (f + k) g \ 0 \ f \ 0" for f g :: "nat \ real" apply (simp add: LIMSEQ_iff bigo_alt_def) apply clarify apply (drule_tac x = "r / c" in spec) apply (drule mp) apply simp apply clarify apply (rule_tac x = no in exI) apply (rule allI) apply (drule_tac x = n in spec)+ apply (rule impI) apply (drule mp) apply assumption apply (rule order_le_less_trans) apply assumption apply (rule order_less_le_trans) apply (subgoal_tac "c * \g n\ < c * (r / c)") apply assumption apply (erule mult_strict_left_mono) apply assumption apply simp done lemma bigo_LIMSEQ2: "f =o g +o O(h) \ h \ 0 \ f \ a \ g \ a" for f g h :: "nat \ real" apply (drule set_plus_imp_minus) apply (drule bigo_LIMSEQ1) apply assumption apply (simp only: fun_diff_def) apply (erule Lim_transform2) apply assumption done end diff --git a/src/HOL/Library/Code_Lazy.thy b/src/HOL/Library/Code_Lazy.thy --- a/src/HOL/Library/Code_Lazy.thy +++ b/src/HOL/Library/Code_Lazy.thy @@ -1,238 +1,238 @@ (* Author: Pascal Stoop, ETH Zurich Author: Andreas Lochbihler, Digital Asset *) section \Lazy types in generated code\ theory Code_Lazy imports Case_Converter keywords "code_lazy_type" "activate_lazy_type" "deactivate_lazy_type" "activate_lazy_types" "deactivate_lazy_types" "print_lazy_types" :: thy_decl begin text \ - This theory and the CodeLazy tool described in @{cite "LochbihlerStoop2018"}. + This theory and the CodeLazy tool described in \<^cite>\"LochbihlerStoop2018"\. It hooks into Isabelle's code generator such that the generated code evaluates a user-specified set of type constructors lazily, even in target languages with eager evaluation. The lazy type must be algebraic, i.e., values must be built from constructors and a corresponding case operator decomposes them. Every datatype and codatatype is algebraic and thus eligible for lazification. \ subsection \The type \lazy\\ typedef 'a lazy = "UNIV :: 'a set" .. setup_lifting type_definition_lazy lift_definition delay :: "(unit \ 'a) \ 'a lazy" is "\f. f ()" . lift_definition force :: "'a lazy \ 'a" is "\x. x" . code_datatype delay lemma force_delay [code]: "force (delay f) = f ()" by transfer (rule refl) lemma delay_force: "delay (\_. force s) = s" by transfer (rule refl) definition termify_lazy2 :: "'a :: typerep lazy \ term" where "termify_lazy2 x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Code_Lazy.delay'') (TYPEREP((unit \ 'a) \ 'a lazy))) (Code_Evaluation.Const (STR ''Pure.dummy_pattern'') (TYPEREP((unit \ 'a))))" definition termify_lazy :: "(String.literal \ 'typerep \ 'term) \ ('term \ 'term \ 'term) \ (String.literal \ 'typerep \ 'term \ 'term) \ 'typerep \ ('typerep \ 'typerep \ 'typerep) \ ('typerep \ 'typerep) \ ('a \ 'term) \ 'typerep \ 'a :: typerep lazy \ 'term \ term" where "termify_lazy _ _ _ _ _ _ _ _ x _ = termify_lazy2 x" declare [[code drop: "Code_Evaluation.term_of :: _ lazy \ _"]] lemma term_of_lazy_code [code]: "Code_Evaluation.term_of x \ termify_lazy Code_Evaluation.Const Code_Evaluation.App Code_Evaluation.Abs TYPEREP(unit) (\T U. typerep.Typerep (STR ''fun'') [T, U]) (\T. typerep.Typerep (STR ''Code_Lazy.lazy'') [T]) Code_Evaluation.term_of TYPEREP('a) x (Code_Evaluation.Const (STR '''') (TYPEREP(unit)))" for x :: "'a :: {typerep, term_of} lazy" by (rule term_of_anything) text \ The implementations of \<^typ>\_ lazy\ using language primitives cache forced values. Term reconstruction for lazy looks into the lazy value and reconstructs it to the depth it has been evaluated. This is not done for Haskell as we do not know of any portable way to inspect whether a lazy value has been evaluated to or not. \ code_printing code_module Lazy \ (SML) \signature LAZY = sig type 'a lazy; val lazy : (unit -> 'a) -> 'a lazy; val force : 'a lazy -> 'a; val peek : 'a lazy -> 'a option val termify_lazy : (string -> 'typerep -> 'term) -> ('term -> 'term -> 'term) -> (string -> 'typerep -> 'term -> 'term) -> 'typerep -> ('typerep -> 'typerep -> 'typerep) -> ('typerep -> 'typerep) -> ('a -> 'term) -> 'typerep -> 'a lazy -> 'term -> 'term; end; structure Lazy : LAZY = struct datatype 'a content = Delay of unit -> 'a | Value of 'a | Exn of exn; datatype 'a lazy = Lazy of 'a content ref; fun lazy f = Lazy (ref (Delay f)); fun force (Lazy x) = case !x of Delay f => ( let val res = f (); val _ = x := Value res; in res end handle exn => (x := Exn exn; raise exn)) | Value x => x | Exn exn => raise exn; fun peek (Lazy x) = case !x of Value x => SOME x | _ => NONE; fun termify_lazy const app abs unitT funT lazyT term_of T x _ = app (const "Code_Lazy.delay" (funT (funT unitT T) (lazyT T))) (case peek x of SOME y => abs "_" unitT (term_of y) | _ => const "Pure.dummy_pattern" (funT unitT T)); end;\ for type_constructor lazy constant delay force termify_lazy | type_constructor lazy \ (SML) "_ Lazy.lazy" | constant delay \ (SML) "Lazy.lazy" | constant force \ (SML) "Lazy.force" | constant termify_lazy \ (SML) "Lazy.termify'_lazy" code_reserved SML Lazy code_printing \ \For code generation within the Isabelle environment, we reuse the thread-safe implementation of lazy from \<^file>\~~/src/Pure/Concurrent/lazy.ML\\ code_module Lazy \ (Eval) \\ for constant undefined | type_constructor lazy \ (Eval) "_ Lazy.lazy" | constant delay \ (Eval) "Lazy.lazy" | constant force \ (Eval) "Lazy.force" | code_module Termify_Lazy \ (Eval) \structure Termify_Lazy = struct fun termify_lazy (_: string -> typ -> term) (_: term -> term -> term) (_: string -> typ -> term -> term) (_: typ) (_: typ -> typ -> typ) (_: typ -> typ) (term_of: 'a -> term) (T: typ) (x: 'a Lazy.lazy) (_: term) = Const ("Code_Lazy.delay", (HOLogic.unitT --> T) --> Type ("Code_Lazy.lazy", [T])) $ (case Lazy.peek x of SOME (Exn.Res x) => absdummy HOLogic.unitT (term_of x) | _ => Const ("Pure.dummy_pattern", HOLogic.unitT --> T)); end;\ for constant termify_lazy | constant termify_lazy \ (Eval) "Termify'_Lazy.termify'_lazy" code_reserved Eval Termify_Lazy code_printing type_constructor lazy \ (OCaml) "_ Lazy.t" | constant delay \ (OCaml) "Lazy.from'_fun" | constant force \ (OCaml) "Lazy.force" | code_module Termify_Lazy \ (OCaml) \module Termify_Lazy : sig val termify_lazy : (string -> 'typerep -> 'term) -> ('term -> 'term -> 'term) -> (string -> 'typerep -> 'term -> 'term) -> 'typerep -> ('typerep -> 'typerep -> 'typerep) -> ('typerep -> 'typerep) -> ('a -> 'term) -> 'typerep -> 'a Lazy.t -> 'term -> 'term end = struct let termify_lazy const app abs unitT funT lazyT term_of ty x _ = app (const "Code_Lazy.delay" (funT (funT unitT ty) (lazyT ty))) (if Lazy.is_val x then abs "_" unitT (term_of (Lazy.force x)) else const "Pure.dummy_pattern" (funT unitT ty));; end;;\ for constant termify_lazy | constant termify_lazy \ (OCaml) "Termify'_Lazy.termify'_lazy" code_reserved OCaml Lazy Termify_Lazy code_printing code_module Lazy \ (Haskell) \ module Lazy(Lazy, delay, force) where newtype Lazy a = Lazy a delay f = Lazy (f ()) force (Lazy x) = x\ for type_constructor lazy constant delay force | type_constructor lazy \ (Haskell) "Lazy.Lazy _" | constant delay \ (Haskell) "Lazy.delay" | constant force \ (Haskell) "Lazy.force" code_reserved Haskell Lazy code_printing code_module Lazy \ (Scala) \object Lazy { final class Lazy[A] (f: Unit => A) { var evaluated = false; lazy val x: A = f(()) def get() : A = { evaluated = true; return x } } def force[A] (x: Lazy[A]) : A = { return x.get() } def delay[A] (f: Unit => A) : Lazy[A] = { return new Lazy[A] (f) } def termify_lazy[Typerep, Term, A] ( const: String => Typerep => Term, app: Term => Term => Term, abs: String => Typerep => Term => Term, unitT: Typerep, funT: Typerep => Typerep => Typerep, lazyT: Typerep => Typerep, term_of: A => Term, ty: Typerep, x: Lazy[A], dummy: Term) : Term = { if (x.evaluated) app(const("Code_Lazy.delay")(funT(funT(unitT)(ty))(lazyT(ty))))(abs("_")(unitT)(term_of(x.get()))) else app(const("Code_Lazy.delay")(funT(funT(unitT)(ty))(lazyT(ty))))(const("Pure.dummy_pattern")(funT(unitT)(ty))) } }\ for type_constructor lazy constant delay force termify_lazy | type_constructor lazy \ (Scala) "Lazy.Lazy[_]" | constant delay \ (Scala) "Lazy.delay" | constant force \ (Scala) "Lazy.force" | constant termify_lazy \ (Scala) "Lazy.termify'_lazy" code_reserved Scala Lazy text \Make evaluation with the simplifier respect \<^term>\delay\s.\ lemma delay_lazy_cong: "delay f = delay f" by simp setup \Code_Simp.map_ss (Simplifier.add_cong @{thm delay_lazy_cong})\ subsection \Implementation\ ML_file \code_lazy.ML\ setup \ Code_Preproc.add_functrans ("lazy_datatype", Code_Lazy.transform_code_eqs) \ end diff --git a/src/HOL/Library/Poly_Mapping.thy b/src/HOL/Library/Poly_Mapping.thy --- a/src/HOL/Library/Poly_Mapping.thy +++ b/src/HOL/Library/Poly_Mapping.thy @@ -1,1878 +1,1878 @@ (* Author: Andreas Lochbihler, ETH Zurich Author: Florian Haftmann, TU Muenchen with some material ported from HOL Light by LCP *) section \Polynomial mapping: combination of almost everywhere zero functions with an algebraic view\ theory Poly_Mapping imports Groups_Big_Fun Fun_Lexorder More_List begin subsection \Preliminary: auxiliary operations for \emph{almost everywhere zero}\ text \ A central notion for polynomials are functions being \emph{almost everywhere zero}. For these we provide some auxiliary definitions and lemmas. \ lemma finite_mult_not_eq_zero_leftI: fixes f :: "'b \ 'a :: mult_zero" assumes "finite {a. f a \ 0}" shows "finite {a. g a * f a \ 0}" proof - have "{a. g a * f a \ 0} \ {a. f a \ 0}" by auto then show ?thesis using assms by (rule finite_subset) qed lemma finite_mult_not_eq_zero_rightI: fixes f :: "'b \ 'a :: mult_zero" assumes "finite {a. f a \ 0}" shows "finite {a. f a * g a \ 0}" proof - have "{a. f a * g a \ 0} \ {a. f a \ 0}" by auto then show ?thesis using assms by (rule finite_subset) qed lemma finite_mult_not_eq_zero_prodI: fixes f g :: "'a \ 'b::semiring_0" assumes "finite {a. f a \ 0}" (is "finite ?F") assumes "finite {b. g b \ 0}" (is "finite ?G") shows "finite {(a, b). f a * g b \ 0}" proof - from assms have "finite (?F \ ?G)" by blast then have "finite {(a, b). f a \ 0 \ g b \ 0}" by simp then show ?thesis by (rule rev_finite_subset) auto qed lemma finite_not_eq_zero_sumI: fixes f g :: "'a::monoid_add \ 'b::semiring_0" assumes "finite {a. f a \ 0}" (is "finite ?F") assumes "finite {b. g b \ 0}" (is "finite ?G") shows "finite {a + b | a b. f a \ 0 \ g b \ 0}" (is "finite ?FG") proof - from assms have "finite (?F \ ?G)" by (simp add: finite_cartesian_product_iff) then have "finite (case_prod plus ` (?F \ ?G))" by (rule finite_imageI) also have "case_prod plus ` (?F \ ?G) = ?FG" by auto finally show ?thesis by simp qed lemma finite_mult_not_eq_zero_sumI: fixes f g :: "'a::monoid_add \ 'b::semiring_0" assumes "finite {a. f a \ 0}" assumes "finite {b. g b \ 0}" shows "finite {a + b | a b. f a * g b \ 0}" proof - from assms have "finite {a + b | a b. f a \ 0 \ g b \ 0}" by (rule finite_not_eq_zero_sumI) then show ?thesis by (rule rev_finite_subset) (auto dest: mult_not_zero) qed lemma finite_Sum_any_not_eq_zero_weakenI: assumes "finite {a. \b. f a b \ 0}" shows "finite {a. Sum_any (f a) \ 0}" proof - have "{a. Sum_any (f a) \ 0} \ {a. \b. f a b \ 0}" by (auto elim: Sum_any.not_neutral_obtains_not_neutral) then show ?thesis using assms by (rule finite_subset) qed context zero begin definition "when" :: "'a \ bool \ 'a" (infixl "when" 20) where "(a when P) = (if P then a else 0)" text \ Case distinctions always complicate matters, particularly when nested. The @{const "when"} operation allows to minimise these if @{term 0} is the false-case value and makes proof obligations much more readable. \ lemma "when" [simp]: "P \ (a when P) = a" "\ P \ (a when P) = 0" by (simp_all add: when_def) lemma when_simps [simp]: "(a when True) = a" "(a when False) = 0" by simp_all lemma when_cong: assumes "P \ Q" and "Q \ a = b" shows "(a when P) = (b when Q)" using assms by (simp add: when_def) lemma zero_when [simp]: "(0 when P) = 0" by (simp add: when_def) lemma when_when: "(a when P when Q) = (a when P \ Q)" by (cases Q) simp_all lemma when_commute: "(a when Q when P) = (a when P when Q)" by (simp add: when_when conj_commute) lemma when_neq_zero [simp]: "(a when P) \ 0 \ P \ a \ 0" by (cases P) simp_all end context monoid_add begin lemma when_add_distrib: "(a + b when P) = (a when P) + (b when P)" by (simp add: when_def) end context semiring_1 begin lemma zero_power_eq: "0 ^ n = (1 when n = 0)" by (simp add: power_0_left) end context comm_monoid_add begin lemma Sum_any_when_equal [simp]: "(\a. (f a when a = b)) = f b" by (simp add: when_def) lemma Sum_any_when_equal' [simp]: "(\a. (f a when b = a)) = f b" by (simp add: when_def) lemma Sum_any_when_independent: "(\a. g a when P) = ((\a. g a) when P)" by (cases P) simp_all lemma Sum_any_when_dependent_prod_right: "(\(a, b). g a when b = h a) = (\a. g a)" proof - have "inj_on (\a. (a, h a)) {a. g a \ 0}" by (rule inj_onI) auto then show ?thesis unfolding Sum_any.expand_set by (rule sum.reindex_cong) auto qed lemma Sum_any_when_dependent_prod_left: "(\(a, b). g b when a = h b) = (\b. g b)" proof - have "(\(a, b). g b when a = h b) = (\(b, a). g b when a = h b)" by (rule Sum_any.reindex_cong [of prod.swap]) (simp_all add: fun_eq_iff) then show ?thesis by (simp add: Sum_any_when_dependent_prod_right) qed end context cancel_comm_monoid_add begin lemma when_diff_distrib: "(a - b when P) = (a when P) - (b when P)" by (simp add: when_def) end context group_add begin lemma when_uminus_distrib: "(- a when P) = - (a when P)" by (simp add: when_def) end context mult_zero begin lemma mult_when: "a * (b when P) = (a * b when P)" by (cases P) simp_all lemma when_mult: "(a when P) * b = (a * b when P)" by (cases P) simp_all end subsection \Type definition\ text \ The following type is of central importance: \ typedef (overloaded) ('a, 'b) poly_mapping ("(_ \\<^sub>0 /_)" [1, 0] 0) = "{f :: 'a \ 'b::zero. finite {x. f x \ 0}}" morphisms lookup Abs_poly_mapping proof - have "(\_::'a. (0 :: 'b)) \ ?poly_mapping" by simp then show ?thesis by (blast intro!: exI) qed declare lookup_inverse [simp] declare lookup_inject [simp] lemma lookup_Abs_poly_mapping [simp]: "finite {x. f x \ 0} \ lookup (Abs_poly_mapping f) = f" using Abs_poly_mapping_inverse [of f] by simp lemma finite_lookup [simp]: "finite {k. lookup f k \ 0}" using poly_mapping.lookup [of f] by simp lemma finite_lookup_nat [simp]: fixes f :: "'a \\<^sub>0 nat" shows "finite {k. 0 < lookup f k}" using poly_mapping.lookup [of f] by simp lemma poly_mapping_eqI: assumes "\k. lookup f k = lookup g k" shows "f = g" using assms unfolding poly_mapping.lookup_inject [symmetric] by blast lemma poly_mapping_eq_iff: "a = b \ lookup a = lookup b" by auto text \ We model the universe of functions being \emph{almost everywhere zero} by means of a separate type @{typ "('a, 'b) poly_mapping"}. For convenience we provide a suggestive infix syntax which is a variant of the usual function space syntax. Conversion between both types happens through the morphisms \begin{quote} @{term_type lookup} \end{quote} \begin{quote} @{term_type Abs_poly_mapping} \end{quote} satisfying \begin{quote} @{thm lookup_inverse} \end{quote} \begin{quote} @{thm lookup_Abs_poly_mapping} \end{quote} Luckily, we have rarely to deal with those low-level morphisms explicitly but rely on Isabelle's \emph{lifting} package with its method \transfer\ and its specification tool \lift_definition\. \ setup_lifting type_definition_poly_mapping code_datatype Abs_poly_mapping\\FIXME? workaround for preventing \code_abstype\ setup\ text \ @{typ "'a \\<^sub>0 'b"} serves distinctive purposes: \begin{enumerate} \item A clever nesting as @{typ "(nat \\<^sub>0 nat) \\<^sub>0 'a"} later in theory \MPoly\ gives a suitable representation type for polynomials \emph{almost for free}: Interpreting @{typ "nat \\<^sub>0 nat"} as a mapping from variable identifiers to exponents yields monomials, and the whole type maps monomials to coefficients. Lets call this the \emph{ultimate interpretation}. \item A further more specialised type isomorphic to @{typ "nat \\<^sub>0 'a"} is apt to direct implementation using code generation - \cite{Haftmann-Nipkow:2010:code}. + \<^cite>\"Haftmann-Nipkow:2010:code"\. \end{enumerate} Note that despite the names \emph{mapping} and \emph{lookup} suggest something implementation-near, it is best to keep @{typ "'a \\<^sub>0 'b"} as an abstract \emph{algebraic} type providing operations like \emph{addition}, \emph{multiplication} without any notion of key-order, data structures etc. This implementations-specific notions are easily introduced later for particular implementations but do not provide any gain for specifying logical properties of polynomials. \ subsection \Additive structure\ text \ The additive structure covers the usual operations \0\, \+\ and (unary and binary) \-\. Recalling the ultimate interpretation, it is obvious that these have just lift the corresponding operations on values to mappings. Isabelle has a rich hierarchy of algebraic type classes, and in such situations of pointwise lifting a typical pattern is to have instantiations for a considerable number of type classes. The operations themselves are specified using \lift_definition\, where the proofs of the \emph{almost everywhere zero} property can be significantly involved. The @{const lookup} operation is supposed to be usable explicitly (unless in other situations where the morphisms between types are somehow internal to the \emph{lifting} package). Hence it is good style to provide explicit rewrite rules how @{const lookup} acts on operations immediately. \ instantiation poly_mapping :: (type, zero) zero begin lift_definition zero_poly_mapping :: "'a \\<^sub>0 'b" is "\k. 0" by simp instance .. end lemma Abs_poly_mapping [simp]: "Abs_poly_mapping (\k. 0) = 0" by (simp add: zero_poly_mapping.abs_eq) lemma lookup_zero [simp]: "lookup 0 k = 0" by transfer rule instantiation poly_mapping :: (type, monoid_add) monoid_add begin lift_definition plus_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is "\f1 f2 k. f1 k + f2 k" proof - fix f1 f2 :: "'a \ 'b" assume "finite {k. f1 k \ 0}" and "finite {k. f2 k \ 0}" then have "finite ({k. f1 k \ 0} \ {k. f2 k \ 0})" by auto moreover have "{x. f1 x + f2 x \ 0} \ {k. f1 k \ 0} \ {k. f2 k \ 0}" by auto ultimately show "finite {x. f1 x + f2 x \ 0}" by (blast intro: finite_subset) qed instance by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ end lemma lookup_add: "lookup (f + g) k = lookup f k + lookup g k" by transfer rule instance poly_mapping :: (type, comm_monoid_add) comm_monoid_add by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ lemma lookup_sum: "lookup (sum pp X) i = sum (\x. lookup (pp x) i) X" by (induction rule: infinite_finite_induct) (auto simp: lookup_add) (*instance poly_mapping :: (type, "{monoid_add, cancel_semigroup_add}") cancel_semigroup_add by intro_classes (transfer, simp add: fun_eq_iff)+*) instantiation poly_mapping :: (type, cancel_comm_monoid_add) cancel_comm_monoid_add begin lift_definition minus_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is "\f1 f2 k. f1 k - f2 k" proof - fix f1 f2 :: "'a \ 'b" assume "finite {k. f1 k \ 0}" and "finite {k. f2 k \ 0}" then have "finite ({k. f1 k \ 0} \ {k. f2 k \ 0})" by auto moreover have "{x. f1 x - f2 x \ 0} \ {k. f1 k \ 0} \ {k. f2 k \ 0}" by auto ultimately show "finite {x. f1 x - f2 x \ 0}" by (blast intro: finite_subset) qed instance by intro_classes (transfer, simp add: fun_eq_iff diff_diff_add)+ end instantiation poly_mapping :: (type, ab_group_add) ab_group_add begin lift_definition uminus_poly_mapping :: "('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is uminus by simp instance by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ end lemma lookup_uminus [simp]: "lookup (- f) k = - lookup f k" by transfer simp lemma lookup_minus: "lookup (f - g) k = lookup f k - lookup g k" by transfer rule subsection \Multiplicative structure\ instantiation poly_mapping :: (zero, zero_neq_one) zero_neq_one begin lift_definition one_poly_mapping :: "'a \\<^sub>0 'b" is "\k. 1 when k = 0" by simp instance by intro_classes (transfer, simp add: fun_eq_iff) end lemma lookup_one: "lookup 1 k = (1 when k = 0)" by transfer rule lemma lookup_one_zero [simp]: "lookup 1 0 = 1" by transfer simp definition prod_fun :: "('a \ 'b) \ ('a \ 'b) \ 'a::monoid_add \ 'b::semiring_0" where "prod_fun f1 f2 k = (\l. f1 l * (\q. (f2 q when k = l + q)))" lemma prod_fun_unfold_prod: fixes f g :: "'a :: monoid_add \ 'b::semiring_0" assumes fin_f: "finite {a. f a \ 0}" assumes fin_g: "finite {b. g b \ 0}" shows "prod_fun f g k = (\(a, b). f a * g b when k = a + b)" proof - let ?C = "{a. f a \ 0} \ {b. g b \ 0}" from fin_f fin_g have "finite ?C" by blast moreover have "{a. \b. (f a * g b when k = a + b) \ 0} \ {b. \a. (f a * g b when k = a + b) \ 0} \ {a. f a \ 0} \ {b. g b \ 0}" by auto ultimately show ?thesis using fin_g by (auto simp add: prod_fun_def Sum_any.cartesian_product [of "{a. f a \ 0} \ {b. g b \ 0}"] Sum_any_right_distrib mult_when) qed lemma finite_prod_fun: fixes f1 f2 :: "'a :: monoid_add \ 'b :: semiring_0" assumes fin1: "finite {l. f1 l \ 0}" and fin2: "finite {q. f2 q \ 0}" shows "finite {k. prod_fun f1 f2 k \ 0}" proof - have *: "finite {k. (\l. f1 l \ 0 \ (\q. f2 q \ 0 \ k = l + q))}" using assms by simp { fix k l have "{q. (f2 q when k = l + q) \ 0} \ {q. f2 q \ 0 \ k = l + q}" by auto with fin2 have "sum f2 {q. f2 q \ 0 \ k = l + q} = (\q. (f2 q when k = l + q))" by (simp add: Sum_any.expand_superset [of "{q. f2 q \ 0 \ k = l + q}"]) } note aux = this have "{k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q}) \ 0} \ {k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q} \ 0)}" by (auto elim!: Sum_any.not_neutral_obtains_not_neutral) also have "\ \ {k. (\l. f1 l \ 0 \ sum f2 {q. f2 q \ 0 \ k = l + q} \ 0)}" by (auto dest: mult_not_zero) also have "\ \ {k. (\l. f1 l \ 0 \ (\q. f2 q \ 0 \ k = l + q))}" by (auto elim!: sum.not_neutral_contains_not_neutral) finally have "finite {k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q}) \ 0}" using * by (rule finite_subset) with aux have "finite {k. (\l. f1 l * (\q. (f2 q when k = l + q))) \ 0}" by simp with fin2 show ?thesis by (simp add: prod_fun_def) qed instantiation poly_mapping :: (monoid_add, semiring_0) semiring_0 begin lift_definition times_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is prod_fun by(rule finite_prod_fun) instance proof fix a b c :: "'a \\<^sub>0 'b" show "a * b * c = a * (b * c)" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {a. f a \ 0}" (is "finite ?F") assume fin_g: "finite {b. g b \ 0}" (is "finite ?G") assume fin_h: "finite {c. h c \ 0}" (is "finite ?H") from fin_f fin_g have fin_fg: "finite {(a, b). f a * g b \ 0}" (is "finite ?FG") by (rule finite_mult_not_eq_zero_prodI) from fin_g fin_h have fin_gh: "finite {(b, c). g b * h c \ 0}" (is "finite ?GH") by (rule finite_mult_not_eq_zero_prodI) from fin_f fin_g have fin_fg': "finite {a + b | a b. f a * g b \ 0}" (is "finite ?FG'") by (rule finite_mult_not_eq_zero_sumI) then have fin_fg'': "finite {d. (\(a, b). f a * g b when d = a + b) \ 0}" by (auto intro: finite_Sum_any_not_eq_zero_weakenI) from fin_g fin_h have fin_gh': "finite {b + c | b c. g b * h c \ 0}" (is "finite ?GH'") by (rule finite_mult_not_eq_zero_sumI) then have fin_gh'': "finite {d. (\(b, c). g b * h c when d = b + c) \ 0}" by (auto intro: finite_Sum_any_not_eq_zero_weakenI) show "prod_fun (prod_fun f g) h = prod_fun f (prod_fun g h)" (is "?lhs = ?rhs") proof fix k from fin_f fin_g fin_h fin_fg'' have "?lhs k = (\(ab, c). (\(a, b). f a * g b when ab = a + b) * h c when k = ab + c)" by (simp add: prod_fun_unfold_prod) also have "\ = (\(ab, c). (\(a, b). f a * g b * h c when k = ab + c when ab = a + b))" apply (subst Sum_any_left_distrib) using fin_fg apply (simp add: split_def) apply (subst Sum_any_when_independent [symmetric]) apply (simp add: when_when when_mult mult_when split_def conj_commute) done also have "\ = (\(ab, c, a, b). f a * g b * h c when k = ab + c when ab = a + b)" apply (subst Sum_any.cartesian_product2 [of "(?FG' \ ?H) \ ?FG"]) apply (auto simp add: finite_cartesian_product_iff fin_fg fin_fg' fin_h dest: mult_not_zero) done also have "\ = (\(ab, c, a, b). f a * g b * h c when k = a + b + c when ab = a + b)" by (rule Sum_any.cong) (simp add: split_def when_def) also have "\ = (\(ab, cab). (case cab of (c, a, b) \ f a * g b * h c when k = a + b + c) when ab = (case cab of (c, a, b) \ a + b))" by (simp add: split_def) also have "\ = (\(c, a, b). f a * g b * h c when k = a + b + c)" by (simp add: Sum_any_when_dependent_prod_left) also have "\ = (\(bc, cab). (case cab of (c, a, b) \ f a * g b * h c when k = a + b + c) when bc = (case cab of (c, a, b) \ b + c))" by (simp add: Sum_any_when_dependent_prod_left) also have "\ = (\(bc, c, a, b). f a * g b * h c when k = a + b + c when bc = b + c)" by (simp add: split_def) also have "\ = (\(bc, c, a, b). f a * g b * h c when bc = b + c when k = a + bc)" by (rule Sum_any.cong) (simp add: split_def when_def ac_simps) also have "\ = (\(a, bc, b, c). f a * g b * h c when bc = b + c when k = a + bc)" proof - have "bij (\(a, d, b, c). (d, c, a, b))" by (auto intro!: bijI injI surjI [of _ "\(d, c, a, b). (a, d, b, c)"] simp add: split_def) then show ?thesis by (rule Sum_any.reindex_cong) auto qed also have "\ = (\(a, bc). (\(b, c). f a * g b * h c when bc = b + c when k = a + bc))" apply (subst Sum_any.cartesian_product2 [of "(?F \ ?GH') \ ?GH"]) apply (auto simp add: finite_cartesian_product_iff fin_f fin_gh fin_gh' ac_simps dest: mult_not_zero) done also have "\ = (\(a, bc). f a * (\(b, c). g b * h c when bc = b + c) when k = a + bc)" apply (subst Sum_any_right_distrib) using fin_gh apply (simp add: split_def) apply (subst Sum_any_when_independent [symmetric]) apply (simp add: when_when when_mult mult_when split_def ac_simps) done also from fin_f fin_g fin_h fin_gh'' have "\ = ?rhs k" by (simp add: prod_fun_unfold_prod) finally show "?lhs k = ?rhs k" . qed qed show "(a + b) * c = a * c + b * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun (\k. f k + g k) h = (\k. prod_fun f h k + prod_fun g h k)" apply (rule ext) apply (auto simp add: prod_fun_def algebra_simps) apply (subst Sum_any.distrib) using fin_f fin_g apply (auto intro: finite_mult_not_eq_zero_rightI) done qed show "a * (b + c) = a * b + a * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun f (\k. g k + h k) = (\k. prod_fun f g k + prod_fun f h k)" apply (rule ext) apply (auto simp add: prod_fun_def Sum_any.distrib algebra_simps when_add_distrib) apply (subst Sum_any.distrib) apply (simp_all add: algebra_simps) apply (auto intro: fin_g fin_h) apply (subst Sum_any.distrib) apply (simp_all add: algebra_simps) using fin_f apply (rule finite_mult_not_eq_zero_rightI) using fin_f apply (rule finite_mult_not_eq_zero_rightI) done qed show "0 * a = 0" by transfer (simp add: prod_fun_def [abs_def]) show "a * 0 = 0" by transfer (simp add: prod_fun_def [abs_def]) qed end lemma lookup_mult: "lookup (f * g) k = (\l. lookup f l * (\q. lookup g q when k = l + q))" by transfer (simp add: prod_fun_def) instance poly_mapping :: (comm_monoid_add, comm_semiring_0) comm_semiring_0 proof fix a b c :: "'a \\<^sub>0 'b" show "a * b = b * a" proof transfer fix f g :: "'a \ 'b" assume fin_f: "finite {a. f a \ 0}" assume fin_g: "finite {b. g b \ 0}" show "prod_fun f g = prod_fun g f" proof fix k have fin1: "\l. finite {a. (f a when k = l + a) \ 0}" using fin_f by auto have fin2: "\l. finite {b. (g b when k = l + b) \ 0}" using fin_g by auto from fin_f fin_g have "finite {(a, b). f a \ 0 \ g b \ 0}" (is "finite ?AB") by simp show "prod_fun f g k = prod_fun g f k" apply (simp add: prod_fun_def) apply (subst Sum_any_right_distrib) apply (rule fin2) apply (subst Sum_any_right_distrib) apply (rule fin1) apply (subst Sum_any.swap [of ?AB]) apply (fact \finite ?AB\) apply (auto simp add: mult_when ac_simps) done qed qed show "(a + b) * c = a * c + b * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun (\k. f k + g k) h = (\k. prod_fun f h k + prod_fun g h k)" apply (auto simp add: prod_fun_def fun_eq_iff algebra_simps) apply (subst Sum_any.distrib) using fin_f apply (rule finite_mult_not_eq_zero_rightI) using fin_g apply (rule finite_mult_not_eq_zero_rightI) apply simp_all done qed qed instance poly_mapping :: (monoid_add, semiring_0_cancel) semiring_0_cancel .. instance poly_mapping :: (comm_monoid_add, comm_semiring_0_cancel) comm_semiring_0_cancel .. instance poly_mapping :: (monoid_add, semiring_1) semiring_1 proof fix a :: "'a \\<^sub>0 'b" show "1 * a = a" by transfer (simp add: prod_fun_def [abs_def] when_mult) show "a * 1 = a" apply transfer apply (simp add: prod_fun_def [abs_def] Sum_any_right_distrib Sum_any_left_distrib mult_when) apply (subst when_commute) apply simp done qed instance poly_mapping :: (comm_monoid_add, comm_semiring_1) comm_semiring_1 proof fix a :: "'a \\<^sub>0 'b" show "1 * a = a" by transfer (simp add: prod_fun_def [abs_def]) qed instance poly_mapping :: (monoid_add, semiring_1_cancel) semiring_1_cancel .. instance poly_mapping :: (monoid_add, ring) ring .. instance poly_mapping :: (comm_monoid_add, comm_ring) comm_ring .. instance poly_mapping :: (monoid_add, ring_1) ring_1 .. instance poly_mapping :: (comm_monoid_add, comm_ring_1) comm_ring_1 .. subsection \Single-point mappings\ lift_definition single :: "'a \ 'b \ 'a \\<^sub>0 'b::zero" is "\k v k'. (v when k = k')" by simp lemma inj_single [iff]: "inj (single k)" proof (rule injI, transfer) fix k :: 'b and a b :: "'a::zero" assume "(\k'. a when k = k') = (\k'. b when k = k')" then have "(\k'. a when k = k') k = (\k'. b when k = k') k" by (rule arg_cong) then show "a = b" by simp qed lemma lookup_single: "lookup (single k v) k' = (v when k = k')" by transfer rule lemma lookup_single_eq [simp]: "lookup (single k v) k = v" by transfer simp lemma lookup_single_not_eq: "k \ k' \ lookup (single k v) k' = 0" by transfer simp lemma single_zero [simp]: "single k 0 = 0" by transfer simp lemma single_one [simp]: "single 0 1 = 1" by transfer simp lemma single_add: "single k (a + b) = single k a + single k b" by transfer (simp add: fun_eq_iff when_add_distrib) lemma single_uminus: "single k (- a) = - single k a" by transfer (simp add: fun_eq_iff when_uminus_distrib) lemma single_diff: "single k (a - b) = single k a - single k b" by transfer (simp add: fun_eq_iff when_diff_distrib) lemma single_numeral [simp]: "single 0 (numeral n) = numeral n" by (induct n) (simp_all only: numeral.simps numeral_add single_zero single_one single_add) lemma lookup_numeral: "lookup (numeral n) k = (numeral n when k = 0)" proof - have "lookup (numeral n) k = lookup (single 0 (numeral n)) k" by simp then show ?thesis unfolding lookup_single by simp qed lemma single_of_nat [simp]: "single 0 (of_nat n) = of_nat n" by (induct n) (simp_all add: single_add) lemma lookup_of_nat: "lookup (of_nat n) k = (of_nat n when k = 0)" proof - have "lookup (of_nat n) k = lookup (single 0 (of_nat n)) k" by simp then show ?thesis unfolding lookup_single by simp qed lemma of_nat_single: "of_nat = single 0 \ of_nat" by (simp add: fun_eq_iff) lemma mult_single: "single k a * single l b = single (k + l) (a * b)" proof transfer fix k l :: 'a and a b :: 'b show "prod_fun (\k'. a when k = k') (\k'. b when l = k') = (\k'. a * b when k + l = k')" proof fix k' have "prod_fun (\k'. a when k = k') (\k'. b when l = k') k' = (\n. a * b when l = n when k' = k + n)" by (simp add: prod_fun_def Sum_any_right_distrib mult_when when_mult) also have "\ = (\n. a * b when k' = k + n when l = n)" by (simp add: when_when conj_commute) also have "\ = (a * b when k' = k + l)" by simp also have "\ = (a * b when k + l = k')" by (simp add: when_def) finally show "prod_fun (\k'. a when k = k') (\k'. b when l = k') k' = (\k'. a * b when k + l = k') k'" . qed qed instance poly_mapping :: (monoid_add, semiring_char_0) semiring_char_0 by intro_classes (auto intro: inj_compose inj_of_nat simp add: of_nat_single) instance poly_mapping :: (monoid_add, ring_char_0) ring_char_0 .. lemma single_of_int [simp]: "single 0 (of_int k) = of_int k" by (cases k) (simp_all add: single_diff single_uminus) lemma lookup_of_int: "lookup (of_int l) k = (of_int l when k = 0)" proof - have "lookup (of_int l) k = lookup (single 0 (of_int l)) k" by simp then show ?thesis unfolding lookup_single by simp qed subsection \Integral domains\ instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", semiring_no_zero_divisors) semiring_no_zero_divisors text \The @{class "linorder"} constraint is a pragmatic device for the proof --- maybe it can be dropped\ proof fix f g :: "'a \\<^sub>0 'b" assume "f \ 0" and "g \ 0" then show "f * g \ 0" proof transfer fix f g :: "'a \ 'b" define F where "F = {a. f a \ 0}" moreover define G where "G = {a. g a \ 0}" ultimately have [simp]: "\a. f a \ 0 \ a \ F" "\b. g b \ 0 \ b \ G" by simp_all assume "finite {a. f a \ 0}" then have [simp]: "finite F" by simp assume "finite {a. g a \ 0}" then have [simp]: "finite G" by simp assume "f \ (\a. 0)" then obtain a where "f a \ 0" by (auto simp add: fun_eq_iff) assume "g \ (\b. 0)" then obtain b where "g b \ 0" by (auto simp add: fun_eq_iff) from \f a \ 0\ and \g b \ 0\ have "F \ {}" and "G \ {}" by auto note Max_F = \finite F\ \F \ {}\ note Max_G = \finite G\ \G \ {}\ from Max_F and Max_G have [simp]: "Max F \ F" "Max G \ G" by auto from Max_F Max_G have [dest!]: "\a. a \ F \ a \ Max F" "\b. b \ G \ b \ Max G" by auto define q where "q = Max F + Max G" have "(\(a, b). f a * g b when q = a + b) = (\(a, b). f a * g b when q = a + b when a \ F \ b \ G)" by (rule Sum_any.cong) (auto simp add: split_def when_def q_def intro: ccontr) also have "\ = (\(a, b). f a * g b when (Max F, Max G) = (a, b))" proof (rule Sum_any.cong) fix ab :: "'a \ 'a" obtain a b where [simp]: "ab = (a, b)" by (cases ab) simp_all have [dest!]: "a \ Max F \ Max F \ a \ a < Max F" "b \ Max G \ Max G \ b \ b < Max G" by auto show "(case ab of (a, b) \ f a * g b when q = a + b when a \ F \ b \ G) = (case ab of (a, b) \ f a * g b when (Max F, Max G) = (a, b))" by (auto simp add: split_def when_def q_def dest: add_strict_mono [of a "Max F" b "Max G"]) qed also have "\ = (\ab. (case ab of (a, b) \ f a * g b) when (Max F, Max G) = ab)" unfolding split_def when_def by auto also have "\ \ 0" by simp finally have "prod_fun f g q \ 0" by (simp add: prod_fun_unfold_prod) then show "prod_fun f g \ (\k. 0)" by (auto simp add: fun_eq_iff) qed qed instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_no_zero_divisors) ring_no_zero_divisors .. instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_1_no_zero_divisors) ring_1_no_zero_divisors .. instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", idom) idom .. subsection \Mapping order\ instantiation poly_mapping :: (linorder, "{zero, linorder}") linorder begin lift_definition less_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is less_fun . lift_definition less_eq_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is "\f g. less_fun f g \ f = g" . instance proof (rule class.Orderings.linorder.of_class.intro) show "class.linorder (less_eq :: (_ \\<^sub>0 _) \ _) less" proof (rule linorder_strictI, rule order_strictI) fix f g h :: "'a \\<^sub>0 'b" show "f \ g \ f < g \ f = g" by transfer (rule refl) show "\ f < f" by transfer (rule less_fun_irrefl) show "f < g \ f = g \ g < f" proof transfer fix f g :: "'a \ 'b" assume "finite {k. f k \ 0}" and "finite {k. g k \ 0}" then have "finite ({k. f k \ 0} \ {k. g k \ 0})" by simp moreover have "{k. f k \ g k} \ {k. f k \ 0} \ {k. g k \ 0}" by auto ultimately have "finite {k. f k \ g k}" by (rule rev_finite_subset) then show "less_fun f g \ f = g \ less_fun g f" by (rule less_fun_trichotomy) qed assume "f < g" then show "\ g < f" by transfer (rule less_fun_asym) note \f < g\ moreover assume "g < h" ultimately show "f < h" by transfer (rule less_fun_trans) qed qed end instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, linorder}") ordered_ab_semigroup_add proof (intro_classes, transfer) fix f g h :: "'a \ 'b" assume *: "less_fun f g \ f = g" { assume "less_fun f g" then obtain k where "f k < g k" "(\k'. k' < k \ f k' = g k')" by (blast elim!: less_funE) then have "h k + f k < h k + g k" "(\k'. k' < k \ h k' + f k' = h k' + g k')" by simp_all then have "less_fun (\k. h k + f k) (\k. h k + g k)" by (blast intro: less_funI) } with * show "less_fun (\k. h k + f k) (\k. h k + g k) \ (\k. h k + f k) = (\k. h k + g k)" by (auto simp add: fun_eq_iff) qed instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") linordered_cancel_ab_semigroup_add .. instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_comm_monoid_add .. instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_cancel_comm_monoid_add .. instance poly_mapping :: (linorder, linordered_ab_group_add) linordered_ab_group_add .. text \ For pragmatism we leave out the final elements in the hierarchy: @{class linordered_ring}, @{class linordered_ring_strict}, @{class linordered_idom}; remember that the order instance is a mere technical device, not a deeper algebraic property. \ subsection \Fundamental mapping notions\ lift_definition keys :: "('a \\<^sub>0 'b::zero) \ 'a set" is "\f. {k. f k \ 0}" . lift_definition range :: "('a \\<^sub>0 'b::zero) \ 'b set" is "\f :: 'a \ 'b. Set.range f - {0}" . lemma finite_keys [simp]: "finite (keys f)" by transfer lemma not_in_keys_iff_lookup_eq_zero: "k \ keys f \ lookup f k = 0" by transfer simp lemma lookup_not_eq_zero_eq_in_keys: "lookup f k \ 0 \ k \ keys f" by transfer simp lemma lookup_eq_zero_in_keys_contradict [dest]: "lookup f k = 0 \ \ k \ keys f" by (simp add: not_in_keys_iff_lookup_eq_zero) lemma finite_range [simp]: "finite (Poly_Mapping.range p)" proof transfer fix f :: "'b \ 'a" assume *: "finite {x. f x \ 0}" have "Set.range f - {0} \ f ` {x. f x \ 0}" by auto thus "finite (Set.range f - {0})" by(rule finite_subset)(rule finite_imageI[OF *]) qed lemma in_keys_lookup_in_range [simp]: "k \ keys f \ lookup f k \ range f" by transfer simp lemma in_keys_iff: "x \ (keys s) = (lookup s x \ 0)" by (transfer, simp) lemma keys_zero [simp]: "keys 0 = {}" by transfer simp lemma range_zero [simp]: "range 0 = {}" by transfer auto lemma keys_add: "keys (f + g) \ keys f \ keys g" by transfer auto lemma keys_one [simp]: "keys 1 = {0}" by transfer simp lemma range_one [simp]: "range 1 = {1}" by transfer (auto simp add: when_def) lemma keys_single [simp]: "keys (single k v) = (if v = 0 then {} else {k})" by transfer simp lemma range_single [simp]: "range (single k v) = (if v = 0 then {} else {v})" by transfer (auto simp add: when_def) lemma keys_mult: "keys (f * g) \ {a + b | a b. a \ keys f \ b \ keys g}" apply transfer apply (auto simp add: prod_fun_def dest!: mult_not_zero elim!: Sum_any.not_neutral_obtains_not_neutral) apply blast done lemma setsum_keys_plus_distrib: assumes hom_0: "\k. f k 0 = 0" and hom_plus: "\k. k \ Poly_Mapping.keys p \ Poly_Mapping.keys q \ f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k) = f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k)" shows "(\k\Poly_Mapping.keys (p + q). f k (Poly_Mapping.lookup (p + q) k)) = (\k\Poly_Mapping.keys p. f k (Poly_Mapping.lookup p k)) + (\k\Poly_Mapping.keys q. f k (Poly_Mapping.lookup q k))" (is "?lhs = ?p + ?q") proof - let ?A = "Poly_Mapping.keys p \ Poly_Mapping.keys q" have "?lhs = (\k\?A. f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k))" apply(rule sum.mono_neutral_cong_left) apply(simp_all add: Poly_Mapping.keys_add) apply(transfer fixing: f) apply(auto simp add: hom_0)[1] apply(transfer fixing: f) apply(auto simp add: hom_0)[1] done also have "\ = (\k\?A. f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k))" by(rule sum.cong)(simp_all add: hom_plus) also have "\ = (\k\?A. f k (Poly_Mapping.lookup p k)) + (\k\?A. f k (Poly_Mapping.lookup q k))" (is "_ = ?p' + ?q'") by(simp add: sum.distrib) also have "?p' = ?p" by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right) also have "?q' = ?q" by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right) finally show ?thesis . qed subsection \Degree\ definition degree :: "(nat \\<^sub>0 'a::zero) \ nat" where "degree f = Max (insert 0 (Suc ` keys f))" lemma degree_zero [simp]: "degree 0 = 0" unfolding degree_def by transfer simp lemma degree_one [simp]: "degree 1 = 1" unfolding degree_def by transfer simp lemma degree_single_zero [simp]: "degree (single k 0) = 0" unfolding degree_def by transfer simp lemma degree_single_not_zero [simp]: "v \ 0 \ degree (single k v) = Suc k" unfolding degree_def by transfer simp lemma degree_zero_iff [simp]: "degree f = 0 \ f = 0" unfolding degree_def proof transfer fix f :: "nat \ 'a" assume "finite {n. f n \ 0}" then have fin: "finite (insert 0 (Suc ` {n. f n \ 0}))" by auto show "Max (insert 0 (Suc ` {n. f n \ 0})) = 0 \ f = (\n. 0)" (is "?P \ ?Q") proof assume ?P have "{n. f n \ 0} = {}" proof (rule ccontr) assume "{n. f n \ 0} \ {}" then obtain n where "n \ {n. f n \ 0}" by blast then have "{n. f n \ 0} = insert n {n. f n \ 0}" by auto then have "Suc ` {n. f n \ 0} = insert (Suc n) (Suc ` {n. f n \ 0})" by auto with \?P\ have "Max (insert 0 (insert (Suc n) (Suc ` {n. f n \ 0}))) = 0" by simp then have "Max (insert (Suc n) (insert 0 (Suc ` {n. f n \ 0}))) = 0" by (simp add: insert_commute) with fin have "max (Suc n) (Max (insert 0 (Suc ` {n. f n \ 0}))) = 0" by simp then show False by simp qed then show ?Q by (simp add: fun_eq_iff) next assume ?Q then show ?P by simp qed qed lemma degree_greater_zero_in_keys: assumes "0 < degree f" shows "degree f - 1 \ keys f" proof - from assms have "keys f \ {}" by (auto simp add: degree_def) then show ?thesis unfolding degree_def by (simp add: mono_Max_commute [symmetric] mono_Suc) qed lemma in_keys_less_degree: "n \ keys f \ n < degree f" unfolding degree_def by transfer (auto simp add: Max_gr_iff) lemma beyond_degree_lookup_zero: "degree f \ n \ lookup f n = 0" unfolding degree_def by transfer auto lemma degree_add: "degree (f + g) \ max (degree f) (Poly_Mapping.degree g)" unfolding degree_def proof transfer fix f g :: "nat \ 'a" assume f: "finite {x. f x \ 0}" assume g: "finite {x. g x \ 0}" let ?f = "Max (insert 0 (Suc ` {k. f k \ 0}))" let ?g = "Max (insert 0 (Suc ` {k. g k \ 0}))" have "Max (insert 0 (Suc ` {k. f k + g k \ 0})) \ Max (insert 0 (Suc ` ({k. f k \ 0} \ {k. g k \ 0})))" by (rule Max.subset_imp) (insert f g, auto) also have "\ = max ?f ?g" using f g by (simp_all add: image_Un Max_Un [symmetric]) finally show "Max (insert 0 (Suc ` {k. f k + g k \ 0})) \ max (Max (insert 0 (Suc ` {k. f k \ 0}))) (Max (insert 0 (Suc ` {k. g k \ 0})))" . qed lemma sorted_list_of_set_keys: "sorted_list_of_set (keys f) = filter (\k. k \ keys f) [0..Inductive structure\ lift_definition update :: "'a \ 'b \ ('a \\<^sub>0 'b::zero) \ 'a \\<^sub>0 'b" is "\k v f. f(k := v)" proof - fix f :: "'a \ 'b" and k' v assume "finite {k. f k \ 0}" then have "finite (insert k' {k. f k \ 0})" by simp then show "finite {k. (f(k' := v)) k \ 0}" by (rule rev_finite_subset) auto qed lemma update_induct [case_names const update]: assumes const': "P 0" assumes update': "\f a b. a \ keys f \ b \ 0 \ P f \ P (update a b f)" shows "P f" proof - obtain g where "f = Abs_poly_mapping g" and "finite {a. g a \ 0}" by (cases f) simp_all define Q where "Q g = P (Abs_poly_mapping g)" for g from \finite {a. g a \ 0}\ have "Q g" proof (induct g rule: finite_update_induct) case const with const' Q_def show ?case by simp next case (update a b g) from \finite {a. g a \ 0}\ \g a = 0\ have "a \ keys (Abs_poly_mapping g)" by (simp add: Abs_poly_mapping_inverse keys.rep_eq) moreover note \b \ 0\ moreover from \Q g\ have "P (Abs_poly_mapping g)" by (simp add: Q_def) ultimately have "P (update a b (Abs_poly_mapping g))" by (rule update') also from \finite {a. g a \ 0}\ have "update a b (Abs_poly_mapping g) = Abs_poly_mapping (g(a := b))" by (simp add: update.abs_eq eq_onp_same_args) finally show ?case by (simp add: Q_def fun_upd_def) qed then show ?thesis by (simp add: Q_def \f = Abs_poly_mapping g\) qed lemma lookup_update: "lookup (update k v f) k' = (if k = k' then v else lookup f k')" by transfer simp lemma keys_update: "keys (update k v f) = (if v = 0 then keys f - {k} else insert k (keys f))" by transfer auto subsection \Quasi-functorial structure\ lift_definition map :: "('b::zero \ 'c::zero) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'c::zero)" is "\g f k. g (f k) when f k \ 0" by simp context fixes f :: "'b \ 'a" assumes inj_f: "inj f" begin lift_definition map_key :: "('a \\<^sub>0 'c::zero) \ 'b \\<^sub>0 'c" is "\p. p \ f" proof - fix g :: "'c \ 'd" and p :: "'a \ 'c" assume "finite {x. p x \ 0}" hence "finite (f ` {y. p (f y) \ 0})" by(rule finite_subset[rotated]) auto thus "finite {x. (p \ f) x \ 0}" unfolding o_def by(rule finite_imageD)(rule subset_inj_on[OF inj_f], simp) qed end lemma map_key_compose: assumes [transfer_rule]: "inj f" "inj g" shows "map_key f (map_key g p) = map_key (g \ f) p" proof - from assms have [transfer_rule]: "inj (g \ f)" by(simp add: inj_compose) show ?thesis by transfer(simp add: o_assoc) qed lemma map_key_id: "map_key (\x. x) p = p" proof - have [transfer_rule]: "inj (\x. x)" by simp show ?thesis by transfer(simp add: o_def) qed context fixes f :: "'a \ 'b" assumes inj_f [transfer_rule]: "inj f" begin lemma map_key_map: "map_key f (map g p) = map g (map_key f p)" by transfer (simp add: fun_eq_iff) lemma map_key_plus: "map_key f (p + q) = map_key f p + map_key f q" by transfer (simp add: fun_eq_iff) lemma keys_map_key: "keys (map_key f p) = f -` keys p" by transfer auto lemma map_key_zero [simp]: "map_key f 0 = 0" by transfer (simp add: fun_eq_iff) lemma map_key_single [simp]: "map_key f (single (f k) v) = single k v" by transfer (simp add: fun_eq_iff inj_onD [OF inj_f] when_def) end lemma mult_map_scale_conv_mult: "map ((*) s) p = single 0 s * p" proof(transfer fixing: s) fix p :: "'a \ 'b" assume *: "finite {x. p x \ 0}" { fix x have "prod_fun (\k'. s when 0 = k') p x = (\l :: 'a. if l = 0 then s * (\q. p q when x = q) else 0)" by(auto simp add: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta) also have "\ = (\k. s * p k when p k \ 0) x" by(simp add: when_def) also note calculation } then show "(\k. s * p k when p k \ 0) = prod_fun (\k'. s when 0 = k') p" by(simp add: fun_eq_iff) qed lemma map_single [simp]: "(c = 0 \ f 0 = 0) \ map f (single x c) = single x (f c)" by transfer(auto simp add: fun_eq_iff when_def) lemma map_eq_zero_iff: "map f p = 0 \ (\k \ keys p. f (lookup p k) = 0)" by transfer(auto simp add: fun_eq_iff when_def) subsection \Canonical dense representation of @{typ "nat \\<^sub>0 'a"}\ abbreviation no_trailing_zeros :: "'a :: zero list \ bool" where "no_trailing_zeros \ no_trailing ((=) 0)" lift_definition "nth" :: "'a list \ (nat \\<^sub>0 'a::zero)" is "nth_default 0" by (fact finite_nth_default_neq_default) text \ The opposite direction is directly specified on (later) type \nat_mapping\. \ lemma nth_Nil [simp]: "nth [] = 0" by transfer (simp add: fun_eq_iff) lemma nth_singleton [simp]: "nth [v] = single 0 v" proof (transfer, rule ext) fix n :: nat and v :: 'a show "nth_default 0 [v] n = (v when 0 = n)" by (auto simp add: nth_default_def nth_append) qed lemma nth_replicate [simp]: "nth (replicate n 0 @ [v]) = single n v" proof (transfer, rule ext) fix m n :: nat and v :: 'a show "nth_default 0 (replicate n 0 @ [v]) m = (v when n = m)" by (auto simp add: nth_default_def nth_append) qed lemma nth_strip_while [simp]: "nth (strip_while ((=) 0) xs) = nth xs" by transfer (fact nth_default_strip_while_dflt) lemma nth_strip_while' [simp]: "nth (strip_while (\k. k = 0) xs) = nth xs" by (subst eq_commute) (fact nth_strip_while) lemma nth_eq_iff: "nth xs = nth ys \ strip_while (HOL.eq 0) xs = strip_while (HOL.eq 0) ys" by transfer (simp add: nth_default_eq_iff) lemma lookup_nth [simp]: "lookup (nth xs) = nth_default 0 xs" by (fact nth.rep_eq) lemma keys_nth [simp]: "keys (nth xs) = fst ` {(n, v) \ set (enumerate 0 xs). v \ 0}" proof transfer fix xs :: "'a list" { fix n assume "nth_default 0 xs n \ 0" then have "n < length xs" and "xs ! n \ 0" by (auto simp add: nth_default_def split: if_splits) then have "(n, xs ! n) \ {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" (is "?x \ ?A") by (auto simp add: in_set_conv_nth enumerate_eq_zip) then have "fst ?x \ fst ` ?A" by blast then have "n \ fst ` {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" by simp } then show "{k. nth_default 0 xs k \ 0} = fst ` {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" by (auto simp add: in_enumerate_iff_nth_default_eq) qed lemma range_nth [simp]: "range (nth xs) = set xs - {0}" by transfer simp lemma degree_nth: "no_trailing_zeros xs \ degree (nth xs) = length xs" unfolding degree_def proof transfer fix xs :: "'a list" assume *: "no_trailing_zeros xs" let ?A = "{n. nth_default 0 xs n \ 0}" let ?f = "nth_default 0 xs" let ?bound = "Max (insert 0 (Suc ` {n. ?f n \ 0}))" show "?bound = length xs" proof (cases "xs = []") case False with * obtain n where n: "n < length xs" "xs ! n \ 0" by (fastforce simp add: no_trailing_unfold last_conv_nth neq_Nil_conv) then have "?bound = Max (Suc ` {k. (k < length xs \ xs ! k \ (0::'a)) \ k < length xs})" by (subst Max_insert) (auto simp add: nth_default_def) also let ?A = "{k. k < length xs \ xs ! k \ 0}" have "{k. (k < length xs \ xs ! k \ (0::'a)) \ k < length xs} = ?A" by auto also have "Max (Suc ` ?A) = Suc (Max ?A)" using n by (subst mono_Max_commute [where f = Suc, symmetric]) (auto simp add: mono_Suc) also { have "Max ?A \ ?A" using n Max_in [of ?A] by fastforce hence "Suc (Max ?A) \ length xs" by simp moreover from * False have "length xs - 1 \ ?A" by(auto simp add: no_trailing_unfold last_conv_nth) hence "length xs - 1 \ Max ?A" using Max_ge[of ?A "length xs - 1"] by auto hence "length xs \ Suc (Max ?A)" by simp ultimately have "Suc (Max ?A) = length xs" by simp } finally show ?thesis . qed simp qed lemma nth_trailing_zeros [simp]: "nth (xs @ replicate n 0) = nth xs" by transfer simp lemma nth_idem: "nth (List.map (lookup f) [0.. n" shows "nth (List.map (lookup f) [0..Canonical sparse representation of @{typ "'a \\<^sub>0 'b"}\ lift_definition the_value :: "('a \ 'b) list \ 'a \\<^sub>0 'b::zero" is "\xs k. case map_of xs k of None \ 0 | Some v \ v" proof - fix xs :: "('a \ 'b) list" have fin: "finite {k. \v. map_of xs k = Some v}" using finite_dom_map_of [of xs] unfolding dom_def by auto then show "finite {k. (case map_of xs k of None \ 0 | Some v \ v) \ 0}" using fin by (simp split: option.split) qed definition items :: "('a::linorder \\<^sub>0 'b::zero) \ ('a \ 'b) list" where "items f = List.map (\k. (k, lookup f k)) (sorted_list_of_set (keys f))" text \ For the canonical sparse representation we provide both directions of morphisms since the specification of ordered association lists in theory \OAList\ will support arbitrary linear orders @{class linorder} as keys, not just natural numbers @{typ nat}. \ lemma the_value_items [simp]: "the_value (items f) = f" unfolding items_def by transfer (simp add: fun_eq_iff map_of_map_restrict restrict_map_def) lemma lookup_the_value: "lookup (the_value xs) k = (case map_of xs k of None \ 0 | Some v \ v)" by transfer rule lemma items_the_value: assumes "sorted (List.map fst xs)" and "distinct (List.map fst xs)" and "0 \ snd ` set xs" shows "items (the_value xs) = xs" proof - from assms have "sorted_list_of_set (set (List.map fst xs)) = List.map fst xs" unfolding sorted_list_of_set_sort_remdups by (simp add: distinct_remdups_id sort_key_id_if_sorted) moreover from assms have "keys (the_value xs) = fst ` set xs" by transfer (auto simp add: image_def split: option.split dest: set_map_of_compr) ultimately show ?thesis unfolding items_def using assms by (auto simp add: lookup_the_value intro: map_idI) qed lemma the_value_Nil [simp]: "the_value [] = 0" by transfer (simp add: fun_eq_iff) lemma the_value_Cons [simp]: "the_value (x # xs) = update (fst x) (snd x) (the_value xs)" by transfer (simp add: fun_eq_iff) lemma items_zero [simp]: "items 0 = []" unfolding items_def by simp lemma items_one [simp]: "items 1 = [(0, 1)]" unfolding items_def by transfer simp lemma items_single [simp]: "items (single k v) = (if v = 0 then [] else [(k, v)])" unfolding items_def by simp lemma in_set_items_iff [simp]: "(k, v) \ set (items f) \ k \ keys f \ lookup f k = v" unfolding items_def by transfer auto subsection \Size estimation\ context fixes f :: "'a \ nat" and g :: "'b :: zero \ nat" begin definition poly_mapping_size :: "('a \\<^sub>0 'b) \ nat" where "poly_mapping_size m = g 0 + (\k \ keys m. Suc (f k + g (lookup m k)))" lemma poly_mapping_size_0 [simp]: "poly_mapping_size 0 = g 0" by (simp add: poly_mapping_size_def) lemma poly_mapping_size_single [simp]: "poly_mapping_size (single k v) = (if v = 0 then g 0 else g 0 + f k + g v + 1)" unfolding poly_mapping_size_def by transfer simp lemma keys_less_poly_mapping_size: "k \ keys m \ f k + g (lookup m k) < poly_mapping_size m" unfolding poly_mapping_size_def proof transfer fix k :: 'a and m :: "'a \ 'b" and f :: "'a \ nat" and g let ?keys = "{k. m k \ 0}" assume *: "finite ?keys" "k \ ?keys" then have "f k + g (m k) = (\k' \ ?keys. f k' + g (m k') when k' = k)" by (simp add: sum.delta when_def) also have "\ < (\k' \ ?keys. Suc (f k' + g (m k')))" using * by (intro sum_strict_mono) (auto simp add: when_def) also have "\ \ g 0 + \" by simp finally have "f k + g (m k) < \" . then show "f k + g (m k) < g 0 + (\k | m k \ 0. Suc (f k + g (m k)))" by simp qed lemma lookup_le_poly_mapping_size: "g (lookup m k) \ poly_mapping_size m" proof (cases "k \ keys m") case True with keys_less_poly_mapping_size [of k m] show ?thesis by simp next case False then show ?thesis by (simp add: Poly_Mapping.poly_mapping_size_def in_keys_iff) qed lemma poly_mapping_size_estimation: "k \ keys m \ y \ f k + g (lookup m k) \ y < poly_mapping_size m" using keys_less_poly_mapping_size by (auto intro: le_less_trans) lemma poly_mapping_size_estimation2: assumes "v \ range m" and "y \ g v" shows "y < poly_mapping_size m" proof - from assms obtain k where *: "lookup m k = v" "v \ 0" by transfer blast from * have "k \ keys m" by (simp add: in_keys_iff) then show ?thesis proof (rule poly_mapping_size_estimation) from assms * have "y \ g (lookup m k)" by simp then show "y \ f k + g (lookup m k)" by simp qed qed end lemma poly_mapping_size_one [simp]: "poly_mapping_size f g 1 = g 0 + f 0 + g 1 + 1" unfolding poly_mapping_size_def by transfer simp lemma poly_mapping_size_cong [fundef_cong]: "m = m' \ g 0 = g' 0 \ (\k. k \ keys m' \ f k = f' k) \ (\v. v \ range m' \ g v = g' v) \ poly_mapping_size f g m = poly_mapping_size f' g' m'" by (auto simp add: poly_mapping_size_def intro!: sum.cong) instantiation poly_mapping :: (type, zero) size begin definition "size = poly_mapping_size (\_. 0) (\_. 0)" instance .. end subsection \Further mapping operations and properties\ text \It is like in algebra: there are many definitions, some are also used\ lift_definition mapp :: "('a \ 'b :: zero \ 'c :: zero) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'c)" is "\f p k. (if k \ keys p then f k (lookup p k) else 0)" by simp lemma mapp_cong [fundef_cong]: "\ m = m'; \k. k \ keys m' \ f k (lookup m' k) = f' k (lookup m' k) \ \ mapp f m = mapp f' m'" by transfer (auto simp add: fun_eq_iff) lemma lookup_mapp: "lookup (mapp f p) k = (f k (lookup p k) when k \ keys p)" by (simp add: mapp.rep_eq) lemma keys_mapp_subset: "keys (mapp f p) \ keys p" by (meson in_keys_iff mapp.rep_eq subsetI) subsection\Free Abelian Groups Over a Type\ abbreviation frag_of :: "'a \ 'a \\<^sub>0 int" where "frag_of c \ Poly_Mapping.single c (1::int)" lemma lookup_frag_of [simp]: "Poly_Mapping.lookup(frag_of c) = (\x. if x = c then 1 else 0)" by (force simp add: lookup_single_not_eq) lemma frag_of_nonzero [simp]: "frag_of a \ 0" proof - let ?f = "\x. if x = a then 1 else (0::int)" have "?f \ (\x. 0::int)" by (auto simp: fun_eq_iff) then have "Poly_Mapping.lookup (Abs_poly_mapping ?f) \ Poly_Mapping.lookup (Abs_poly_mapping (\x. 0))" by fastforce then show ?thesis by (metis lookup_single_eq lookup_zero) qed definition frag_cmul :: "int \ ('a \\<^sub>0 int) \ ('a \\<^sub>0 int)" where "frag_cmul c a = Abs_poly_mapping (\x. c * Poly_Mapping.lookup a x)" lemma frag_cmul_zero [simp]: "frag_cmul 0 x = 0" by (simp add: frag_cmul_def) lemma frag_cmul_zero2 [simp]: "frag_cmul c 0 = 0" by (simp add: frag_cmul_def) lemma frag_cmul_one [simp]: "frag_cmul 1 x = x" by (auto simp: frag_cmul_def Poly_Mapping.poly_mapping.lookup_inverse) lemma frag_cmul_minus_one [simp]: "frag_cmul (-1) x = -x" by (simp add: frag_cmul_def uminus_poly_mapping_def poly_mapping_eqI) lemma frag_cmul_cmul [simp]: "frag_cmul c (frag_cmul d x) = frag_cmul (c*d) x" by (simp add: frag_cmul_def mult_ac) lemma lookup_frag_cmul [simp]: "poly_mapping.lookup (frag_cmul c x) i = c * poly_mapping.lookup x i" by (simp add: frag_cmul_def) lemma minus_frag_cmul [simp]: "- frag_cmul k x = frag_cmul (-k) x" by (simp add: poly_mapping_eqI) lemma keys_frag_of: "Poly_Mapping.keys(frag_of a) = {a}" by simp lemma finite_cmul_nonzero: "finite {x. c * Poly_Mapping.lookup a x \ (0::int)}" by simp lemma keys_cmul: "Poly_Mapping.keys(frag_cmul c a) \ Poly_Mapping.keys a" using finite_cmul_nonzero [of c a] by (metis lookup_frag_cmul mult_zero_right not_in_keys_iff_lookup_eq_zero subsetI) lemma keys_cmul_iff [iff]: "i \ Poly_Mapping.keys (frag_cmul c x) \ i \ Poly_Mapping.keys x \ c \ 0" apply auto apply (meson subsetD keys_cmul) by (metis in_keys_iff lookup_frag_cmul mult_eq_0_iff) lemma keys_minus [simp]: "Poly_Mapping.keys(-a) = Poly_Mapping.keys a" by (metis (no_types, opaque_lifting) in_keys_iff lookup_uminus neg_equal_0_iff_equal subsetI subset_antisym) lemma keys_diff: "Poly_Mapping.keys(a - b) \ Poly_Mapping.keys a \ Poly_Mapping.keys b" by (auto simp add: in_keys_iff lookup_minus) lemma keys_eq_empty [simp]: "Poly_Mapping.keys c = {} \ c = 0" by (metis in_keys_iff keys_zero lookup_zero poly_mapping_eqI) lemma frag_cmul_eq_0_iff [simp]: "frag_cmul k c = 0 \ k=0 \ c=0" by auto (metis subsetI subset_antisym keys_cmul_iff keys_eq_empty) lemma frag_of_eq: "frag_of x = frag_of y \ x = y" by (metis lookup_single_eq lookup_single_not_eq zero_neq_one) lemma frag_cmul_distrib: "frag_cmul (c+d) a = frag_cmul c a + frag_cmul d a" by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib) lemma frag_cmul_distrib2: "frag_cmul c (a+b) = frag_cmul c a + frag_cmul c b" proof - have "finite {x. poly_mapping.lookup a x + poly_mapping.lookup b x \ 0}" using keys_add [of a b] by (metis (no_types, lifting) finite_keys finite_subset keys.rep_eq lookup_add mem_Collect_eq subsetI) then show ?thesis by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib) qed lemma frag_cmul_diff_distrib: "frag_cmul (a - b) c = frag_cmul a c - frag_cmul b c" by (auto simp: left_diff_distrib lookup_minus poly_mapping_eqI) lemma frag_cmul_sum: "frag_cmul a (sum b I) = (\i\I. frag_cmul a (b i))" proof (induction rule: infinite_finite_induct) case (insert i I) then show ?case by (auto simp: algebra_simps frag_cmul_distrib2) qed auto lemma keys_sum: "Poly_Mapping.keys(sum b I) \ (\i \I. Poly_Mapping.keys(b i))" proof (induction I rule: infinite_finite_induct) case (insert i I) then show ?case using keys_add [of "b i" "sum b I"] by auto qed auto definition frag_extend :: "('b \ 'a \\<^sub>0 int) \ ('b \\<^sub>0 int) \ 'a \\<^sub>0 int" where "frag_extend b x \ (\i \ Poly_Mapping.keys x. frag_cmul (Poly_Mapping.lookup x i) (b i))" lemma frag_extend_0 [simp]: "frag_extend b 0 = 0" by (simp add: frag_extend_def) lemma frag_extend_of [simp]: "frag_extend f (frag_of a) = f a" by (simp add: frag_extend_def) lemma frag_extend_cmul: "frag_extend f (frag_cmul c x) = frag_cmul c (frag_extend f x)" by (auto simp: frag_extend_def frag_cmul_sum intro: sum.mono_neutral_cong_left) lemma frag_extend_minus: "frag_extend f (- x) = - (frag_extend f x)" using frag_extend_cmul [of f "-1"] by simp lemma frag_extend_add: "frag_extend f (a+b) = (frag_extend f a) + (frag_extend f b)" proof - have *: "(\i\Poly_Mapping.keys a. frag_cmul (poly_mapping.lookup a i) (f i)) = (\i\Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i))" "(\i\Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i)) = (\i\Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))" by (auto simp: in_keys_iff intro: sum.mono_neutral_cong_left) have "frag_extend f (a+b) = (\i\Poly_Mapping.keys (a + b). frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i)) " by (auto simp: frag_extend_def Poly_Mapping.lookup_add frag_cmul_distrib) also have "... = (\i \ Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i))" apply (rule sum.mono_neutral_cong_left) using keys_add [of a b] apply (auto simp add: in_keys_iff plus_poly_mapping.rep_eq frag_cmul_distrib [symmetric]) done also have "... = (frag_extend f a) + (frag_extend f b)" by (auto simp: * sum.distrib frag_extend_def) finally show ?thesis . qed lemma frag_extend_diff: "frag_extend f (a-b) = (frag_extend f a) - (frag_extend f b)" by (metis (no_types, opaque_lifting) add_uminus_conv_diff frag_extend_add frag_extend_minus) lemma frag_extend_sum: "finite I \ frag_extend f (\i\I. g i) = sum (frag_extend f o g) I" by (induction I rule: finite_induct) (simp_all add: frag_extend_add) lemma frag_extend_eq: "(\f. f \ Poly_Mapping.keys c \ g f = h f) \ frag_extend g c = frag_extend h c" by (simp add: frag_extend_def) lemma frag_extend_eq_0: "(\x. x \ Poly_Mapping.keys c \ f x = 0) \ frag_extend f c = 0" by (simp add: frag_extend_def) lemma keys_frag_extend: "Poly_Mapping.keys(frag_extend f c) \ (\x \ Poly_Mapping.keys c. Poly_Mapping.keys(f x))" unfolding frag_extend_def using keys_sum by fastforce lemma frag_expansion: "a = frag_extend frag_of a" proof - have *: "finite I \ Poly_Mapping.lookup (\i\I. frag_cmul (Poly_Mapping.lookup a i) (frag_of i)) j = (if j \ I then Poly_Mapping.lookup a j else 0)" for I j by (induction I rule: finite_induct) (auto simp: lookup_single lookup_add) show ?thesis unfolding frag_extend_def by (rule poly_mapping_eqI) (fastforce simp add: in_keys_iff *) qed lemma frag_closure_minus_cmul: assumes "P 0" and P: "\x y. \P x; P y\ \ P(x - y)" "P c" shows "P(frag_cmul k c)" proof - have "P (frag_cmul (int n) c)" for n apply (induction n) apply (simp_all add: assms frag_cmul_distrib) by (metis add.left_neutral add_diff_cancel_right' add_uminus_conv_diff P) then show ?thesis by (metis (no_types, opaque_lifting) add_diff_eq assms(2) diff_add_cancel frag_cmul_distrib int_diff_cases) qed lemma frag_induction [consumes 1, case_names zero one diff]: assumes supp: "Poly_Mapping.keys c \ S" and 0: "P 0" and sing: "\x. x \ S \ P(frag_of x)" and diff: "\a b. \P a; P b\ \ P(a - b)" shows "P c" proof - have "P (\i\I. frag_cmul (poly_mapping.lookup c i) (frag_of i))" if "I \ Poly_Mapping.keys c" for I using finite_subset [OF that finite_keys [of c]] that supp proof (induction I arbitrary: c rule: finite_induct) case empty then show ?case by (auto simp: 0) next case (insert i I c) have ab: "a+b = a - (0 - b)" for a b :: "'a \\<^sub>0 int" by simp have Pfrag: "P (frag_cmul (poly_mapping.lookup c i) (frag_of i))" by (metis "0" diff frag_closure_minus_cmul insert.prems insert_subset sing subset_iff) show ?case apply (simp add: insert.hyps) apply (subst ab) using insert apply (blast intro: assms Pfrag) done qed then show ?thesis by (subst frag_expansion) (auto simp add: frag_extend_def) qed lemma frag_extend_compose: "frag_extend f (frag_extend (frag_of o g) c) = frag_extend (f o g) c" using subset_UNIV by (induction c rule: frag_induction) (auto simp: frag_extend_diff) lemma frag_split: fixes c :: "'a \\<^sub>0 int" assumes "Poly_Mapping.keys c \ S \ T" obtains d e where "Poly_Mapping.keys d \ S" "Poly_Mapping.keys e \ T" "d + e = c" proof let ?d = "frag_extend (\f. if f \ S then frag_of f else 0) c" let ?e = "frag_extend (\f. if f \ S then 0 else frag_of f) c" show "Poly_Mapping.keys ?d \ S" "Poly_Mapping.keys ?e \ T" using assms by (auto intro!: order_trans [OF keys_frag_extend] split: if_split_asm) show "?d + ?e = c" using assms proof (induction c rule: frag_induction) case (diff a b) then show ?case by (metis (no_types, lifting) frag_extend_diff add_diff_eq diff_add_eq diff_add_eq_diff_diff_swap) qed auto qed hide_const (open) lookup single update keys range map map_key degree nth the_value items foldr mapp end \ No newline at end of file diff --git a/src/HOL/Library/Ramsey.thy b/src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy +++ b/src/HOL/Library/Ramsey.thy @@ -1,1002 +1,1002 @@ (* Title: HOL/Library/Ramsey.thy Author: Tom Ridge. Full finite version by L C Paulson. *) section \Ramsey's Theorem\ theory Ramsey imports Infinite_Set FuncSet begin subsection \Preliminary definitions\ abbreviation strict_sorted :: "'a::linorder list \ bool" where "strict_sorted \ sorted_wrt (<)" subsubsection \The $n$-element subsets of a set $A$\ definition nsets :: "['a set, nat] \ 'a set set" ("([_]\<^bsup>_\<^esup>)" [0,999] 999) where "nsets A n \ {N. N \ A \ finite N \ card N = n}" lemma nsets_mono: "A \ B \ nsets A n \ nsets B n" by (auto simp: nsets_def) lemma nsets_Pi_contra: "A' \ A \ Pi ([A]\<^bsup>n\<^esup>) B \ Pi ([A']\<^bsup>n\<^esup>) B" by (auto simp: nsets_def) lemma nsets_2_eq: "nsets A 2 = (\x\A. \y\A - {x}. {{x, y}})" by (auto simp: nsets_def card_2_iff) lemma nsets_doubleton_2_eq [simp]: "[{x, y}]\<^bsup>2\<^esup> = (if x=y then {} else {{x, y}})" by (auto simp: nsets_2_eq) lemma doubleton_in_nsets_2 [simp]: "{x,y} \ [A]\<^bsup>2\<^esup> \ x \ A \ y \ A \ x \ y" by (auto simp: nsets_2_eq Set.doubleton_eq_iff) lemma nsets_3_eq: "nsets A 3 = (\x\A. \y\A - {x}. \z\A - {x,y}. {{x,y,z}})" by (simp add: eval_nat_numeral nsets_def card_Suc_eq) blast lemma nsets_4_eq: "[A]\<^bsup>4\<^esup> = (\u\A. \x\A - {u}. \y\A - {u,x}. \z\A - {u,x,y}. {{u,x,y,z}})" (is "_ = ?rhs") proof show "[A]\<^bsup>4\<^esup> \ ?rhs" by (clarsimp simp add: nsets_def eval_nat_numeral card_Suc_eq) blast show "?rhs \ [A]\<^bsup>4\<^esup>" apply (clarsimp simp add: nsets_def eval_nat_numeral card_Suc_eq) by (metis insert_iff singletonD) qed lemma nsets_disjoint_2: "X \ Y = {} \ [X \ Y]\<^bsup>2\<^esup> = [X]\<^bsup>2\<^esup> \ [Y]\<^bsup>2\<^esup> \ (\x\X. \y\Y. {{x,y}})" by (fastforce simp: nsets_2_eq Set.doubleton_eq_iff) lemma ordered_nsets_2_eq: fixes A :: "'a::linorder set" shows "nsets A 2 = {{x,y} | x y. x \ A \ y \ A \ x ?rhs" unfolding numeral_nat apply (clarsimp simp add: nsets_def card_Suc_eq Set.doubleton_eq_iff not_less) by (metis antisym) show "?rhs \ nsets A 2" unfolding numeral_nat by (auto simp: nsets_def card_Suc_eq) qed lemma ordered_nsets_3_eq: fixes A :: "'a::linorder set" shows "nsets A 3 = {{x,y,z} | x y z. x \ A \ y \ A \ z \ A \ x y ?rhs" apply (clarsimp simp add: nsets_def card_Suc_eq eval_nat_numeral) by (metis insert_commute linorder_cases) show "?rhs \ nsets A 3" apply (clarsimp simp add: nsets_def card_Suc_eq eval_nat_numeral) by (metis empty_iff insert_iff not_less_iff_gr_or_eq) qed lemma ordered_nsets_4_eq: fixes A :: "'a::linorder set" shows "[A]\<^bsup>4\<^esup> = {U. \u x y z. U = {u,x,y,z} \ u \ A \ x \ A \ y \ A \ z \ A \ u < x \ x < y \ y < z}" (is "_ = Collect ?RHS") proof - { fix U assume "U \ [A]\<^bsup>4\<^esup>" then obtain l where "strict_sorted l" "List.set l = U" "length l = 4" "U \ A" by (simp add: nsets_def) (metis finite_set_strict_sorted) then have "?RHS U" unfolding numeral_nat length_Suc_conv by auto blast } moreover have "Collect ?RHS \ [A]\<^bsup>4\<^esup>" apply (clarsimp simp add: nsets_def eval_nat_numeral) apply (subst card_insert_disjoint, auto)+ done ultimately show ?thesis by auto qed lemma ordered_nsets_5_eq: fixes A :: "'a::linorder set" shows "[A]\<^bsup>5\<^esup> = {U. \u v x y z. U = {u,v,x,y,z} \ u \ A \ v \ A \ x \ A \ y \ A \ z \ A \ u < v \ v < x \ x < y \ y < z}" (is "_ = Collect ?RHS") proof - { fix U assume "U \ [A]\<^bsup>5\<^esup>" then obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \ A" apply (simp add: nsets_def) by (metis finite_set_strict_sorted) then have "?RHS U" unfolding numeral_nat length_Suc_conv by auto blast } moreover have "Collect ?RHS \ [A]\<^bsup>5\<^esup>" apply (clarsimp simp add: nsets_def eval_nat_numeral) apply (subst card_insert_disjoint, auto)+ done ultimately show ?thesis by auto qed lemma binomial_eq_nsets: "n choose k = card (nsets {0.. finite A \ card A < r" unfolding nsets_def proof (intro iffI conjI) assume that: "{N. N \ A \ finite N \ card N = r} = {}" show "finite A" using infinite_arbitrarily_large that by auto then have "\ r \ card A" using that by (simp add: set_eq_iff) (metis obtain_subset_with_card_n) then show "card A < r" using not_less by blast next show "{N. N \ A \ finite N \ card N = r} = {}" if "finite A \ card A < r" using that card_mono leD by auto qed lemma nsets_eq_empty: "\finite A; card A < r\ \ nsets A r = {}" by (simp add: nsets_eq_empty_iff) lemma nsets_empty_iff: "nsets {} r = (if r=0 then {{}} else {})" by (auto simp: nsets_def) lemma nsets_singleton_iff: "nsets {a} r = (if r=0 then {{}} else if r=1 then {{a}} else {})" by (auto simp: nsets_def card_gt_0_iff subset_singleton_iff) lemma nsets_self [simp]: "nsets {..x. {x}) ` A" using card_eq_SucD by (force simp: nsets_def) lemma inj_on_nsets: assumes "inj_on f A" shows "inj_on (\X. f ` X) ([A]\<^bsup>n\<^esup>)" using assms unfolding nsets_def by (metis (no_types, lifting) inj_on_inverseI inv_into_image_cancel mem_Collect_eq) lemma bij_betw_nsets: assumes "bij_betw f A B" shows "bij_betw (\X. f ` X) ([A]\<^bsup>n\<^esup>) ([B]\<^bsup>n\<^esup>)" proof - have "(`) f ` [A]\<^bsup>n\<^esup> = [f ` A]\<^bsup>n\<^esup>" using assms apply (auto simp: nsets_def bij_betw_def image_iff card_image inj_on_subset) by (metis card_image inj_on_finite order_refl subset_image_inj) with assms show ?thesis by (auto simp: bij_betw_def inj_on_nsets) qed lemma nset_image_obtains: assumes "X \ [f`A]\<^bsup>k\<^esup>" "inj_on f A" obtains Y where "Y \ [A]\<^bsup>k\<^esup>" "X = f ` Y" using assms apply (clarsimp simp add: nsets_def subset_image_iff) by (metis card_image finite_imageD inj_on_subset) lemma nsets_image_funcset: assumes "g \ S \ T" and "inj_on g S" shows "(\X. g ` X) \ [S]\<^bsup>k\<^esup> \ [T]\<^bsup>k\<^esup>" using assms by (fastforce simp add: nsets_def card_image inj_on_subset subset_iff simp flip: image_subset_iff_funcset) lemma nsets_compose_image_funcset: assumes f: "f \ [T]\<^bsup>k\<^esup> \ D" and "g \ S \ T" and "inj_on g S" shows "f \ (\X. g ` X) \ [S]\<^bsup>k\<^esup> \ D" proof - have "(\X. g ` X) \ [S]\<^bsup>k\<^esup> \ [T]\<^bsup>k\<^esup>" using assms by (simp add: nsets_image_funcset) then show ?thesis using f by fastforce qed subsubsection \Partition predicates\ definition partn :: "'a set \ nat \ nat \ 'b set \ bool" where "partn \ \ \ \ \ \f \ nsets \ \ \ \. \H \ nsets \ \. \\\\. f ` (nsets H \) \ {\}" definition partn_lst :: "'a set \ nat list \ nat \ bool" where "partn_lst \ \ \ \ \f \ nsets \ \ \ {..}. \i < length \. \H \ nsets \ (\!i). f ` (nsets H \) \ {i}" lemma partn_lst_greater_resource: fixes M::nat assumes M: "partn_lst {.. \" and "M \ N" shows "partn_lst {.. \" proof (clarsimp simp: partn_lst_def) fix f assume "f \ nsets {.. \ {..}" then have "f \ nsets {.. \ {..}" by (meson Pi_anti_mono \M \ N\ lessThan_subset_iff nsets_mono subsetD) then obtain i H where i: "i < length \" and H: "H \ nsets {.. ! i)" and subi: "f ` nsets H \ \ {i}" using M partn_lst_def by blast have "H \ nsets {.. ! i)" using \M \ N\ H by (auto simp: nsets_def subset_iff) then show "\i. \H\nsets {.. ! i). f ` nsets H \ \ {i}" using i subi by blast qed lemma partn_lst_fewer_colours: assumes major: "partn_lst \ (n#\) \" and "n \ \" shows "partn_lst \ \ \" proof (clarsimp simp: partn_lst_def) fix f :: "'a set \ nat" assume f: "f \ [\]\<^bsup>\\<^esup> \ {..}" then obtain i H where i: "i < Suc (length \)" and H: "H \ [\]\<^bsup>((n # \) ! i)\<^esup>" and hom: "\x\[H]\<^bsup>\\<^esup>. Suc (f x) = i" using \n \ \\ major [unfolded partn_lst_def, rule_format, of "Suc o f"] by (fastforce simp: image_subset_iff nsets_eq_empty_iff) show "\i. \H\nsets \ (\ ! i). f ` [H]\<^bsup>\\<^esup> \ {i}" proof (cases i) case 0 then have "[H]\<^bsup>\\<^esup> = {}" using hom by blast then show ?thesis using 0 H \n \ \\ by (simp add: nsets_eq_empty_iff) (simp add: nsets_def) next case (Suc i') then show ?thesis using i H hom by auto qed qed lemma partn_lst_eq_partn: "partn_lst {..Finite versions of Ramsey's theorem\ text \ To distinguish the finite and infinite ones, lower and upper case names are used. \ subsubsection \Trivial cases\ text \Vacuous, since we are dealing with 0-sets!\ lemma ramsey0: "\N::nat. partn_lst {..Just the pigeon hole principle, since we are dealing with 1-sets\ lemma ramsey1: "\N::nat. partn_lst {..iH\nsets {.. {i}" if "f \ nsets {.. {.. \i. {q. q \ q0+q1 \ f {q} = i}" have "A 0 \ A 1 = {..q0 + q1}" using that by (auto simp: A_def PiE_iff nsets_one lessThan_Suc_atMost le_Suc_eq) moreover have "A 0 \ A 1 = {}" by (auto simp: A_def) ultimately have "q0 + q1 \ card (A 0) + card (A 1)" by (metis card_Un_le card_atMost eq_imp_le le_SucI le_trans) then consider "card (A 0) \ q0" | "card (A 1) \ q1" by linarith then obtain i where "i < Suc (Suc 0)" "card (A i) \ [q0, q1] ! i" by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc zero_less_Suc) then obtain B where "B \ A i" "card B = [q0, q1] ! i" "finite B" by (meson obtain_subset_with_card_n) then have "B \ nsets {.. f ` nsets B (Suc 0) \ {i}" by (auto simp: A_def nsets_def card_1_singleton_iff) then show ?thesis using \i < Suc (Suc 0)\ by auto qed then show ?thesis by (clarsimp simp: partn_lst_def) blast qed subsubsection \Ramsey's theorem with two colours and arbitrary exponents (hypergraph version)\ proposition ramsey2_full: "\N::nat. partn_lst {.. 0" by simp show ?thesis using Suc.prems proof (induct k \ "q1 + q2" arbitrary: q1 q2) case 0 show ?case proof show "partn_lst {..<1::nat} [q1, q2] (Suc r)" using nsets_empty_iff subset_insert 0 by (fastforce simp: partn_lst_def funcset_to_empty_iff nsets_eq_empty image_subset_iff) qed next case (Suc k) consider "q1 = 0 \ q2 = 0" | "q1 \ 0" "q2 \ 0" by auto then show ?case proof cases case 1 then have "partn_lst {..< Suc 0} [q1, q2] (Suc r)" unfolding partn_lst_def using \r > 0\ by (fastforce simp add: nsets_empty_iff nsets_singleton_iff lessThan_Suc) then show ?thesis by blast next case 2 with Suc have "k = (q1 - 1) + q2" "k = q1 + (q2 - 1)" by auto then obtain p1 p2::nat where p1: "partn_lst {..iH\nsets {..p} ([q1,q2] ! i). f ` nsets H (Suc r) \ {i}" if f: "f \ nsets {..p} (Suc r) \ {.. \R. f (insert p R)" have "f (insert p i) \ {.. nsets {.. nsets {.. {.. {i}" and U: "U \ nsets {.. {.. nsets {.. nsets {..p} n" if "X \ nsets {.. \R. f (u ` R)" have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. {.. u ` {.. nsets {.. nsets {.. nsets {..p} q1" unfolding nsets_def using \q1 \ 0\ card_insert_if by fastforce have invu_nsets: "inv_into {.. nsets V r" if "X \ nsets (u ` V) r" for X r proof - have "X \ u ` V \ finite X \ card X = r" using nsets_def that by auto then have [simp]: "card (inv_into {.. nsets ?W (Suc r)" for X proof (cases "p \ X") case True then have Xp: "X - {p} \ nsets (u ` V) r" using X by (auto simp: nsets_def) moreover have "u ` V \ U" using Vsub bij_betwE u by blast ultimately have "X - {p} \ nsets U r" by (meson in_mono nsets_mono) then have "g (X - {p}) = i" using gi by blast have "f X = i" using gi True \X - {p} \ nsets U r\ insert_Diff by (fastforce simp add: g_def image_subset_iff) then show ?thesis by (simp add: \f X = i\ \g (X - {p}) = i\) next case False then have Xim: "X \ nsets (u ` V) (Suc r)" using X by (auto simp: nsets_def subset_insert) then have "u ` inv_into {.. nsets {..p} q1" by (simp add: izero inq1) ultimately show ?thesis by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc) next case jone then have "u ` V \ nsets {..p} q2" using V u_nsets by auto moreover have "f ` nsets (u ` V) (Suc r) \ {j}" using hj by (force simp add: h_def image_subset_iff nsets_def subset_image_inj card_image dest: finite_imageD) ultimately show ?thesis using jone not_less_eq by fastforce qed next case ione then have "U \ nsets {.. nsets {..p} n" if "X \ nsets {.. \R. f (u ` R)" have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. {.. u ` {.. nsets {.. nsets {.. nsets {..p} q2" unfolding nsets_def using \q2 \ 0\ card_insert_if by fastforce have invu_nsets: "inv_into {.. nsets V r" if "X \ nsets (u ` V) r" for X r proof - have "X \ u ` V \ finite X \ card X = r" using nsets_def that by auto then have [simp]: "card (inv_into {.. nsets ?W (Suc r)" for X proof (cases "p \ X") case True then have Xp: "X - {p} \ nsets (u ` V) r" using X by (auto simp: nsets_def) moreover have "u ` V \ U" using Vsub bij_betwE u by blast ultimately have "X - {p} \ nsets U r" by (meson in_mono nsets_mono) then have "g (X - {p}) = i" using gi by blast have "f X = i" using gi True \X - {p} \ nsets U r\ insert_Diff by (fastforce simp add: g_def image_subset_iff) then show ?thesis by (simp add: \f X = i\ \g (X - {p}) = i\) next case False then have Xim: "X \ nsets (u ` V) (Suc r)" using X by (auto simp: nsets_def subset_insert) then have "u ` inv_into {.. nsets {..p} q2" by (simp add: ione inq1) ultimately show ?thesis by (metis ione image_subsetI insertI1 lessI nth_Cons_0 nth_Cons_Suc) next case jzero then have "u ` V \ nsets {..p} q1" using V u_nsets by auto moreover have "f ` nsets (u ` V) (Suc r) \ {j}" using hj apply (clarsimp simp add: h_def image_subset_iff nsets_def) by (metis Zero_not_Suc card_eq_0_iff card_image subset_image_inj) ultimately show ?thesis using jzero not_less_eq by fastforce qed qed qed then show "partn_lst {..Full Ramsey's theorem with multiple colours and arbitrary exponents\ theorem ramsey_full: "\N::nat. partn_lst {.. "length qs" arbitrary: qs) case 0 then show ?case by (rule_tac x=" r" in exI) (simp add: partn_lst_def) next case (Suc k) note IH = this show ?case proof (cases k) case 0 with Suc obtain q where "qs = [q]" by (metis length_0_conv length_Suc_conv) then show ?thesis by (rule_tac x=q in exI) (auto simp: partn_lst_def funcset_to_empty_iff) next case (Suc k') then obtain q1 q2 l where qs: "qs = q1#q2#l" by (metis Suc.hyps(2) length_Suc_conv) then obtain q::nat where q: "partn_lst {..qs = q1 # q2 # l\ by fastforce have keq: "Suc (length l) = k" using IH qs by auto show ?thesis proof (intro exI conjI) show "partn_lst {.. nsets {.. {.. \X. if f X < Suc (Suc 0) then 0 else f X - Suc 0" have "g \ nsets {.. {.. {i}" and U: "U \ nsets {..iH\nsets {.. {i}" proof (cases "i = 0") case True then have "U \ nsets {.. {0, Suc 0}" using U gi unfolding g_def by (auto simp: image_subset_iff) then obtain u where u: "bij_betw u {.. {.. nsets {.. nsets {.. \X. f (u ` X)" have "f (u ` X) < Suc (Suc 0)" if "X \ nsets {.. nsets U r" using u u_nsets that by (auto simp: nsets_def bij_betwE subset_eq) then show ?thesis using f01 by auto qed then have "h \ nsets {.. {.. {j}" and V: "V \ nsets {.. (\x. (u ` x)) ` nsets V r" apply (clarsimp simp add: nsets_def image_iff) by (metis card_eq_0_iff card_image image_is_empty subset_image_inj) then have "f ` nsets (u ` V) r \ h ` nsets V r" by (auto simp: h_def) then show "f ` nsets (u ` V) r \ {j}" using hj by auto show "(u ` V) \ nsets {.. {Suc i}" using i gi False apply (auto simp: g_def image_subset_iff) by (metis Suc_lessD Suc_pred g_def gi image_subset_iff not_less_eq singleton_iff) show "U \ nsets {..Simple graph version\ text \This is the most basic version in terms of cliques and independent sets, i.e. the version for graphs and 2 colours. \ definition "clique V E \ (\v\V. \w\V. v \ w \ {v, w} \ E)" definition "indep V E \ (\v\V. \w\V. v \ w \ {v, w} \ E)" lemma ramsey2: "\r\1. \(V::'a set) (E::'a set set). finite V \ card V \ r \ (\R \ V. card R = m \ clique R E \ card R = n \ indep R E)" proof - obtain N where "N \ Suc 0" and N: "partn_lst {..R\V. card R = m \ clique R E \ card R = n \ indep R E" if "finite V" "N \ card V" for V :: "'a set" and E :: "'a set set" proof - from that obtain v where u: "inj_on v {.. V" by (metis card_le_inj card_lessThan finite_lessThan) define f where "f \ \e. if v ` e \ E then 0 else Suc 0" have f: "f \ nsets {.. {.. {i}" and U: "U \ nsets {.. V" using U u by (auto simp: image_subset_iff nsets_def) show "card (v ` U) = m \ clique (v ` U) E \ card (v ` U) = n \ indep (v ` U) E" using i unfolding numeral_2_eq_2 using gi U u apply (simp add: image_subset_iff nsets_2_eq clique_def indep_def less_Suc_eq) apply (auto simp: f_def nsets_def card_image inj_on_subset split: if_split_asm) done qed qed then show ?thesis using \Suc 0 \ N\ by auto qed subsection \Preliminaries\ subsubsection \``Axiom'' of Dependent Choice\ primrec choice :: "('a \ bool) \ ('a \ 'a) set \ nat \ 'a" where \ \An integer-indexed chain of choices\ choice_0: "choice P r 0 = (SOME x. P x)" | choice_Suc: "choice P r (Suc n) = (SOME y. P y \ (choice P r n, y) \ r)" lemma choice_n: assumes P0: "P x0" and Pstep: "\x. P x \ \y. P y \ (x, y) \ r" shows "P (choice P r n)" proof (induct n) case 0 show ?case by (force intro: someI P0) next case Suc then show ?case by (auto intro: someI2_ex [OF Pstep]) qed lemma dependent_choice: assumes trans: "trans r" and P0: "P x0" and Pstep: "\x. P x \ \y. P y \ (x, y) \ r" obtains f :: "nat \ 'a" where "\n. P (f n)" and "\n m. n < m \ (f n, f m) \ r" proof fix n show "P (choice P r n)" by (blast intro: choice_n [OF P0 Pstep]) next fix n m :: nat assume "n < m" from Pstep [OF choice_n [OF P0 Pstep]] have "(choice P r k, choice P r (Suc k)) \ r" for k by (auto intro: someI2_ex) then show "(choice P r n, choice P r m) \ r" by (auto intro: less_Suc_induct [OF \n < m\] transD [OF trans]) qed subsubsection \Partition functions\ definition part_fn :: "nat \ nat \ 'a set \ ('a set \ nat) \ bool" \ \the function \<^term>\f\ partitions the \<^term>\r\-subsets of the typically infinite set \<^term>\Y\ into \<^term>\s\ distinct categories.\ where "part_fn r s Y f \ (f \ nsets Y r \ {..For induction, we decrease the value of \<^term>\r\ in partitions.\ lemma part_fn_Suc_imp_part_fn: "\infinite Y; part_fn (Suc r) s Y f; y \ Y\ \ part_fn r s (Y - {y}) (\u. f (insert y u))" by (simp add: part_fn_def nsets_def Pi_def subset_Diff_insert) lemma part_fn_subset: "part_fn r s YY f \ Y \ YY \ part_fn r s Y f" unfolding part_fn_def nsets_def by blast subsection \Ramsey's Theorem: Infinitary Version\ lemma Ramsey_induction: fixes s r :: nat and YY :: "'a set" and f :: "'a set \ nat" assumes "infinite YY" "part_fn r s YY f" shows "\Y' t'. Y' \ YY \ infinite Y' \ t' < s \ (\X. X \ Y' \ finite X \ card X = r \ f X = t')" using assms proof (induct r arbitrary: YY f) case 0 then show ?case by (auto simp add: part_fn_def card_eq_0_iff cong: conj_cong) next case (Suc r) show ?case proof - from Suc.prems infinite_imp_nonempty obtain yy where yy: "yy \ YY" by blast let ?ramr = "{((y, Y, t), (y', Y', t')). y' \ Y \ Y' \ Y}" let ?propr = "\(y, Y, t). y \ YY \ y \ Y \ Y \ YY \ infinite Y \ t < s \ (\X. X\Y \ finite X \ card X = r \ (f \ insert y) X = t)" from Suc.prems have infYY': "infinite (YY - {yy})" by auto from Suc.prems have partf': "part_fn r s (YY - {yy}) (f \ insert yy)" by (simp add: o_def part_fn_Suc_imp_part_fn yy) have transr: "trans ?ramr" by (force simp add: trans_def) from Suc.hyps [OF infYY' partf'] obtain Y0 and t0 where "Y0 \ YY - {yy}" "infinite Y0" "t0 < s" "X \ Y0 \ finite X \ card X = r \ (f \ insert yy) X = t0" for X by blast with yy have propr0: "?propr(yy, Y0, t0)" by blast have proprstep: "\y. ?propr y \ (x, y) \ ?ramr" if x: "?propr x" for x proof (cases x) case (fields yx Yx tx) with x obtain yx' where yx': "yx' \ Yx" by (blast dest: infinite_imp_nonempty) from fields x have infYx': "infinite (Yx - {yx'})" by auto with fields x yx' Suc.prems have partfx': "part_fn r s (Yx - {yx'}) (f \ insert yx')" by (simp add: o_def part_fn_Suc_imp_part_fn part_fn_subset [where YY=YY and Y=Yx]) from Suc.hyps [OF infYx' partfx'] obtain Y' and t' where Y': "Y' \ Yx - {yx'}" "infinite Y'" "t' < s" "X \ Y' \ finite X \ card X = r \ (f \ insert yx') X = t'" for X by blast from fields x Y' yx' have "?propr (yx', Y', t') \ (x, (yx', Y', t')) \ ?ramr" by blast then show ?thesis .. qed from dependent_choice [OF transr propr0 proprstep] obtain g where pg: "?propr (g n)" and rg: "n < m \ (g n, g m) \ ?ramr" for n m :: nat by blast let ?gy = "fst \ g" let ?gt = "snd \ snd \ g" have rangeg: "\k. range ?gt \ {.. range ?gt" then obtain n where "x = ?gt n" .. with pg [of n] show "x \ {.. ?gy m'" by (cases "g m", cases "g m'") auto qed show ?thesis proof (intro exI conjI) from pg show "?gy ` {n. ?gt n = s'} \ YY" by (auto simp add: Let_def split_beta) from infeqs' show "infinite (?gy ` {n. ?gt n = s'})" by (blast intro: inj_gy [THEN subset_inj_on] dest: finite_imageD) show "s' < s" by (rule less') show "\X. X \ ?gy ` {n. ?gt n = s'} \ finite X \ card X = Suc r \ f X = s'" proof - have "f X = s'" if X: "X \ ?gy ` {n. ?gt n = s'}" and cardX: "finite X" "card X = Suc r" for X proof - from X obtain AA where AA: "AA \ {n. ?gt n = s'}" and Xeq: "X = ?gy`AA" by (auto simp add: subset_image_iff) with cardX have "AA \ {}" by auto then have AAleast: "(LEAST x. x \ AA) \ AA" by (auto intro: LeastI_ex) show ?thesis proof (cases "g (LEAST x. x \ AA)") case (fields ya Ya ta) with AAleast Xeq have ya: "ya \ X" by (force intro!: rev_image_eqI) then have "f X = f (insert ya (X - {ya}))" by (simp add: insert_absorb) also have "\ = ta" proof - have *: "X - {ya} \ Ya" proof fix x assume x: "x \ X - {ya}" then obtain a' where xeq: "x = ?gy a'" and a': "a' \ AA" by (auto simp add: Xeq) with fields x have "a' \ (LEAST x. x \ AA)" by auto with Least_le [of "\x. x \ AA", OF a'] have "(LEAST x. x \ AA) < a'" by arith from xeq fields rg [OF this] show "x \ Ya" by auto qed have "card (X - {ya}) = r" by (simp add: cardX ya) with pg [of "LEAST x. x \ AA"] fields cardX * show ?thesis by (auto simp del: insert_Diff_single) qed also from AA AAleast fields have "\ = s'" by auto finally show ?thesis . qed qed then show ?thesis by blast qed qed qed qed theorem Ramsey: fixes s r :: nat and Z :: "'a set" and f :: "'a set \ nat" shows "\infinite Z; \X. X \ Z \ finite X \ card X = r \ f X < s\ \ \Y t. Y \ Z \ infinite Y \ t < s \ (\X. X \ Y \ finite X \ card X = r \ f X = t)" by (blast intro: Ramsey_induction [unfolded part_fn_def nsets_def]) corollary Ramsey2: fixes s :: nat and Z :: "'a set" and f :: "'a set \ nat" assumes infZ: "infinite Z" and part: "\x\Z. \y\Z. x \ y \ f {x, y} < s" shows "\Y t. Y \ Z \ infinite Y \ t < s \ (\x\Y. \y\Y. x\y \ f {x, y} = t)" proof - from part have part2: "\X. X \ Z \ finite X \ card X = 2 \ f X < s" by (fastforce simp add: eval_nat_numeral card_Suc_eq) obtain Y t where *: "Y \ Z" "infinite Y" "t < s" "(\X. X \ Y \ finite X \ card X = 2 \ f X = t)" by (insert Ramsey [OF infZ part2]) auto then have "\x\Y. \y\Y. x \ y \ f {x, y} = t" by auto with * show ?thesis by iprover qed corollary Ramsey_nsets: fixes f :: "'a set \ nat" assumes "infinite Z" "f ` nsets Z r \ {.. Z" "infinite Y" "t < s" "f ` nsets Y r \ {t}" using Ramsey [of Z r f s] assms by (auto simp: nsets_def image_subset_iff) subsection \Disjunctive Well-Foundedness\ text \ An application of Ramsey's theorem to program termination. See - @{cite "Podelski-Rybalchenko"}. + \<^cite>\"Podelski-Rybalchenko"\. \ definition disj_wf :: "('a \ 'a) set \ bool" where "disj_wf r \ (\T. \n::nat. (\i r = (\i 'a) \ (nat \ ('a \ 'a) set) \ nat set \ nat" where "transition_idx s T A = (LEAST k. \i j. A = {i, j} \ i < j \ (s j, s i) \ T k)" lemma transition_idx_less: assumes "i < j" "(s j, s i) \ T k" "k < n" shows "transition_idx s T {i, j} < n" proof - from assms(1,2) have "transition_idx s T {i, j} \ k" by (simp add: transition_idx_def, blast intro: Least_le) with assms(3) show ?thesis by simp qed lemma transition_idx_in: assumes "i < j" "(s j, s i) \ T k" shows "(s j, s i) \ T (transition_idx s T {i, j})" using assms by (simp add: transition_idx_def doubleton_eq_iff conj_disj_distribR cong: conj_cong) (erule LeastI) text \To be equal to the union of some well-founded relations is equivalent to being the subset of such a union.\ lemma disj_wf: "disj_wf r \ (\T. \n::nat. (\i r \ (\iT n. \\i \ (T ` {.. \ (\i r)) \ r = (\i r)" by (force simp add: wf_Int1) show ?thesis unfolding disj_wf_def by auto (metis "*") qed theorem trans_disj_wf_implies_wf: assumes "trans r" and "disj_wf r" shows "wf r" proof (simp only: wf_iff_no_infinite_down_chain, rule notI) assume "\s. \i. (s (Suc i), s i) \ r" then obtain s where sSuc: "\i. (s (Suc i), s i) \ r" .. from \disj_wf r\ obtain T and n :: nat where wfT: "\kkk. (s j, s i) \ T k \ ki < j\ have "(s j, s i) \ r" proof (induct rule: less_Suc_induct) case 1 then show ?case by (simp add: sSuc) next case 2 with \trans r\ show ?case unfolding trans_def by blast qed then show ?thesis by (auto simp add: r) qed have "i < j \ transition_idx s T {i, j} < n" for i j using s_in_T transition_idx_less by blast then have trless: "i \ j \ transition_idx s T {i, j} < n" for i j by (metis doubleton_eq_iff less_linear) have "\K k. K \ UNIV \ infinite K \ k < n \ (\i\K. \j\K. i \ j \ transition_idx s T {i, j} = k)" by (rule Ramsey2) (auto intro: trless infinite_UNIV_nat) then obtain K and k where infK: "infinite K" and "k < n" and allk: "\i\K. \j\K. i \ j \ transition_idx s T {i, j} = k" by auto have "(s (enumerate K (Suc m)), s (enumerate K m)) \ T k" for m :: nat proof - let ?j = "enumerate K (Suc m)" let ?i = "enumerate K m" have ij: "?i < ?j" by (simp add: enumerate_step infK) have "?j \ K" "?i \ K" by (simp_all add: enumerate_in_set infK) with ij have k: "k = transition_idx s T {?i, ?j}" by (simp add: allk) from s_in_T [OF ij] obtain k' where "(s ?j, s ?i) \ T k'" "k' T k" by (simp add: k transition_idx_in ij) qed then have "\ wf (T k)" unfolding wf_iff_no_infinite_down_chain by iprover with wfT \k < n\ show False by blast qed end diff --git a/src/HOL/Proofs/Extraction/Euclid.thy b/src/HOL/Proofs/Extraction/Euclid.thy --- a/src/HOL/Proofs/Extraction/Euclid.thy +++ b/src/HOL/Proofs/Extraction/Euclid.thy @@ -1,273 +1,273 @@ (* Title: HOL/Proofs/Extraction/Euclid.thy Author: Markus Wenzel, TU Muenchen Author: Freek Wiedijk, Radboud University Nijmegen Author: Stefan Berghofer, TU Muenchen *) section \Euclid's theorem\ theory Euclid imports "HOL-Computational_Algebra.Primes" Util "HOL-Library.Code_Target_Numeral" "HOL-Library.Realizers" begin text \ A constructive version of the proof of Euclid's theorem by - Markus Wenzel and Freek Wiedijk @{cite "Wenzel-Wiedijk-JAR2002"}. + Markus Wenzel and Freek Wiedijk \<^cite>\"Wenzel-Wiedijk-JAR2002"\. \ lemma factor_greater_one1: "n = m * k \ m < n \ k < n \ Suc 0 < m" by (induct m) auto lemma factor_greater_one2: "n = m * k \ m < n \ k < n \ Suc 0 < k" by (induct k) auto lemma prod_mn_less_k: "0 < n \ 0 < k \ Suc 0 < m \ m * n = k \ n < k" by (induct m) auto lemma prime_eq: "prime (p::nat) \ 1 < p \ (\m. m dvd p \ 1 < m \ m = p)" apply (simp add: prime_nat_iff) apply (rule iffI) apply blast apply (erule conjE) apply (rule conjI) apply assumption apply (rule allI impI)+ apply (erule allE) apply (erule impE) apply assumption apply (case_tac "m = 0") apply simp apply (case_tac "m = Suc 0") apply simp apply simp done lemma prime_eq': "prime (p::nat) \ 1 < p \ (\m k. p = m * k \ 1 < m \ m = p)" by (simp add: prime_eq dvd_def HOL.all_simps [symmetric] del: HOL.all_simps) lemma not_prime_ex_mk: assumes n: "Suc 0 < n" shows "(\m k. Suc 0 < m \ Suc 0 < k \ m < n \ k < n \ n = m * k) \ prime n" proof - from nat_eq_dec have "(\m \ (\mkm \ (\kmkm (\kmkm m * k" by iprover have "\m k. n = m * k \ Suc 0 < m \ m = n" proof (intro allI impI) fix m k assume nmk: "n = m * k" assume m: "Suc 0 < m" from n m nmk have k: "0 < k" by (cases k) auto moreover from n have n: "0 < n" by simp moreover note m moreover from nmk have "m * k = n" by simp ultimately have kn: "k < n" by (rule prod_mn_less_k) show "m = n" proof (cases "k = Suc 0") case True with nmk show ?thesis by (simp only: mult_Suc_right) next case False from m have "0 < m" by simp moreover note n moreover from False n nmk k have "Suc 0 < k" by auto moreover from nmk have "k * m = n" by (simp only: ac_simps) ultimately have mn: "m < n" by (rule prod_mn_less_k) with kn A nmk show ?thesis by iprover qed qed with n have "prime n" by (simp only: prime_eq' One_nat_def simp_thms) then show ?thesis .. qed qed lemma dvd_factorial: "0 < m \ m \ n \ m dvd fact n" proof (induct n rule: nat_induct) case 0 then show ?case by simp next case (Suc n) from \m \ Suc n\ show ?case proof (rule le_SucE) assume "m \ n" with \0 < m\ have "m dvd fact n" by (rule Suc) then have "m dvd (fact n * Suc n)" by (rule dvd_mult2) then show ?thesis by (simp add: mult.commute) next assume "m = Suc n" then have "m dvd (fact n * Suc n)" by (auto intro: dvdI simp: ac_simps) then show ?thesis by (simp add: mult.commute) qed qed lemma dvd_prod [iff]: "n dvd (\m::nat \# mset (n # ns). m)" by (simp add: prod_mset_Un) definition all_prime :: "nat list \ bool" where "all_prime ps \ (\p\set ps. prime p)" lemma all_prime_simps: "all_prime []" "all_prime (p # ps) \ prime p \ all_prime ps" by (simp_all add: all_prime_def) lemma all_prime_append: "all_prime (ps @ qs) \ all_prime ps \ all_prime qs" by (simp add: all_prime_def ball_Un) lemma split_all_prime: assumes "all_prime ms" and "all_prime ns" shows "\qs. all_prime qs \ (\m::nat \# mset qs. m) = (\m::nat \# mset ms. m) * (\m::nat \# mset ns. m)" (is "\qs. ?P qs \ ?Q qs") proof - from assms have "all_prime (ms @ ns)" by (simp add: all_prime_append) moreover have "(\m::nat \# mset (ms @ ns). m) = (\m::nat \# mset ms. m) * (\m::nat \# mset ns. m)" using assms by (simp add: prod_mset_Un) ultimately have "?P (ms @ ns) \ ?Q (ms @ ns)" .. then show ?thesis .. qed lemma all_prime_nempty_g_one: assumes "all_prime ps" and "ps \ []" shows "Suc 0 < (\m::nat \# mset ps. m)" using \ps \ []\ \all_prime ps\ unfolding One_nat_def [symmetric] by (induct ps rule: list_nonempty_induct) (simp_all add: all_prime_simps prod_mset_Un prime_gt_1_nat less_1_mult del: One_nat_def) lemma factor_exists: "Suc 0 < n \ (\ps. all_prime ps \ (\m::nat \# mset ps. m) = n)" proof (induct n rule: nat_wf_ind) case (1 n) from \Suc 0 < n\ have "(\m k. Suc 0 < m \ Suc 0 < k \ m < n \ k < n \ n = m * k) \ prime n" by (rule not_prime_ex_mk) then show ?case proof assume "\m k. Suc 0 < m \ Suc 0 < k \ m < n \ k < n \ n = m * k" then obtain m k where m: "Suc 0 < m" and k: "Suc 0 < k" and mn: "m < n" and kn: "k < n" and nmk: "n = m * k" by iprover from mn and m have "\ps. all_prime ps \ (\m::nat \# mset ps. m) = m" by (rule 1) then obtain ps1 where "all_prime ps1" and prod_ps1_m: "(\m::nat \# mset ps1. m) = m" by iprover from kn and k have "\ps. all_prime ps \ (\m::nat \# mset ps. m) = k" by (rule 1) then obtain ps2 where "all_prime ps2" and prod_ps2_k: "(\m::nat \# mset ps2. m) = k" by iprover from \all_prime ps1\ \all_prime ps2\ have "\ps. all_prime ps \ (\m::nat \# mset ps. m) = (\m::nat \# mset ps1. m) * (\m::nat \# mset ps2. m)" by (rule split_all_prime) with prod_ps1_m prod_ps2_k nmk show ?thesis by simp next assume "prime n" then have "all_prime [n]" by (simp add: all_prime_simps) moreover have "(\m::nat \# mset [n]. m) = n" by (simp) ultimately have "all_prime [n] \ (\m::nat \# mset [n]. m) = n" .. then show ?thesis .. qed qed lemma prime_factor_exists: assumes N: "(1::nat) < n" shows "\p. prime p \ p dvd n" proof - from N obtain ps where "all_prime ps" and prod_ps: "n = (\m::nat \# mset ps. m)" using factor_exists by simp iprover with N have "ps \ []" by (auto simp add: all_prime_nempty_g_one) then obtain p qs where ps: "ps = p # qs" by (cases ps) simp with \all_prime ps\ have "prime p" by (simp add: all_prime_simps) moreover from \all_prime ps\ ps prod_ps have "p dvd n" by (simp only: dvd_prod) ultimately show ?thesis by iprover qed text \Euclid's theorem: there are infinitely many primes.\ lemma Euclid: "\p::nat. prime p \ n < p" proof - let ?k = "fact n + (1::nat)" have "1 < ?k" by simp then obtain p where prime: "prime p" and dvd: "p dvd ?k" using prime_factor_exists by iprover have "n < p" proof - have "\ p \ n" proof assume pn: "p \ n" from \prime p\ have "0 < p" by (rule prime_gt_0_nat) then have "p dvd fact n" using pn by (rule dvd_factorial) with dvd have "p dvd ?k - fact n" by (rule dvd_diff_nat) then have "p dvd 1" by simp with prime show False by auto qed then show ?thesis by simp qed with prime show ?thesis by iprover qed extract Euclid text \ The program extracted from the proof of Euclid's theorem looks as follows. @{thm [display] Euclid_def} The program corresponding to the proof of the factorization theorem is @{thm [display] factor_exists_def} \ instantiation nat :: default begin definition "default = (0::nat)" instance .. end instantiation list :: (type) default begin definition "default = []" instance .. end primrec iterate :: "nat \ ('a \ 'a) \ 'a \ 'a list" where "iterate 0 f x = []" | "iterate (Suc n) f x = (let y = f x in y # iterate n f y)" lemma "factor_exists 1007 = [53, 19]" by eval lemma "factor_exists 567 = [7, 3, 3, 3, 3]" by eval lemma "factor_exists 345 = [23, 5, 3]" by eval lemma "factor_exists 999 = [37, 3, 3, 3]" by eval lemma "factor_exists 876 = [73, 3, 2, 2]" by eval lemma "iterate 4 Euclid 0 = [2, 3, 7, 71]" by eval end diff --git a/src/HOL/Proofs/Extraction/Higman.thy b/src/HOL/Proofs/Extraction/Higman.thy --- a/src/HOL/Proofs/Extraction/Higman.thy +++ b/src/HOL/Proofs/Extraction/Higman.thy @@ -1,327 +1,327 @@ (* Title: HOL/Proofs/Extraction/Higman.thy Author: Stefan Berghofer, TU Muenchen Author: Monika Seisenberger, LMU Muenchen *) section \Higman's lemma\ theory Higman imports Main begin text \ Formalization by Stefan Berghofer and Monika Seisenberger, - based on Coquand and Fridlender @{cite Coquand93}. + based on Coquand and Fridlender \<^cite>\Coquand93\. \ datatype letter = A | B inductive emb :: "letter list \ letter list \ bool" where emb0 [Pure.intro]: "emb [] bs" | emb1 [Pure.intro]: "emb as bs \ emb as (b # bs)" | emb2 [Pure.intro]: "emb as bs \ emb (a # as) (a # bs)" inductive L :: "letter list \ letter list list \ bool" for v :: "letter list" where L0 [Pure.intro]: "emb w v \ L v (w # ws)" | L1 [Pure.intro]: "L v ws \ L v (w # ws)" inductive good :: "letter list list \ bool" where good0 [Pure.intro]: "L w ws \ good (w # ws)" | good1 [Pure.intro]: "good ws \ good (w # ws)" inductive R :: "letter \ letter list list \ letter list list \ bool" for a :: letter where R0 [Pure.intro]: "R a [] []" | R1 [Pure.intro]: "R a vs ws \ R a (w # vs) ((a # w) # ws)" inductive T :: "letter \ letter list list \ letter list list \ bool" for a :: letter where T0 [Pure.intro]: "a \ b \ R b ws zs \ T a (w # zs) ((a # w) # zs)" | T1 [Pure.intro]: "T a ws zs \ T a (w # ws) ((a # w) # zs)" | T2 [Pure.intro]: "a \ b \ T a ws zs \ T a ws ((b # w) # zs)" inductive bar :: "letter list list \ bool" where bar1 [Pure.intro]: "good ws \ bar ws" | bar2 [Pure.intro]: "(\w. bar (w # ws)) \ bar ws" theorem prop1: "bar ([] # ws)" by iprover theorem lemma1: "L as ws \ L (a # as) ws" by (erule L.induct) iprover+ lemma lemma2': "R a vs ws \ L as vs \ L (a # as) ws" supply [[simproc del: defined_all]] apply (induct set: R) apply (erule L.cases) apply simp+ apply (erule L.cases) apply simp_all apply (rule L0) apply (erule emb2) apply (erule L1) done lemma lemma2: "R a vs ws \ good vs \ good ws" supply [[simproc del: defined_all]] apply (induct set: R) apply iprover apply (erule good.cases) apply simp_all apply (rule good0) apply (erule lemma2') apply assumption apply (erule good1) done lemma lemma3': "T a vs ws \ L as vs \ L (a # as) ws" supply [[simproc del: defined_all]] apply (induct set: T) apply (erule L.cases) apply simp_all apply (rule L0) apply (erule emb2) apply (rule L1) apply (erule lemma1) apply (erule L.cases) apply simp_all apply iprover+ done lemma lemma3: "T a ws zs \ good ws \ good zs" supply [[simproc del: defined_all]] apply (induct set: T) apply (erule good.cases) apply simp_all apply (rule good0) apply (erule lemma1) apply (erule good1) apply (erule good.cases) apply simp_all apply (rule good0) apply (erule lemma3') apply iprover+ done lemma lemma4: "R a ws zs \ ws \ [] \ T a ws zs" supply [[simproc del: defined_all]] apply (induct set: R) apply iprover apply (case_tac vs) apply (erule R.cases) apply simp apply (case_tac a) apply (rule_tac b=B in T0) apply simp apply (rule R0) apply (rule_tac b=A in T0) apply simp apply (rule R0) apply simp apply (rule T1) apply simp done lemma letter_neq: "a \ b \ c \ a \ c = b" for a b c :: letter apply (case_tac a) apply (case_tac b) apply (case_tac c, simp, simp) apply (case_tac c, simp, simp) apply (case_tac b) apply (case_tac c, simp, simp) apply (case_tac c, simp, simp) done lemma letter_eq_dec: "a = b \ a \ b" for a b :: letter apply (case_tac a) apply (case_tac b) apply simp apply simp apply (case_tac b) apply simp apply simp done theorem prop2: assumes ab: "a \ b" and bar: "bar xs" shows "\ys zs. bar ys \ T a xs zs \ T b ys zs \ bar zs" using bar proof induct fix xs zs assume "T a xs zs" and "good xs" then have "good zs" by (rule lemma3) then show "bar zs" by (rule bar1) next fix xs ys assume I: "\w ys zs. bar ys \ T a (w # xs) zs \ T b ys zs \ bar zs" assume "bar ys" then show "\zs. T a xs zs \ T b ys zs \ bar zs" proof induct fix ys zs assume "T b ys zs" and "good ys" then have "good zs" by (rule lemma3) then show "bar zs" by (rule bar1) next fix ys zs assume I': "\w zs. T a xs zs \ T b (w # ys) zs \ bar zs" and ys: "\w. bar (w # ys)" and Ta: "T a xs zs" and Tb: "T b ys zs" show "bar zs" proof (rule bar2) fix w show "bar (w # zs)" proof (cases w) case Nil then show ?thesis by simp (rule prop1) next case (Cons c cs) from letter_eq_dec show ?thesis proof assume ca: "c = a" from ab have "bar ((a # cs) # zs)" by (iprover intro: I ys Ta Tb) then show ?thesis by (simp add: Cons ca) next assume "c \ a" with ab have cb: "c = b" by (rule letter_neq) from ab have "bar ((b # cs) # zs)" by (iprover intro: I' Ta Tb) then show ?thesis by (simp add: Cons cb) qed qed qed qed qed theorem prop3: assumes bar: "bar xs" shows "\zs. xs \ [] \ R a xs zs \ bar zs" using bar proof induct fix xs zs assume "R a xs zs" and "good xs" then have "good zs" by (rule lemma2) then show "bar zs" by (rule bar1) next fix xs zs assume I: "\w zs. w # xs \ [] \ R a (w # xs) zs \ bar zs" and xsb: "\w. bar (w # xs)" and xsn: "xs \ []" and R: "R a xs zs" show "bar zs" proof (rule bar2) fix w show "bar (w # zs)" proof (induct w) case Nil show ?case by (rule prop1) next case (Cons c cs) from letter_eq_dec show ?case proof assume "c = a" then show ?thesis by (iprover intro: I [simplified] R) next from R xsn have T: "T a xs zs" by (rule lemma4) assume "c \ a" then show ?thesis by (iprover intro: prop2 Cons xsb xsn R T) qed qed qed qed theorem higman: "bar []" proof (rule bar2) fix w show "bar [w]" proof (induct w) show "bar [[]]" by (rule prop1) next fix c cs assume "bar [cs]" then show "bar [c # cs]" by (rule prop3) (simp, iprover) qed qed primrec is_prefix :: "'a list \ (nat \ 'a) \ bool" where "is_prefix [] f = True" | "is_prefix (x # xs) f = (x = f (length xs) \ is_prefix xs f)" theorem L_idx: assumes L: "L w ws" shows "is_prefix ws f \ \i. emb (f i) w \ i < length ws" using L proof induct case (L0 v ws) then have "emb (f (length ws)) w" by simp moreover have "length ws < length (v # ws)" by simp ultimately show ?case by iprover next case (L1 ws v) then obtain i where emb: "emb (f i) w" and "i < length ws" by simp iprover then have "i < length (v # ws)" by simp with emb show ?case by iprover qed theorem good_idx: assumes good: "good ws" shows "is_prefix ws f \ \i j. emb (f i) (f j) \ i < j" using good proof induct case (good0 w ws) then have "w = f (length ws)" and "is_prefix ws f" by simp_all with good0 show ?case by (iprover dest: L_idx) next case (good1 ws w) then show ?case by simp qed theorem bar_idx: assumes bar: "bar ws" shows "is_prefix ws f \ \i j. emb (f i) (f j) \ i < j" using bar proof induct case (bar1 ws) then show ?case by (rule good_idx) next case (bar2 ws) then have "is_prefix (f (length ws) # ws) f" by simp then show ?case by (rule bar2) qed text \ Strong version: yields indices of words that can be embedded into each other. \ theorem higman_idx: "\(i::nat) j. emb (f i) (f j) \ i < j" proof (rule bar_idx) show "bar []" by (rule higman) show "is_prefix [] f" by simp qed text \ Weak version: only yield sequence containing words that can be embedded into each other. \ theorem good_prefix_lemma: assumes bar: "bar ws" shows "is_prefix ws f \ \vs. is_prefix vs f \ good vs" using bar proof induct case bar1 then show ?case by iprover next case (bar2 ws) from bar2.prems have "is_prefix (f (length ws) # ws) f" by simp then show ?case by (iprover intro: bar2) qed theorem good_prefix: "\vs. is_prefix vs f \ good vs" using higman by (rule good_prefix_lemma) simp+ end diff --git a/src/HOL/Proofs/Extraction/Pigeonhole.thy b/src/HOL/Proofs/Extraction/Pigeonhole.thy --- a/src/HOL/Proofs/Extraction/Pigeonhole.thy +++ b/src/HOL/Proofs/Extraction/Pigeonhole.thy @@ -1,247 +1,247 @@ (* Title: HOL/Proofs/Extraction/Pigeonhole.thy Author: Stefan Berghofer, TU Muenchen *) section \The pigeonhole principle\ theory Pigeonhole imports Util "HOL-Library.Realizers" "HOL-Library.Code_Target_Numeral" begin text \ We formalize two proofs of the pigeonhole principle, which lead to extracted programs of quite different complexity. The original formalization of these proofs in {\sc Nuprl} is due to - Aleksey Nogin @{cite "Nogin-ENTCS-2000"}. + Aleksey Nogin \<^cite>\"Nogin-ENTCS-2000"\. This proof yields a polynomial program. \ theorem pigeonhole: "\f. (\i. i \ Suc n \ f i \ n) \ \i j. i \ Suc n \ j < i \ f i = f j" proof (induct n) case 0 then have "Suc 0 \ Suc 0 \ 0 < Suc 0 \ f (Suc 0) = f 0" by simp then show ?case by iprover next case (Suc n) have r: "k \ Suc (Suc n) \ (\i j. Suc k \ i \ i \ Suc (Suc n) \ j < i \ f i \ f j) \ (\i j. i \ k \ j < i \ f i = f j)" for k proof (induct k) case 0 let ?f = "\i. if f i = Suc n then f (Suc (Suc n)) else f i" have "\ (\i j. i \ Suc n \ j < i \ ?f i = ?f j)" proof assume "\i j. i \ Suc n \ j < i \ ?f i = ?f j" then obtain i j where i: "i \ Suc n" and j: "j < i" and f: "?f i = ?f j" by iprover from j have i_nz: "Suc 0 \ i" by simp from i have iSSn: "i \ Suc (Suc n)" by simp have S0SSn: "Suc 0 \ Suc (Suc n)" by simp show False proof cases assume fi: "f i = Suc n" show False proof cases assume fj: "f j = Suc n" from i_nz and iSSn and j have "f i \ f j" by (rule 0) moreover from fi have "f i = f j" by (simp add: fj [symmetric]) ultimately show ?thesis .. next from i and j have "j < Suc (Suc n)" by simp with S0SSn and le_refl have "f (Suc (Suc n)) \ f j" by (rule 0) moreover assume "f j \ Suc n" with fi and f have "f (Suc (Suc n)) = f j" by simp ultimately show False .. qed next assume fi: "f i \ Suc n" show False proof cases from i have "i < Suc (Suc n)" by simp with S0SSn and le_refl have "f (Suc (Suc n)) \ f i" by (rule 0) moreover assume "f j = Suc n" with fi and f have "f (Suc (Suc n)) = f i" by simp ultimately show False .. next from i_nz and iSSn and j have "f i \ f j" by (rule 0) moreover assume "f j \ Suc n" with fi and f have "f i = f j" by simp ultimately show False .. qed qed qed moreover have "?f i \ n" if "i \ Suc n" for i proof - from that have i: "i < Suc (Suc n)" by simp have "f (Suc (Suc n)) \ f i" by (rule 0) (simp_all add: i) moreover have "f (Suc (Suc n)) \ Suc n" by (rule Suc) simp moreover from i have "i \ Suc (Suc n)" by simp then have "f i \ Suc n" by (rule Suc) ultimately show ?thesis by simp qed then have "\i j. i \ Suc n \ j < i \ ?f i = ?f j" by (rule Suc) ultimately show ?case .. next case (Suc k) from search [OF nat_eq_dec] show ?case proof assume "\j (\ji j. i \ k \ j < i \ f i = f j" proof (rule Suc) from Suc show "k \ Suc (Suc n)" by simp fix i j assume k: "Suc k \ i" and i: "i \ Suc (Suc n)" and j: "j < i" show "f i \ f j" proof cases assume eq: "i = Suc k" show ?thesis proof assume "f i = f j" then have "f (Suc k) = f j" by (simp add: eq) with nex and j and eq show False by iprover qed next assume "i \ Suc k" with k have "Suc (Suc k) \ i" by simp then show ?thesis using i and j by (rule Suc) qed qed then show ?thesis by (iprover intro: le_SucI) qed qed show ?case by (rule r) simp_all qed text \ The following proof, although quite elegant from a mathematical point of view, leads to an exponential program: \ theorem pigeonhole_slow: "\f. (\i. i \ Suc n \ f i \ n) \ \i j. i \ Suc n \ j < i \ f i = f j" proof (induct n) case 0 have "Suc 0 \ Suc 0" .. moreover have "0 < Suc 0" .. moreover from 0 have "f (Suc 0) = f 0" by simp ultimately show ?case by iprover next case (Suc n) from search [OF nat_eq_dec] show ?case proof assume "\j < Suc (Suc n). f (Suc (Suc n)) = f j" then show ?case by (iprover intro: le_refl) next assume "\ (\j < Suc (Suc n). f (Suc (Suc n)) = f j)" then have nex: "\j < Suc (Suc n). f (Suc (Suc n)) \ f j" by iprover let ?f = "\i. if f i = Suc n then f (Suc (Suc n)) else f i" have "\i. i \ Suc n \ ?f i \ n" proof - fix i assume i: "i \ Suc n" show "?thesis i" proof (cases "f i = Suc n") case True from i and nex have "f (Suc (Suc n)) \ f i" by simp with True have "f (Suc (Suc n)) \ Suc n" by simp moreover from Suc have "f (Suc (Suc n)) \ Suc n" by simp ultimately have "f (Suc (Suc n)) \ n" by simp with True show ?thesis by simp next case False from Suc and i have "f i \ Suc n" by simp with False show ?thesis by simp qed qed then have "\i j. i \ Suc n \ j < i \ ?f i = ?f j" by (rule Suc) then obtain i j where i: "i \ Suc n" and ji: "j < i" and f: "?f i = ?f j" by iprover have "f i = f j" proof (cases "f i = Suc n") case True show ?thesis proof (cases "f j = Suc n") assume "f j = Suc n" with True show ?thesis by simp next assume "f j \ Suc n" moreover from i ji nex have "f (Suc (Suc n)) \ f j" by simp ultimately show ?thesis using True f by simp qed next case False show ?thesis proof (cases "f j = Suc n") assume "f j = Suc n" moreover from i nex have "f (Suc (Suc n)) \ f i" by simp ultimately show ?thesis using False f by simp next assume "f j \ Suc n" with False f show ?thesis by simp qed qed moreover from i have "i \ Suc (Suc n)" by simp ultimately show ?thesis using ji by iprover qed qed extract pigeonhole pigeonhole_slow text \ The programs extracted from the above proofs look as follows: @{thm [display] pigeonhole_def} @{thm [display] pigeonhole_slow_def} The program for searching for an element in an array is @{thm [display,eta_contract=false] search_def} The correctness statement for \<^term>\pigeonhole\ is @{thm [display] pigeonhole_correctness [no_vars]} In order to analyze the speed of the above programs, we generate ML code from them. \ instantiation nat :: default begin definition "default = (0::nat)" instance .. end instantiation prod :: (default, default) default begin definition "default = (default, default)" instance .. end definition "test n u = pigeonhole (nat_of_integer n) (\m. m - 1)" definition "test' n u = pigeonhole_slow (nat_of_integer n) (\m. m - 1)" definition "test'' u = pigeonhole 8 (List.nth [0, 1, 2, 3, 4, 5, 6, 3, 7, 8])" ML_val "timeit (@{code test} 10)" ML_val "timeit (@{code test'} 10)" ML_val "timeit (@{code test} 20)" ML_val "timeit (@{code test'} 20)" ML_val "timeit (@{code test} 25)" ML_val "timeit (@{code test'} 25)" ML_val "timeit (@{code test} 500)" ML_val "timeit @{code test''}" end diff --git a/src/HOL/Proofs/Extraction/Warshall.thy b/src/HOL/Proofs/Extraction/Warshall.thy --- a/src/HOL/Proofs/Extraction/Warshall.thy +++ b/src/HOL/Proofs/Extraction/Warshall.thy @@ -1,255 +1,255 @@ (* Title: HOL/Proofs/Extraction/Warshall.thy Author: Stefan Berghofer, TU Muenchen *) section \Warshall's algorithm\ theory Warshall imports "HOL-Library.Realizers" begin text \ Derivation of Warshall's algorithm using program extraction, - based on Berger, Schwichtenberg and Seisenberger @{cite "Berger-JAR-2001"}. + based on Berger, Schwichtenberg and Seisenberger \<^cite>\"Berger-JAR-2001"\. \ datatype b = T | F primrec is_path' :: "('a \ 'a \ b) \ 'a \ 'a list \ 'a \ bool" where "is_path' r x [] z \ r x z = T" | "is_path' r x (y # ys) z \ r x y = T \ is_path' r y ys z" definition is_path :: "(nat \ nat \ b) \ (nat * nat list * nat) \ nat \ nat \ nat \ bool" where "is_path r p i j k \ fst p = j \ snd (snd p) = k \ list_all (\x. x < i) (fst (snd p)) \ is_path' r (fst p) (fst (snd p)) (snd (snd p))" definition conc :: "'a \ 'a list \ 'a \ 'a \ 'a list \ 'a \ 'a \ 'a list * 'a" where "conc p q = (fst p, fst (snd p) @ fst q # fst (snd q), snd (snd q))" theorem is_path'_snoc [simp]: "\x. is_path' r x (ys @ [y]) z = (is_path' r x ys y \ r y z = T)" by (induct ys) simp+ theorem list_all_scoc [simp]: "list_all P (xs @ [x]) \ P x \ list_all P xs" by (induct xs) (simp+, iprover) theorem list_all_lemma: "list_all P xs \ (\x. P x \ Q x) \ list_all Q xs" proof - assume PQ: "\x. P x \ Q x" show "list_all P xs \ list_all Q xs" proof (induct xs) case Nil show ?case by simp next case (Cons y ys) then have Py: "P y" by simp from Cons have Pys: "list_all P ys" by simp show ?case by simp (rule conjI PQ Py Cons Pys)+ qed qed theorem lemma1: "\p. is_path r p i j k \ is_path r p (Suc i) j k" unfolding is_path_def apply (simp cong add: conj_cong add: split_paired_all) apply (erule conjE)+ apply (erule list_all_lemma) apply simp done theorem lemma2: "\p. is_path r p 0 j k \ r j k = T" unfolding is_path_def apply (simp cong add: conj_cong add: split_paired_all) apply (case_tac a) apply simp_all done theorem is_path'_conc: "is_path' r j xs i \ is_path' r i ys k \ is_path' r j (xs @ i # ys) k" proof - assume pys: "is_path' r i ys k" show "\j. is_path' r j xs i \ is_path' r j (xs @ i # ys) k" proof (induct xs) case (Nil j) then have "r j i = T" by simp with pys show ?case by simp next case (Cons z zs j) then have jzr: "r j z = T" by simp from Cons have pzs: "is_path' r z zs i" by simp show ?case by simp (rule conjI jzr Cons pzs)+ qed qed theorem lemma3: "\p q. is_path r p i j i \ is_path r q i i k \ is_path r (conc p q) (Suc i) j k" apply (unfold is_path_def conc_def) apply (simp cong add: conj_cong add: split_paired_all) apply (erule conjE)+ apply (rule conjI) apply (erule list_all_lemma) apply simp apply (rule conjI) apply (erule list_all_lemma) apply simp apply (rule is_path'_conc) apply assumption+ done theorem lemma5: "\p. is_path r p (Suc i) j k \ \ is_path r p i j k \ (\q. is_path r q i j i) \ (\q'. is_path r q' i i k)" proof (simp cong add: conj_cong add: split_paired_all is_path_def, (erule conjE)+) fix xs assume asms: "list_all (\x. x < Suc i) xs" "is_path' r j xs k" "\ list_all (\x. x < i) xs" show "(\ys. list_all (\x. x < i) ys \ is_path' r j ys i) \ (\ys. list_all (\x. x < i) ys \ is_path' r i ys k)" proof have "\j. list_all (\x. x < Suc i) xs \ is_path' r j xs k \ \ list_all (\x. x < i) xs \ \ys. list_all (\x. x < i) ys \ is_path' r j ys i" (is "PROP ?ih xs") proof (induct xs) case Nil then show ?case by simp next case (Cons a as j) show ?case proof (cases "a=i") case True show ?thesis proof from True and Cons have "r j i = T" by simp then show "list_all (\x. x < i) [] \ is_path' r j [] i" by simp qed next case False have "PROP ?ih as" by (rule Cons) then obtain ys where ys: "list_all (\x. x < i) ys \ is_path' r a ys i" proof from Cons show "list_all (\x. x < Suc i) as" by simp from Cons show "is_path' r a as k" by simp from Cons and False show "\ list_all (\x. x < i) as" by (simp) qed show ?thesis proof from Cons False ys show "list_all (\x. x is_path' r j (a#ys) i" by simp qed qed qed from this asms show "\ys. list_all (\x. x < i) ys \ is_path' r j ys i" . have "\k. list_all (\x. x < Suc i) xs \ is_path' r j xs k \ \ list_all (\x. x < i) xs \ \ys. list_all (\x. x < i) ys \ is_path' r i ys k" (is "PROP ?ih xs") proof (induct xs rule: rev_induct) case Nil then show ?case by simp next case (snoc a as k) show ?case proof (cases "a=i") case True show ?thesis proof from True and snoc have "r i k = T" by simp then show "list_all (\x. x < i) [] \ is_path' r i [] k" by simp qed next case False have "PROP ?ih as" by (rule snoc) then obtain ys where ys: "list_all (\x. x < i) ys \ is_path' r i ys a" proof from snoc show "list_all (\x. x < Suc i) as" by simp from snoc show "is_path' r j as a" by simp from snoc and False show "\ list_all (\x. x < i) as" by simp qed show ?thesis proof from snoc False ys show "list_all (\x. x < i) (ys @ [a]) \ is_path' r i (ys @ [a]) k" by simp qed qed qed from this asms show "\ys. list_all (\x. x < i) ys \ is_path' r i ys k" . qed qed theorem lemma5': "\p. is_path r p (Suc i) j k \ \ is_path r p i j k \ \ (\q. \ is_path r q i j i) \ \ (\q'. \ is_path r q' i i k)" by (iprover dest: lemma5) theorem warshall: "\j k. \ (\p. is_path r p i j k) \ (\p. is_path r p i j k)" proof (induct i) case (0 j k) show ?case proof (cases "r j k") assume "r j k = T" then have "is_path r (j, [], k) 0 j k" by (simp add: is_path_def) then have "\p. is_path r p 0 j k" .. then show ?thesis .. next assume "r j k = F" then have "r j k \ T" by simp then have "\ (\p. is_path r p 0 j k)" by (iprover dest: lemma2) then show ?thesis .. qed next case (Suc i j k) then show ?case proof assume h1: "\ (\p. is_path r p i j k)" from Suc show ?case proof assume "\ (\p. is_path r p i j i)" with h1 have "\ (\p. is_path r p (Suc i) j k)" by (iprover dest: lemma5') then show ?case .. next assume "\p. is_path r p i j i" then obtain p where h2: "is_path r p i j i" .. from Suc show ?case proof assume "\ (\p. is_path r p i i k)" with h1 have "\ (\p. is_path r p (Suc i) j k)" by (iprover dest: lemma5') then show ?case .. next assume "\q. is_path r q i i k" then obtain q where "is_path r q i i k" .. with h2 have "is_path r (conc p q) (Suc i) j k" by (rule lemma3) then have "\pq. is_path r pq (Suc i) j k" .. then show ?case .. qed qed next assume "\p. is_path r p i j k" then have "\p. is_path r p (Suc i) j k" by (iprover intro: lemma1) then show ?case .. qed qed extract warshall text \ The program extracted from the above proof looks as follows @{thm [display, eta_contract=false] warshall_def [no_vars]} The corresponding correctness theorem is @{thm [display] warshall_correctness [no_vars]} \ ML_val "@{code warshall}" end diff --git a/src/HOL/Proofs/Lambda/Eta.thy b/src/HOL/Proofs/Lambda/Eta.thy --- a/src/HOL/Proofs/Lambda/Eta.thy +++ b/src/HOL/Proofs/Lambda/Eta.thy @@ -1,390 +1,390 @@ (* Title: HOL/Proofs/Lambda/Eta.thy Author: Tobias Nipkow and Stefan Berghofer Copyright 1995, 2005 TU Muenchen *) section \Eta-reduction\ theory Eta imports ParRed begin subsection \Definition of eta-reduction and relatives\ primrec free :: "dB => nat => bool" where "free (Var j) i = (j = i)" | "free (s \ t) i = (free s i \ free t i)" | "free (Abs s) i = free s (i + 1)" inductive eta :: "[dB, dB] => bool" (infixl "\\<^sub>\" 50) where eta [simp, intro]: "\ free s 0 ==> Abs (s \ Var 0) \\<^sub>\ s[dummy/0]" | appL [simp, intro]: "s \\<^sub>\ t ==> s \ u \\<^sub>\ t \ u" | appR [simp, intro]: "s \\<^sub>\ t ==> u \ s \\<^sub>\ u \ t" | abs [simp, intro]: "s \\<^sub>\ t ==> Abs s \\<^sub>\ Abs t" abbreviation eta_reds :: "[dB, dB] => bool" (infixl "\\<^sub>\\<^sup>*" 50) where "s \\<^sub>\\<^sup>* t == eta\<^sup>*\<^sup>* s t" abbreviation eta_red0 :: "[dB, dB] => bool" (infixl "\\<^sub>\\<^sup>=" 50) where "s \\<^sub>\\<^sup>= t == eta\<^sup>=\<^sup>= s t" inductive_cases eta_cases [elim!]: "Abs s \\<^sub>\ z" "s \ t \\<^sub>\ u" "Var i \\<^sub>\ t" subsection \Properties of \eta\, \subst\ and \free\\ lemma subst_not_free [simp]: "\ free s i \ s[t/i] = s[u/i]" by (induct s arbitrary: i t u) (simp_all add: subst_Var) lemma free_lift [simp]: "free (lift t k) i = (i < k \ free t i \ k < i \ free t (i - 1))" apply (induct t arbitrary: i k) apply (auto cong: conj_cong) done lemma free_subst [simp]: "free (s[t/k]) i = (free s k \ free t i \ free s (if i < k then i else i + 1))" apply (induct s arbitrary: i k t) prefer 2 apply simp apply blast prefer 2 apply simp apply (simp add: diff_Suc subst_Var split: nat.split) done lemma free_eta: "s \\<^sub>\ t ==> free t i = free s i" by (induct arbitrary: i set: eta) (simp_all cong: conj_cong) lemma not_free_eta: "[| s \\<^sub>\ t; \ free s i |] ==> \ free t i" by (simp add: free_eta) lemma eta_subst [simp]: "s \\<^sub>\ t ==> s[u/i] \\<^sub>\ t[u/i]" by (induct arbitrary: u i set: eta) (simp_all add: subst_subst [symmetric]) theorem lift_subst_dummy: "\ free s i \ lift (s[dummy/i]) i = s" by (induct s arbitrary: i dummy) simp_all subsection \Confluence of \eta\\ lemma square_eta: "square eta eta (eta\<^sup>=\<^sup>=) (eta\<^sup>=\<^sup>=)" apply (unfold square_def id_def) apply (rule impI [THEN allI [THEN allI]]) apply (erule eta.induct) apply (slowsimp intro: subst_not_free eta_subst free_eta [THEN iffD1]) apply safe prefer 5 apply (blast intro!: eta_subst intro: free_eta [THEN iffD1]) apply blast+ done theorem eta_confluent: "confluent eta" apply (rule square_eta [THEN square_reflcl_confluent]) done subsection \Congruence rules for \eta\<^sup>*\\ lemma rtrancl_eta_Abs: "s \\<^sub>\\<^sup>* s' ==> Abs s \\<^sub>\\<^sup>* Abs s'" by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+ lemma rtrancl_eta_AppL: "s \\<^sub>\\<^sup>* s' ==> s \ t \\<^sub>\\<^sup>* s' \ t" by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+ lemma rtrancl_eta_AppR: "t \\<^sub>\\<^sup>* t' ==> s \ t \\<^sub>\\<^sup>* s \ t'" by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+ lemma rtrancl_eta_App: "[| s \\<^sub>\\<^sup>* s'; t \\<^sub>\\<^sup>* t' |] ==> s \ t \\<^sub>\\<^sup>* s' \ t'" by (blast intro!: rtrancl_eta_AppL rtrancl_eta_AppR intro: rtranclp_trans) subsection \Commutation of \beta\ and \eta\\ lemma free_beta: "s \\<^sub>\ t ==> free t i \ free s i" by (induct arbitrary: i set: beta) auto lemma beta_subst [intro]: "s \\<^sub>\ t ==> s[u/i] \\<^sub>\ t[u/i]" by (induct arbitrary: u i set: beta) (simp_all add: subst_subst [symmetric]) lemma subst_Var_Suc [simp]: "t[Var i/i] = t[Var(i)/i + 1]" by (induct t arbitrary: i) (auto elim!: linorder_neqE simp: subst_Var) lemma eta_lift [simp]: "s \\<^sub>\ t ==> lift s i \\<^sub>\ lift t i" by (induct arbitrary: i set: eta) simp_all lemma rtrancl_eta_subst: "s \\<^sub>\ t \ u[s/i] \\<^sub>\\<^sup>* u[t/i]" apply (induct u arbitrary: s t i) apply (simp_all add: subst_Var) apply blast apply (blast intro: rtrancl_eta_App) apply (blast intro!: rtrancl_eta_Abs eta_lift) done lemma rtrancl_eta_subst': fixes s t :: dB assumes eta: "s \\<^sub>\\<^sup>* t" shows "s[u/i] \\<^sub>\\<^sup>* t[u/i]" using eta by induct (iprover intro: eta_subst)+ lemma rtrancl_eta_subst'': fixes s t :: dB assumes eta: "s \\<^sub>\\<^sup>* t" shows "u[s/i] \\<^sub>\\<^sup>* u[t/i]" using eta by induct (iprover intro: rtrancl_eta_subst rtranclp_trans)+ lemma square_beta_eta: "square beta eta (eta\<^sup>*\<^sup>*) (beta\<^sup>=\<^sup>=)" apply (unfold square_def) apply (rule impI [THEN allI [THEN allI]]) apply (erule beta.induct) apply (slowsimp intro: rtrancl_eta_subst eta_subst) apply (blast intro: rtrancl_eta_AppL) apply (blast intro: rtrancl_eta_AppR) apply simp apply (slowsimp intro: rtrancl_eta_Abs free_beta iff del: dB.distinct simp: dB.distinct) (*23 seconds?*) done lemma confluent_beta_eta: "confluent (sup beta eta)" apply (assumption | rule square_rtrancl_reflcl_commute confluent_Un beta_confluent eta_confluent square_beta_eta)+ done subsection \Implicit definition of \eta\\ text \\<^term>\Abs (lift s 0 \ Var 0) \\<^sub>\ s\\ lemma not_free_iff_lifted: "(\ free s i) = (\t. s = lift t i)" apply (induct s arbitrary: i) apply simp apply (rule iffI) apply (erule linorder_neqE) apply (rename_tac nat a, rule_tac x = "Var nat" in exI) apply simp apply (rename_tac nat a, rule_tac x = "Var (nat - 1)" in exI) apply simp apply clarify apply (rule notE) prefer 2 apply assumption apply (erule thin_rl) apply (case_tac t) apply simp apply simp apply simp apply simp apply (erule thin_rl) apply (erule thin_rl) apply (rule iffI) apply (elim conjE exE) apply (rename_tac u1 u2) apply (rule_tac x = "u1 \ u2" in exI) apply simp apply (erule exE) apply (erule rev_mp) apply (case_tac t) apply simp apply simp apply blast apply simp apply simp apply (erule thin_rl) apply (rule iffI) apply (erule exE) apply (rule_tac x = "Abs t" in exI) apply simp apply (erule exE) apply (erule rev_mp) apply (case_tac t) apply simp apply simp apply simp apply blast done theorem explicit_is_implicit: "(\s u. (\ free s 0) --> R (Abs (s \ Var 0)) (s[u/0])) = (\s. R (Abs (lift s 0 \ Var 0)) s)" by (auto simp add: not_free_iff_lifted) subsection \Eta-postponement theorem\ text \ Based on a paper proof due to Andreas Abel. - Unlike the proof by Masako Takahashi @{cite "Takahashi-IandC"}, it does not + Unlike the proof by Masako Takahashi \<^cite>\"Takahashi-IandC"\, it does not use parallel eta reduction, which only seems to complicate matters unnecessarily. \ theorem eta_case: fixes s :: dB assumes free: "\ free s 0" and s: "s[dummy/0] => u" shows "\t'. Abs (s \ Var 0) => t' \ t' \\<^sub>\\<^sup>* u" proof - from s have "lift (s[dummy/0]) 0 => lift u 0" by (simp del: lift_subst) with free have "s => lift u 0" by (simp add: lift_subst_dummy del: lift_subst) hence "Abs (s \ Var 0) => Abs (lift u 0 \ Var 0)" by simp moreover have "\ free (lift u 0) 0" by simp hence "Abs (lift u 0 \ Var 0) \\<^sub>\ lift u 0[dummy/0]" by (rule eta.eta) hence "Abs (lift u 0 \ Var 0) \\<^sub>\\<^sup>* u" by simp ultimately show ?thesis by iprover qed theorem eta_par_beta: assumes st: "s \\<^sub>\ t" and tu: "t => u" shows "\t'. s => t' \ t' \\<^sub>\\<^sup>* u" using tu st proof (induct arbitrary: s) case (var n) thus ?case by (iprover intro: par_beta_refl) next case (abs s' t) note abs' = this from \s \\<^sub>\ Abs s'\ show ?case proof cases case (eta s'' dummy) from abs have "Abs s' => Abs t" by simp with eta have "s''[dummy/0] => Abs t" by simp with \\ free s'' 0\ have "\t'. Abs (s'' \ Var 0) => t' \ t' \\<^sub>\\<^sup>* Abs t" by (rule eta_case) with eta show ?thesis by simp next case (abs r) from \r \\<^sub>\ s'\ obtain t' where r: "r => t'" and t': "t' \\<^sub>\\<^sup>* t" by (iprover dest: abs') from r have "Abs r => Abs t'" .. moreover from t' have "Abs t' \\<^sub>\\<^sup>* Abs t" by (rule rtrancl_eta_Abs) ultimately show ?thesis using abs by simp iprover qed next case (app u u' t t') from \s \\<^sub>\ u \ t\ show ?case proof cases case (eta s' dummy) from app have "u \ t => u' \ t'" by simp with eta have "s'[dummy/0] => u' \ t'" by simp with \\ free s' 0\ have "\r. Abs (s' \ Var 0) => r \ r \\<^sub>\\<^sup>* u' \ t'" by (rule eta_case) with eta show ?thesis by simp next case (appL s') from \s' \\<^sub>\ u\ obtain r where s': "s' => r" and r: "r \\<^sub>\\<^sup>* u'" by (iprover dest: app) from s' and app have "s' \ t => r \ t'" by simp moreover from r have "r \ t' \\<^sub>\\<^sup>* u' \ t'" by (simp add: rtrancl_eta_AppL) ultimately show ?thesis using appL by simp iprover next case (appR s') from \s' \\<^sub>\ t\ obtain r where s': "s' => r" and r: "r \\<^sub>\\<^sup>* t'" by (iprover dest: app) from s' and app have "u \ s' => u' \ r" by simp moreover from r have "u' \ r \\<^sub>\\<^sup>* u' \ t'" by (simp add: rtrancl_eta_AppR) ultimately show ?thesis using appR by simp iprover qed next case (beta u u' t t') from \s \\<^sub>\ Abs u \ t\ show ?case proof cases case (eta s' dummy) from beta have "Abs u \ t => u'[t'/0]" by simp with eta have "s'[dummy/0] => u'[t'/0]" by simp with \\ free s' 0\ have "\r. Abs (s' \ Var 0) => r \ r \\<^sub>\\<^sup>* u'[t'/0]" by (rule eta_case) with eta show ?thesis by simp next case (appL s') from \s' \\<^sub>\ Abs u\ show ?thesis proof cases case (eta s'' dummy) have "Abs (lift u 1) = lift (Abs u) 0" by simp also from eta have "\ = s''" by (simp add: lift_subst_dummy del: lift_subst) finally have s: "s = Abs (Abs (lift u 1) \ Var 0) \ t" using appL and eta by simp from beta have "lift u 1 => lift u' 1" by simp hence "Abs (lift u 1) \ Var 0 => lift u' 1[Var 0/0]" using par_beta.var .. hence "Abs (Abs (lift u 1) \ Var 0) \ t => lift u' 1[Var 0/0][t'/0]" using \t => t'\ .. with s have "s => u'[t'/0]" by simp thus ?thesis by iprover next case (abs r) from \r \\<^sub>\ u\ obtain r'' where r: "r => r''" and r'': "r'' \\<^sub>\\<^sup>* u'" by (iprover dest: beta) from r and beta have "Abs r \ t => r''[t'/0]" by simp moreover from r'' have "r''[t'/0] \\<^sub>\\<^sup>* u'[t'/0]" by (rule rtrancl_eta_subst') ultimately show ?thesis using abs and appL by simp iprover qed next case (appR s') from \s' \\<^sub>\ t\ obtain r where s': "s' => r" and r: "r \\<^sub>\\<^sup>* t'" by (iprover dest: beta) from s' and beta have "Abs u \ s' => u'[r/0]" by simp moreover from r have "u'[r/0] \\<^sub>\\<^sup>* u'[t'/0]" by (rule rtrancl_eta_subst'') ultimately show ?thesis using appR by simp iprover qed qed theorem eta_postponement': assumes eta: "s \\<^sub>\\<^sup>* t" and beta: "t => u" shows "\t'. s => t' \ t' \\<^sub>\\<^sup>* u" using eta beta proof (induct arbitrary: u) case base thus ?case by blast next case (step s' s'' s''') then obtain t' where s': "s' => t'" and t': "t' \\<^sub>\\<^sup>* s'''" by (auto dest: eta_par_beta) from s' obtain t'' where s: "s => t''" and t'': "t'' \\<^sub>\\<^sup>* t'" using step by blast from t'' and t' have "t'' \\<^sub>\\<^sup>* s'''" by (rule rtranclp_trans) with s show ?case by iprover qed theorem eta_postponement: assumes "(sup beta eta)\<^sup>*\<^sup>* s t" shows "(beta\<^sup>*\<^sup>* OO eta\<^sup>*\<^sup>*) s t" using assms proof induct case base show ?case by blast next case (step s' s'') from step(3) obtain t' where s: "s \\<^sub>\\<^sup>* t'" and t': "t' \\<^sub>\\<^sup>* s'" by blast from step(2) show ?case proof assume "s' \\<^sub>\ s''" with beta_subset_par_beta have "s' => s''" .. with t' obtain t'' where st: "t' => t''" and tu: "t'' \\<^sub>\\<^sup>* s''" by (auto dest: eta_postponement') from par_beta_subset_beta st have "t' \\<^sub>\\<^sup>* t''" .. with s have "s \\<^sub>\\<^sup>* t''" by (rule rtranclp_trans) thus ?thesis using tu .. next assume "s' \\<^sub>\ s''" with t' have "t' \\<^sub>\\<^sup>* s''" .. with s show ?thesis .. qed qed end diff --git a/src/HOL/Proofs/Lambda/Standardization.thy b/src/HOL/Proofs/Lambda/Standardization.thy --- a/src/HOL/Proofs/Lambda/Standardization.thy +++ b/src/HOL/Proofs/Lambda/Standardization.thy @@ -1,360 +1,360 @@ (* Title: HOL/Proofs/Lambda/Standardization.thy Author: Stefan Berghofer Copyright 2005 TU Muenchen *) section \Standardization\ theory Standardization imports NormalForm begin text \ -Based on lecture notes by Ralph Matthes @{cite "Matthes-ESSLLI2000"}, -original proof idea due to Ralph Loader @{cite Loader1998}. +Based on lecture notes by Ralph Matthes \<^cite>\"Matthes-ESSLLI2000"\, +original proof idea due to Ralph Loader \<^cite>\Loader1998\. \ subsection \Standard reduction relation\ declare listrel_mono [mono_set] inductive sred :: "dB \ dB \ bool" (infixl "\\<^sub>s" 50) and sredlist :: "dB list \ dB list \ bool" (infixl "[\\<^sub>s]" 50) where "s [\\<^sub>s] t \ listrelp (\\<^sub>s) s t" | Var: "rs [\\<^sub>s] rs' \ Var x \\ rs \\<^sub>s Var x \\ rs'" | Abs: "r \\<^sub>s r' \ ss [\\<^sub>s] ss' \ Abs r \\ ss \\<^sub>s Abs r' \\ ss'" | Beta: "r[s/0] \\ ss \\<^sub>s t \ Abs r \ s \\ ss \\<^sub>s t" lemma refl_listrelp: "\x\set xs. R x x \ listrelp R xs xs" by (induct xs) (auto intro: listrelp.intros) lemma refl_sred: "t \\<^sub>s t" by (induct t rule: Apps_dB_induct) (auto intro: refl_listrelp sred.intros) lemma refl_sreds: "ts [\\<^sub>s] ts" by (simp add: refl_sred refl_listrelp) lemma listrelp_conj1: "listrelp (\x y. R x y \ S x y) x y \ listrelp R x y" by (erule listrelp.induct) (auto intro: listrelp.intros) lemma listrelp_conj2: "listrelp (\x y. R x y \ S x y) x y \ listrelp S x y" by (erule listrelp.induct) (auto intro: listrelp.intros) lemma listrelp_app: assumes xsys: "listrelp R xs ys" shows "listrelp R xs' ys' \ listrelp R (xs @ xs') (ys @ ys')" using xsys by (induct arbitrary: xs' ys') (auto intro: listrelp.intros) lemma lemma1: assumes r: "r \\<^sub>s r'" and s: "s \\<^sub>s s'" shows "r \ s \\<^sub>s r' \ s'" using r proof induct case (Var rs rs' x) then have "rs [\\<^sub>s] rs'" by (rule listrelp_conj1) moreover have "[s] [\\<^sub>s] [s']" by (iprover intro: s listrelp.intros) ultimately have "rs @ [s] [\\<^sub>s] rs' @ [s']" by (rule listrelp_app) hence "Var x \\ (rs @ [s]) \\<^sub>s Var x \\ (rs' @ [s'])" by (rule sred.Var) thus ?case by (simp only: app_last) next case (Abs r r' ss ss') from Abs(3) have "ss [\\<^sub>s] ss'" by (rule listrelp_conj1) moreover have "[s] [\\<^sub>s] [s']" by (iprover intro: s listrelp.intros) ultimately have "ss @ [s] [\\<^sub>s] ss' @ [s']" by (rule listrelp_app) with \r \\<^sub>s r'\ have "Abs r \\ (ss @ [s]) \\<^sub>s Abs r' \\ (ss' @ [s'])" by (rule sred.Abs) thus ?case by (simp only: app_last) next case (Beta r u ss t) hence "r[u/0] \\ (ss @ [s]) \\<^sub>s t \ s'" by (simp only: app_last) hence "Abs r \ u \\ (ss @ [s]) \\<^sub>s t \ s'" by (rule sred.Beta) thus ?case by (simp only: app_last) qed lemma lemma1': assumes ts: "ts [\\<^sub>s] ts'" shows "r \\<^sub>s r' \ r \\ ts \\<^sub>s r' \\ ts'" using ts by (induct arbitrary: r r') (auto intro: lemma1) lemma lemma2_1: assumes beta: "t \\<^sub>\ u" shows "t \\<^sub>s u" using beta proof induct case (beta s t) have "Abs s \ t \\ [] \\<^sub>s s[t/0] \\ []" by (iprover intro: sred.Beta refl_sred) thus ?case by simp next case (appL s t u) thus ?case by (iprover intro: lemma1 refl_sred) next case (appR s t u) thus ?case by (iprover intro: lemma1 refl_sred) next case (abs s t) hence "Abs s \\ [] \\<^sub>s Abs t \\ []" by (iprover intro: sred.Abs listrelp.Nil) thus ?case by simp qed lemma listrelp_betas: assumes ts: "listrelp (\\<^sub>\\<^sup>*) ts ts'" shows "\t t'. t \\<^sub>\\<^sup>* t' \ t \\ ts \\<^sub>\\<^sup>* t' \\ ts'" using ts by induct auto lemma lemma2_2: assumes t: "t \\<^sub>s u" shows "t \\<^sub>\\<^sup>* u" using t by induct (auto dest: listrelp_conj2 intro: listrelp_betas apps_preserves_beta converse_rtranclp_into_rtranclp) lemma sred_lift: assumes s: "s \\<^sub>s t" shows "lift s i \\<^sub>s lift t i" using s proof (induct arbitrary: i) case (Var rs rs' x) hence "map (\t. lift t i) rs [\\<^sub>s] map (\t. lift t i) rs'" by induct (auto intro: listrelp.intros) thus ?case by (cases "x < i") (auto intro: sred.Var) next case (Abs r r' ss ss') from Abs(3) have "map (\t. lift t i) ss [\\<^sub>s] map (\t. lift t i) ss'" by induct (auto intro: listrelp.intros) thus ?case by (auto intro: sred.Abs Abs) next case (Beta r s ss t) thus ?case by (auto intro: sred.Beta) qed lemma lemma3: assumes r: "r \\<^sub>s r'" shows "s \\<^sub>s s' \ r[s/x] \\<^sub>s r'[s'/x]" using r proof (induct arbitrary: s s' x) case (Var rs rs' y) hence "map (\t. t[s/x]) rs [\\<^sub>s] map (\t. t[s'/x]) rs'" by induct (auto intro: listrelp.intros Var) moreover have "Var y[s/x] \\<^sub>s Var y[s'/x]" proof (cases "y < x") case True thus ?thesis by simp (rule refl_sred) next case False thus ?thesis by (cases "y = x") (auto simp add: Var intro: refl_sred) qed ultimately show ?case by simp (rule lemma1') next case (Abs r r' ss ss') from Abs(4) have "lift s 0 \\<^sub>s lift s' 0" by (rule sred_lift) hence "r[lift s 0/Suc x] \\<^sub>s r'[lift s' 0/Suc x]" by (fast intro: Abs.hyps) moreover from Abs(3) have "map (\t. t[s/x]) ss [\\<^sub>s] map (\t. t[s'/x]) ss'" by induct (auto intro: listrelp.intros Abs) ultimately show ?case by simp (rule sred.Abs) next case (Beta r u ss t) thus ?case by (auto simp add: subst_subst intro: sred.Beta) qed lemma lemma4_aux: assumes rs: "listrelp (\t u. t \\<^sub>s u \ (\r. u \\<^sub>\ r \ t \\<^sub>s r)) rs rs'" shows "rs' => ss \ rs [\\<^sub>s] ss" using rs proof (induct arbitrary: ss) case Nil thus ?case by cases (auto intro: listrelp.Nil) next case (Cons x y xs ys) note Cons' = Cons show ?case proof (cases ss) case Nil with Cons show ?thesis by simp next case (Cons y' ys') hence ss: "ss = y' # ys'" by simp from Cons Cons' have "y \\<^sub>\ y' \ ys' = ys \ y' = y \ ys => ys'" by simp hence "x # xs [\\<^sub>s] y' # ys'" proof assume H: "y \\<^sub>\ y' \ ys' = ys" with Cons' have "x \\<^sub>s y'" by blast moreover from Cons' have "xs [\\<^sub>s] ys" by (iprover dest: listrelp_conj1) ultimately have "x # xs [\\<^sub>s] y' # ys" by (rule listrelp.Cons) with H show ?thesis by simp next assume H: "y' = y \ ys => ys'" with Cons' have "x \\<^sub>s y'" by blast moreover from H have "xs [\\<^sub>s] ys'" by (blast intro: Cons') ultimately show ?thesis by (rule listrelp.Cons) qed with ss show ?thesis by simp qed qed lemma lemma4: assumes r: "r \\<^sub>s r'" shows "r' \\<^sub>\ r'' \ r \\<^sub>s r''" using r proof (induct arbitrary: r'') case (Var rs rs' x) then obtain ss where rs: "rs' => ss" and r'': "r'' = Var x \\ ss" by (blast dest: head_Var_reduction) from Var(1) rs have "rs [\\<^sub>s] ss" by (rule lemma4_aux) hence "Var x \\ rs \\<^sub>s Var x \\ ss" by (rule sred.Var) with r'' show ?case by simp next case (Abs r r' ss ss') from \Abs r' \\ ss' \\<^sub>\ r''\ show ?case proof fix s assume r'': "r'' = s \\ ss'" assume "Abs r' \\<^sub>\ s" then obtain r''' where s: "s = Abs r'''" and r''': "r' \\<^sub>\ r'''" by cases auto from r''' have "r \\<^sub>s r'''" by (blast intro: Abs) moreover from Abs have "ss [\\<^sub>s] ss'" by (iprover dest: listrelp_conj1) ultimately have "Abs r \\ ss \\<^sub>s Abs r''' \\ ss'" by (rule sred.Abs) with r'' s show "Abs r \\ ss \\<^sub>s r''" by simp next fix rs' assume "ss' => rs'" with Abs(3) have "ss [\\<^sub>s] rs'" by (rule lemma4_aux) with \r \\<^sub>s r'\ have "Abs r \\ ss \\<^sub>s Abs r' \\ rs'" by (rule sred.Abs) moreover assume "r'' = Abs r' \\ rs'" ultimately show "Abs r \\ ss \\<^sub>s r''" by simp next fix t u' us' assume "ss' = u' # us'" with Abs(3) obtain u us where ss: "ss = u # us" and u: "u \\<^sub>s u'" and us: "us [\\<^sub>s] us'" by cases (auto dest!: listrelp_conj1) have "r[u/0] \\<^sub>s r'[u'/0]" using Abs(1) and u by (rule lemma3) with us have "r[u/0] \\ us \\<^sub>s r'[u'/0] \\ us'" by (rule lemma1') hence "Abs r \ u \\ us \\<^sub>s r'[u'/0] \\ us'" by (rule sred.Beta) moreover assume "Abs r' = Abs t" and "r'' = t[u'/0] \\ us'" ultimately show "Abs r \\ ss \\<^sub>s r''" using ss by simp qed next case (Beta r s ss t) show ?case by (rule sred.Beta) (rule Beta)+ qed lemma rtrancl_beta_sred: assumes r: "r \\<^sub>\\<^sup>* r'" shows "r \\<^sub>s r'" using r by induct (iprover intro: refl_sred lemma4)+ subsection \Leftmost reduction and weakly normalizing terms\ inductive lred :: "dB \ dB \ bool" (infixl "\\<^sub>l" 50) and lredlist :: "dB list \ dB list \ bool" (infixl "[\\<^sub>l]" 50) where "s [\\<^sub>l] t \ listrelp (\\<^sub>l) s t" | Var: "rs [\\<^sub>l] rs' \ Var x \\ rs \\<^sub>l Var x \\ rs'" | Abs: "r \\<^sub>l r' \ Abs r \\<^sub>l Abs r'" | Beta: "r[s/0] \\ ss \\<^sub>l t \ Abs r \ s \\ ss \\<^sub>l t" lemma lred_imp_sred: assumes lred: "s \\<^sub>l t" shows "s \\<^sub>s t" using lred proof induct case (Var rs rs' x) then have "rs [\\<^sub>s] rs'" by induct (iprover intro: listrelp.intros)+ then show ?case by (rule sred.Var) next case (Abs r r') from \r \\<^sub>s r'\ have "Abs r \\ [] \\<^sub>s Abs r' \\ []" using listrelp.Nil by (rule sred.Abs) then show ?case by simp next case (Beta r s ss t) from \r[s/0] \\ ss \\<^sub>s t\ show ?case by (rule sred.Beta) qed inductive WN :: "dB => bool" where Var: "listsp WN rs \ WN (Var n \\ rs)" | Lambda: "WN r \ WN (Abs r)" | Beta: "WN ((r[s/0]) \\ ss) \ WN ((Abs r \ s) \\ ss)" lemma listrelp_imp_listsp1: assumes H: "listrelp (\x y. P x) xs ys" shows "listsp P xs" using H by induct auto lemma listrelp_imp_listsp2: assumes H: "listrelp (\x y. P y) xs ys" shows "listsp P ys" using H by induct auto lemma lemma5: assumes lred: "r \\<^sub>l r'" shows "WN r" and "NF r'" using lred by induct (iprover dest: listrelp_conj1 listrelp_conj2 listrelp_imp_listsp1 listrelp_imp_listsp2 intro: WN.intros NF.intros [simplified listall_listsp_eq])+ lemma lemma6: assumes wn: "WN r" shows "\r'. r \\<^sub>l r'" using wn proof induct case (Var rs n) then have "\rs'. rs [\\<^sub>l] rs'" by induct (iprover intro: listrelp.intros)+ then show ?case by (iprover intro: lred.Var) qed (iprover intro: lred.intros)+ lemma lemma7: assumes r: "r \\<^sub>s r'" shows "NF r' \ r \\<^sub>l r'" using r proof induct case (Var rs rs' x) from \NF (Var x \\ rs')\ have "listall NF rs'" by cases simp_all with Var(1) have "rs [\\<^sub>l] rs'" proof induct case Nil show ?case by (rule listrelp.Nil) next case (Cons x y xs ys) hence "x \\<^sub>l y" and "xs [\\<^sub>l] ys" by simp_all thus ?case by (rule listrelp.Cons) qed thus ?case by (rule lred.Var) next case (Abs r r' ss ss') from \NF (Abs r' \\ ss')\ have ss': "ss' = []" by (rule Abs_NF) from Abs(3) have ss: "ss = []" using ss' by cases simp_all from ss' Abs have "NF (Abs r')" by simp hence "NF r'" by cases simp_all with Abs have "r \\<^sub>l r'" by simp hence "Abs r \\<^sub>l Abs r'" by (rule lred.Abs) with ss ss' show ?case by simp next case (Beta r s ss t) hence "r[s/0] \\ ss \\<^sub>l t" by simp thus ?case by (rule lred.Beta) qed lemma WN_eq: "WN t = (\t'. t \\<^sub>\\<^sup>* t' \ NF t')" proof assume "WN t" then have "\t'. t \\<^sub>l t'" by (rule lemma6) then obtain t' where t': "t \\<^sub>l t'" .. then have NF: "NF t'" by (rule lemma5) from t' have "t \\<^sub>s t'" by (rule lred_imp_sred) then have "t \\<^sub>\\<^sup>* t'" by (rule lemma2_2) with NF show "\t'. t \\<^sub>\\<^sup>* t' \ NF t'" by iprover next assume "\t'. t \\<^sub>\\<^sup>* t' \ NF t'" then obtain t' where t': "t \\<^sub>\\<^sup>* t'" and NF: "NF t'" by iprover from t' have "t \\<^sub>s t'" by (rule rtrancl_beta_sred) then have "t \\<^sub>l t'" using NF by (rule lemma7) then show "WN t" by (rule lemma5) qed end diff --git a/src/HOL/Proofs/Lambda/StrongNorm.thy b/src/HOL/Proofs/Lambda/StrongNorm.thy --- a/src/HOL/Proofs/Lambda/StrongNorm.thy +++ b/src/HOL/Proofs/Lambda/StrongNorm.thy @@ -1,284 +1,284 @@ (* Title: HOL/Proofs/Lambda/StrongNorm.thy Author: Stefan Berghofer Copyright 2000 TU Muenchen *) section \Strong normalization for simply-typed lambda calculus\ theory StrongNorm imports LambdaType InductTermi begin text \ Formalization by Stefan Berghofer. Partly based on a paper proof by -Felix Joachimski and Ralph Matthes @{cite "Matthes-Joachimski-AML"}. +Felix Joachimski and Ralph Matthes \<^cite>\"Matthes-Joachimski-AML"\. \ subsection \Properties of \IT\\ lemma lift_IT [intro!]: "IT t \ IT (lift t i)" apply (induct arbitrary: i set: IT) apply (simp (no_asm)) apply (rule conjI) apply (rule impI, rule IT.Var, erule listsp.induct, simp (no_asm), simp (no_asm), rule listsp.Cons, blast, assumption)+ apply auto done lemma lifts_IT: "listsp IT ts \ listsp IT (map (\t. lift t 0) ts)" by (induct ts) auto lemma subst_Var_IT: "IT r \ IT (r[Var i/j])" apply (induct arbitrary: i j set: IT) txt \Case \<^term>\Var\:\ apply (simp (no_asm) add: subst_Var) apply ((rule conjI impI)+, rule IT.Var, erule listsp.induct, simp (no_asm), simp (no_asm), rule listsp.Cons, fast, assumption)+ txt \Case \<^term>\Lambda\:\ apply atomize apply simp apply (rule IT.Lambda) apply fast txt \Case \<^term>\Beta\:\ apply atomize apply (simp (no_asm_use) add: subst_subst [symmetric]) apply (rule IT.Beta) apply auto done lemma Var_IT: "IT (Var n)" apply (subgoal_tac "IT (Var n \\ [])") apply simp apply (rule IT.Var) apply (rule listsp.Nil) done lemma app_Var_IT: "IT t \ IT (t \ Var i)" apply (induct set: IT) apply (subst app_last) apply (rule IT.Var) apply simp apply (rule listsp.Cons) apply (rule Var_IT) apply (rule listsp.Nil) apply (rule IT.Beta [where ?ss = "[]", unfolded foldl_Nil [THEN eq_reflection]]) apply (erule subst_Var_IT) apply (rule Var_IT) apply (subst app_last) apply (rule IT.Beta) apply (subst app_last [symmetric]) apply assumption apply assumption done subsection \Well-typed substitution preserves termination\ lemma subst_type_IT: "\t e T u i. IT t \ e\i:U\ \ t : T \ IT u \ e \ u : U \ IT (t[u/i])" (is "PROP ?P U" is "\t e T u i. _ \ PROP ?Q t e T u i U") proof (induct U) fix T t assume MI1: "\T1 T2. T = T1 \ T2 \ PROP ?P T1" assume MI2: "\T1 T2. T = T1 \ T2 \ PROP ?P T2" assume "IT t" thus "\e T' u i. PROP ?Q t e T' u i T" proof induct fix e T' u i assume uIT: "IT u" assume uT: "e \ u : T" { case (Var rs n e1 T'1 u1 i1) assume nT: "e\i:T\ \ Var n \\ rs : T'" let ?ty = "\t. \T'. e\i:T\ \ t : T'" let ?R = "\t. \e T' u i. e\i:T\ \ t : T' \ IT u \ e \ u : T \ IT (t[u/i])" show "IT ((Var n \\ rs)[u/i])" proof (cases "n = i") case True show ?thesis proof (cases rs) case Nil with uIT True show ?thesis by simp next case (Cons a as) with nT have "e\i:T\ \ Var n \ a \\ as : T'" by simp then obtain Ts where headT: "e\i:T\ \ Var n \ a : Ts \ T'" and argsT: "e\i:T\ \ as : Ts" by (rule list_app_typeE) from headT obtain T'' where varT: "e\i:T\ \ Var n : T'' \ Ts \ T'" and argT: "e\i:T\ \ a : T''" by cases simp_all from varT True have T: "T = T'' \ Ts \ T'" by cases auto with uT have uT': "e \ u : T'' \ Ts \ T'" by simp from T have "IT ((Var 0 \\ map (\t. lift t 0) (map (\t. t[u/i]) as))[(u \ a[u/i])/0])" proof (rule MI2) from T have "IT ((lift u 0 \ Var 0)[a[u/i]/0])" proof (rule MI1) have "IT (lift u 0)" by (rule lift_IT [OF uIT]) thus "IT (lift u 0 \ Var 0)" by (rule app_Var_IT) show "e\0:T''\ \ lift u 0 \ Var 0 : Ts \ T'" proof (rule typing.App) show "e\0:T''\ \ lift u 0 : T'' \ Ts \ T'" by (rule lift_type) (rule uT') show "e\0:T''\ \ Var 0 : T''" by (rule typing.Var) simp qed from Var have "?R a" by cases (simp_all add: Cons) with argT uIT uT show "IT (a[u/i])" by simp from argT uT show "e \ a[u/i] : T''" by (rule subst_lemma) simp qed thus "IT (u \ a[u/i])" by simp from Var have "listsp ?R as" by cases (simp_all add: Cons) moreover from argsT have "listsp ?ty as" by (rule lists_typings) ultimately have "listsp (\t. ?R t \ ?ty t) as" by simp hence "listsp IT (map (\t. lift t 0) (map (\t. t[u/i]) as))" (is "listsp IT (?ls as)") proof induct case Nil show ?case by fastforce next case (Cons b bs) hence I: "?R b" by simp from Cons obtain U where "e\i:T\ \ b : U" by fast with uT uIT I have "IT (b[u/i])" by simp hence "IT (lift (b[u/i]) 0)" by (rule lift_IT) hence "listsp IT (lift (b[u/i]) 0 # ?ls bs)" by (rule listsp.Cons) (rule Cons) thus ?case by simp qed thus "IT (Var 0 \\ ?ls as)" by (rule IT.Var) have "e\0:Ts \ T'\ \ Var 0 : Ts \ T'" by (rule typing.Var) simp moreover from uT argsT have "e \ map (\t. t[u/i]) as : Ts" by (rule substs_lemma) hence "e\0:Ts \ T'\ \ ?ls as : Ts" by (rule lift_types) ultimately show "e\0:Ts \ T'\ \ Var 0 \\ ?ls as : T'" by (rule list_app_typeI) from argT uT have "e \ a[u/i] : T''" by (rule subst_lemma) (rule refl) with uT' show "e \ u \ a[u/i] : Ts \ T'" by (rule typing.App) qed with Cons True show ?thesis by (simp add: comp_def) qed next case False from Var have "listsp ?R rs" by simp moreover from nT obtain Ts where "e\i:T\ \ rs : Ts" by (rule list_app_typeE) hence "listsp ?ty rs" by (rule lists_typings) ultimately have "listsp (\t. ?R t \ ?ty t) rs" by simp hence "listsp IT (map (\x. x[u/i]) rs)" proof induct case Nil show ?case by fastforce next case (Cons a as) hence I: "?R a" by simp from Cons obtain U where "e\i:T\ \ a : U" by fast with uT uIT I have "IT (a[u/i])" by simp hence "listsp IT (a[u/i] # map (\t. t[u/i]) as)" by (rule listsp.Cons) (rule Cons) thus ?case by simp qed with False show ?thesis by (auto simp add: subst_Var) qed next case (Lambda r e1 T'1 u1 i1) assume "e\i:T\ \ Abs r : T'" and "\e T' u i. PROP ?Q r e T' u i T" with uIT uT show "IT (Abs r[u/i])" by fastforce next case (Beta r a as e1 T'1 u1 i1) assume T: "e\i:T\ \ Abs r \ a \\ as : T'" assume SI1: "\e T' u i. PROP ?Q (r[a/0] \\ as) e T' u i T" assume SI2: "\e T' u i. PROP ?Q a e T' u i T" have "IT (Abs (r[lift u 0/Suc i]) \ a[u/i] \\ map (\t. t[u/i]) as)" proof (rule IT.Beta) have "Abs r \ a \\ as \\<^sub>\ r[a/0] \\ as" by (rule apps_preserves_beta) (rule beta.beta) with T have "e\i:T\ \ r[a/0] \\ as : T'" by (rule subject_reduction) hence "IT ((r[a/0] \\ as)[u/i])" using uIT uT by (rule SI1) thus "IT (r[lift u 0/Suc i][a[u/i]/0] \\ map (\t. t[u/i]) as)" by (simp del: subst_map add: subst_subst subst_map [symmetric]) from T obtain U where "e\i:T\ \ Abs r \ a : U" by (rule list_app_typeE) fast then obtain T'' where "e\i:T\ \ a : T''" by cases simp_all thus "IT (a[u/i])" using uIT uT by (rule SI2) qed thus "IT ((Abs r \ a \\ as)[u/i])" by simp } qed qed subsection \Well-typed terms are strongly normalizing\ lemma type_implies_IT: assumes "e \ t : T" shows "IT t" using assms proof induct case Var show ?case by (rule Var_IT) next case Abs show ?case by (rule IT.Lambda) (rule Abs) next case (App e s T U t) have "IT ((Var 0 \ lift t 0)[s/0])" proof (rule subst_type_IT) have "IT (lift t 0)" using \IT t\ by (rule lift_IT) hence "listsp IT [lift t 0]" by (rule listsp.Cons) (rule listsp.Nil) hence "IT (Var 0 \\ [lift t 0])" by (rule IT.Var) also have "Var 0 \\ [lift t 0] = Var 0 \ lift t 0" by simp finally show "IT \" . have "e\0:T \ U\ \ Var 0 : T \ U" by (rule typing.Var) simp moreover have "e\0:T \ U\ \ lift t 0 : T" by (rule lift_type) (rule App.hyps) ultimately show "e\0:T \ U\ \ Var 0 \ lift t 0 : U" by (rule typing.App) show "IT s" by fact show "e \ s : T \ U" by fact qed thus ?case by simp qed theorem type_implies_termi: "e \ t : T \ termip beta t" proof - assume "e \ t : T" hence "IT t" by (rule type_implies_IT) thus ?thesis by (rule IT_implies_termi) qed end diff --git a/src/HOL/Proofs/Lambda/WeakNorm.thy b/src/HOL/Proofs/Lambda/WeakNorm.thy --- a/src/HOL/Proofs/Lambda/WeakNorm.thy +++ b/src/HOL/Proofs/Lambda/WeakNorm.thy @@ -1,442 +1,442 @@ (* Title: HOL/Proofs/Lambda/WeakNorm.thy Author: Stefan Berghofer Copyright 2003 TU Muenchen *) section \Weak normalization for simply-typed lambda calculus\ theory WeakNorm imports LambdaType NormalForm "HOL-Library.Realizers" "HOL-Library.Code_Target_Int" begin text \ Formalization by Stefan Berghofer. Partly based on a paper proof by -Felix Joachimski and Ralph Matthes @{cite "Matthes-Joachimski-AML"}. +Felix Joachimski and Ralph Matthes \<^cite>\"Matthes-Joachimski-AML"\. \ subsection \Main theorems\ lemma norm_list: assumes f_compat: "\t t'. t \\<^sub>\\<^sup>* t' \ f t \\<^sub>\\<^sup>* f t'" and f_NF: "\t. NF t \ NF (f t)" and uNF: "NF u" and uT: "e \ u : T" shows "\Us. e\i:T\ \ as : Us \ listall (\t. \e T' u i. e\i:T\ \ t : T' \ NF u \ e \ u : T \ (\t'. t[u/i] \\<^sub>\\<^sup>* t' \ NF t')) as \ \as'. \j. Var j \\ map (\t. f (t[u/i])) as \\<^sub>\\<^sup>* Var j \\ map f as' \ NF (Var j \\ map f as')" (is "\Us. _ \ listall ?R as \ \as'. ?ex Us as as'") proof (induct as rule: rev_induct) case (Nil Us) with Var_NF have "?ex Us [] []" by simp thus ?case .. next case (snoc b bs Us) have "e\i:T\ \ bs @ [b] : Us" by fact then obtain Vs W where Us: "Us = Vs @ [W]" and bs: "e\i:T\ \ bs : Vs" and bT: "e\i:T\ \ b : W" by (rule types_snocE) from snoc have "listall ?R bs" by simp with bs have "\bs'. ?ex Vs bs bs'" by (rule snoc) then obtain bs' where bsred: "Var j \\ map (\t. f (t[u/i])) bs \\<^sub>\\<^sup>* Var j \\ map f bs'" and bsNF: "NF (Var j \\ map f bs')" for j by iprover from snoc have "?R b" by simp with bT and uNF and uT have "\b'. b[u/i] \\<^sub>\\<^sup>* b' \ NF b'" by iprover then obtain b' where bred: "b[u/i] \\<^sub>\\<^sup>* b'" and bNF: "NF b'" by iprover from bsNF [of 0] have "listall NF (map f bs')" by (rule App_NF_D) moreover have "NF (f b')" using bNF by (rule f_NF) ultimately have "listall NF (map f (bs' @ [b']))" by simp hence "\j. NF (Var j \\ map f (bs' @ [b']))" by (rule NF.App) moreover from bred have "f (b[u/i]) \\<^sub>\\<^sup>* f b'" by (rule f_compat) with bsred have "\j. (Var j \\ map (\t. f (t[u/i])) bs) \ f (b[u/i]) \\<^sub>\\<^sup>* (Var j \\ map f bs') \ f b'" by (rule rtrancl_beta_App) ultimately have "?ex Us (bs @ [b]) (bs' @ [b'])" by simp thus ?case .. qed lemma subst_type_NF: "\t e T u i. NF t \ e\i:U\ \ t : T \ NF u \ e \ u : U \ \t'. t[u/i] \\<^sub>\\<^sup>* t' \ NF t'" (is "PROP ?P U" is "\t e T u i. _ \ PROP ?Q t e T u i U") proof (induct U) fix T t let ?R = "\t. \e T' u i. e\i:T\ \ t : T' \ NF u \ e \ u : T \ (\t'. t[u/i] \\<^sub>\\<^sup>* t' \ NF t')" assume MI1: "\T1 T2. T = T1 \ T2 \ PROP ?P T1" assume MI2: "\T1 T2. T = T1 \ T2 \ PROP ?P T2" assume "NF t" thus "\e T' u i. PROP ?Q t e T' u i T" proof induct fix e T' u i assume uNF: "NF u" and uT: "e \ u : T" { case (App ts x e1 T'1 u1 i1) assume "e\i:T\ \ Var x \\ ts : T'" then obtain Us where varT: "e\i:T\ \ Var x : Us \ T'" and argsT: "e\i:T\ \ ts : Us" by (rule var_app_typesE) from nat_eq_dec show "\t'. (Var x \\ ts)[u/i] \\<^sub>\\<^sup>* t' \ NF t'" proof assume eq: "x = i" show ?thesis proof (cases ts) case Nil with eq have "(Var x \\ [])[u/i] \\<^sub>\\<^sup>* u" by simp with Nil and uNF show ?thesis by simp iprover next case (Cons a as) with argsT obtain T'' Ts where Us: "Us = T'' # Ts" by (cases Us) (rule FalseE, simp) from varT and Us have varT: "e\i:T\ \ Var x : T'' \ Ts \ T'" by simp from varT eq have T: "T = T'' \ Ts \ T'" by cases auto with uT have uT': "e \ u : T'' \ Ts \ T'" by simp from argsT Us Cons have argsT': "e\i:T\ \ as : Ts" by simp from argsT Us Cons have argT: "e\i:T\ \ a : T''" by simp from argT uT refl have aT: "e \ a[u/i] : T''" by (rule subst_lemma) from App and Cons have "listall ?R as" by simp (iprover dest: listall_conj2) with lift_preserves_beta' lift_NF uNF uT argsT' have "\as'. \j. Var j \\ map (\t. lift (t[u/i]) 0) as \\<^sub>\\<^sup>* Var j \\ map (\t. lift t 0) as' \ NF (Var j \\ map (\t. lift t 0) as')" by (rule norm_list) then obtain as' where asred: "Var 0 \\ map (\t. lift (t[u/i]) 0) as \\<^sub>\\<^sup>* Var 0 \\ map (\t. lift t 0) as'" and asNF: "NF (Var 0 \\ map (\t. lift t 0) as')" by iprover from App and Cons have "?R a" by simp with argT and uNF and uT have "\a'. a[u/i] \\<^sub>\\<^sup>* a' \ NF a'" by iprover then obtain a' where ared: "a[u/i] \\<^sub>\\<^sup>* a'" and aNF: "NF a'" by iprover from uNF have "NF (lift u 0)" by (rule lift_NF) hence "\u'. lift u 0 \ Var 0 \\<^sub>\\<^sup>* u' \ NF u'" by (rule app_Var_NF) then obtain u' where ured: "lift u 0 \ Var 0 \\<^sub>\\<^sup>* u'" and u'NF: "NF u'" by iprover from T and u'NF have "\ua. u'[a'/0] \\<^sub>\\<^sup>* ua \ NF ua" proof (rule MI1) have "e\0:T''\ \ lift u 0 \ Var 0 : Ts \ T'" proof (rule typing.App) from uT' show "e\0:T''\ \ lift u 0 : T'' \ Ts \ T'" by (rule lift_type) show "e\0:T''\ \ Var 0 : T''" by (rule typing.Var) simp qed with ured show "e\0:T''\ \ u' : Ts \ T'" by (rule subject_reduction') from ared aT show "e \ a' : T''" by (rule subject_reduction') show "NF a'" by fact qed then obtain ua where uared: "u'[a'/0] \\<^sub>\\<^sup>* ua" and uaNF: "NF ua" by iprover from ared have "(lift u 0 \ Var 0)[a[u/i]/0] \\<^sub>\\<^sup>* (lift u 0 \ Var 0)[a'/0]" by (rule subst_preserves_beta2') also from ured have "(lift u 0 \ Var 0)[a'/0] \\<^sub>\\<^sup>* u'[a'/0]" by (rule subst_preserves_beta') also note uared finally have "(lift u 0 \ Var 0)[a[u/i]/0] \\<^sub>\\<^sup>* ua" . hence uared': "u \ a[u/i] \\<^sub>\\<^sup>* ua" by simp from T asNF _ uaNF have "\r. (Var 0 \\ map (\t. lift t 0) as')[ua/0] \\<^sub>\\<^sup>* r \ NF r" proof (rule MI2) have "e\0:Ts \ T'\ \ Var 0 \\ map (\t. lift (t[u/i]) 0) as : T'" proof (rule list_app_typeI) show "e\0:Ts \ T'\ \ Var 0 : Ts \ T'" by (rule typing.Var) simp from uT argsT' have "e \ map (\t. t[u/i]) as : Ts" by (rule substs_lemma) hence "e\0:Ts \ T'\ \ map (\t. lift t 0) (map (\t. t[u/i]) as) : Ts" by (rule lift_types) thus "e\0:Ts \ T'\ \ map (\t. lift (t[u/i]) 0) as : Ts" by (simp_all add: o_def) qed with asred show "e\0:Ts \ T'\ \ Var 0 \\ map (\t. lift t 0) as' : T'" by (rule subject_reduction') from argT uT refl have "e \ a[u/i] : T''" by (rule subst_lemma) with uT' have "e \ u \ a[u/i] : Ts \ T'" by (rule typing.App) with uared' show "e \ ua : Ts \ T'" by (rule subject_reduction') qed then obtain r where rred: "(Var 0 \\ map (\t. lift t 0) as')[ua/0] \\<^sub>\\<^sup>* r" and rnf: "NF r" by iprover from asred have "(Var 0 \\ map (\t. lift (t[u/i]) 0) as)[u \ a[u/i]/0] \\<^sub>\\<^sup>* (Var 0 \\ map (\t. lift t 0) as')[u \ a[u/i]/0]" by (rule subst_preserves_beta') also from uared' have "(Var 0 \\ map (\t. lift t 0) as')[u \ a[u/i]/0] \\<^sub>\\<^sup>* (Var 0 \\ map (\t. lift t 0) as')[ua/0]" by (rule subst_preserves_beta2') also note rred finally have "(Var 0 \\ map (\t. lift (t[u/i]) 0) as)[u \ a[u/i]/0] \\<^sub>\\<^sup>* r" . with rnf Cons eq show ?thesis by (simp add: o_def) iprover qed next assume neq: "x \ i" from App have "listall ?R ts" by (iprover dest: listall_conj2) with uNF uT argsT have "\ts'. \j. Var j \\ map (\t. t[u/i]) ts \\<^sub>\\<^sup>* Var j \\ ts' \ NF (Var j \\ ts')" (is "\ts'. ?ex ts'") by (rule norm_list [of "\t. t", simplified]) then obtain ts' where NF: "?ex ts'" .. from nat_le_dec show ?thesis proof assume "i < x" with NF show ?thesis by simp iprover next assume "\ (i < x)" with NF neq show ?thesis by (simp add: subst_Var) iprover qed qed next case (Abs r e1 T'1 u1 i1) assume absT: "e\i:T\ \ Abs r : T'" then obtain R S where "e\0:R\\Suc i:T\ \ r : S" by (rule abs_typeE) simp moreover have "NF (lift u 0)" using \NF u\ by (rule lift_NF) moreover have "e\0:R\ \ lift u 0 : T" using uT by (rule lift_type) ultimately have "\t'. r[lift u 0/Suc i] \\<^sub>\\<^sup>* t' \ NF t'" by (rule Abs) thus "\t'. Abs r[u/i] \\<^sub>\\<^sup>* t' \ NF t'" by simp (iprover intro: rtrancl_beta_Abs NF.Abs) } qed qed \ \A computationally relevant copy of @{term "e \ t : T"}\ inductive rtyping :: "(nat \ type) \ dB \ type \ bool" ("_ \\<^sub>R _ : _" [50, 50, 50] 50) where Var: "e x = T \ e \\<^sub>R Var x : T" | Abs: "e\0:T\ \\<^sub>R t : U \ e \\<^sub>R Abs t : (T \ U)" | App: "e \\<^sub>R s : T \ U \ e \\<^sub>R t : T \ e \\<^sub>R (s \ t) : U" lemma rtyping_imp_typing: "e \\<^sub>R t : T \ e \ t : T" apply (induct set: rtyping) apply (erule typing.Var) apply (erule typing.Abs) apply (erule typing.App) apply assumption done theorem type_NF: assumes "e \\<^sub>R t : T" shows "\t'. t \\<^sub>\\<^sup>* t' \ NF t'" using assms proof induct case Var show ?case by (iprover intro: Var_NF) next case Abs thus ?case by (iprover intro: rtrancl_beta_Abs NF.Abs) next case (App e s T U t) from App obtain s' t' where sred: "s \\<^sub>\\<^sup>* s'" and "NF s'" and tred: "t \\<^sub>\\<^sup>* t'" and tNF: "NF t'" by iprover have "\u. (Var 0 \ lift t' 0)[s'/0] \\<^sub>\\<^sup>* u \ NF u" proof (rule subst_type_NF) have "NF (lift t' 0)" using tNF by (rule lift_NF) hence "listall NF [lift t' 0]" by (rule listall_cons) (rule listall_nil) hence "NF (Var 0 \\ [lift t' 0])" by (rule NF.App) thus "NF (Var 0 \ lift t' 0)" by simp show "e\0:T \ U\ \ Var 0 \ lift t' 0 : U" proof (rule typing.App) show "e\0:T \ U\ \ Var 0 : T \ U" by (rule typing.Var) simp from tred have "e \ t' : T" by (rule subject_reduction') (rule rtyping_imp_typing, rule App.hyps) thus "e\0:T \ U\ \ lift t' 0 : T" by (rule lift_type) qed from sred show "e \ s' : T \ U" by (rule subject_reduction') (rule rtyping_imp_typing, rule App.hyps) show "NF s'" by fact qed then obtain u where ured: "s' \ t' \\<^sub>\\<^sup>* u" and unf: "NF u" by simp iprover from sred tred have "s \ t \\<^sub>\\<^sup>* s' \ t'" by (rule rtrancl_beta_App) hence "s \ t \\<^sub>\\<^sup>* u" using ured by (rule rtranclp_trans) with unf show ?case by iprover qed subsection \Extracting the program\ declare NF.induct [ind_realizer] declare rtranclp.induct [ind_realizer irrelevant] declare rtyping.induct [ind_realizer] lemmas [extraction_expand] = conj_assoc listall_cons_eq subst_all equal_allI extract type_NF lemma rtranclR_rtrancl_eq: "rtranclpR r a b = r\<^sup>*\<^sup>* a b" apply (rule iffI) apply (erule rtranclpR.induct) apply (rule rtranclp.rtrancl_refl) apply (erule rtranclp.rtrancl_into_rtrancl) apply assumption apply (erule rtranclp.induct) apply (rule rtranclpR.rtrancl_refl) apply (erule rtranclpR.rtrancl_into_rtrancl) apply assumption done lemma NFR_imp_NF: "NFR nf t \ NF t" apply (erule NFR.induct) apply (rule NF.intros) apply (simp add: listall_def) apply (erule NF.intros) done text_raw \ \begin{figure} \renewcommand{\isastyle}{\scriptsize\it}% @{thm [display,eta_contract=false,margin=100] subst_type_NF_def} \renewcommand{\isastyle}{\small\it}% \caption{Program extracted from \subst_type_NF\} \label{fig:extr-subst-type-nf} \end{figure} \begin{figure} \renewcommand{\isastyle}{\scriptsize\it}% @{thm [display,margin=100] subst_Var_NF_def} @{thm [display,margin=100] app_Var_NF_def} @{thm [display,margin=100] lift_NF_def} @{thm [display,eta_contract=false,margin=100] type_NF_def} \renewcommand{\isastyle}{\small\it}% \caption{Program extracted from lemmas and main theorem} \label{fig:extr-type-nf} \end{figure} \ text \ The program corresponding to the proof of the central lemma, which performs substitution and normalization, is shown in Figure \ref{fig:extr-subst-type-nf}. The correctness theorem corresponding to the program \subst_type_NF\ is @{thm [display,margin=100] subst_type_NF_correctness [simplified rtranclR_rtrancl_eq Collect_mem_eq, no_vars]} where \NFR\ is the realizability predicate corresponding to the datatype \NFT\, which is inductively defined by the rules \pagebreak @{thm [display,margin=90] NFR.App [of ts nfs x] NFR.Abs [of nf t]} The programs corresponding to the main theorem \type_NF\, as well as to some lemmas, are shown in Figure \ref{fig:extr-type-nf}. The correctness statement for the main function \type_NF\ is @{thm [display,margin=100] type_NF_correctness [simplified rtranclR_rtrancl_eq Collect_mem_eq, no_vars]} where the realizability predicate \rtypingR\ corresponding to the computationally relevant version of the typing judgement is inductively defined by the rules @{thm [display,margin=100] rtypingR.Var [no_vars] rtypingR.Abs [of ty, no_vars] rtypingR.App [of ty e s T U ty' t]} \ subsection \Generating executable code\ instantiation NFT :: default begin definition "default = Dummy ()" instance .. end instantiation dB :: default begin definition "default = dB.Var 0" instance .. end instantiation prod :: (default, default) default begin definition "default = (default, default)" instance .. end instantiation list :: (type) default begin definition "default = []" instance .. end instantiation "fun" :: (type, default) default begin definition "default = (\x. default)" instance .. end definition int_of_nat :: "nat \ int" where "int_of_nat = of_nat" text \ The following functions convert between Isabelle's built-in {\tt term} datatype and the generated {\tt dB} datatype. This allows to generate example terms using Isabelle's parser and inspect normalized terms using Isabelle's pretty printer. \ ML \ val nat_of_integer = @{code nat} o @{code int_of_integer}; fun dBtype_of_typ (Type ("fun", [T, U])) = @{code Fun} (dBtype_of_typ T, dBtype_of_typ U) | dBtype_of_typ (TFree (s, _)) = (case raw_explode s of ["'", a] => @{code Atom} (nat_of_integer (ord a - 97)) | _ => error "dBtype_of_typ: variable name") | dBtype_of_typ _ = error "dBtype_of_typ: bad type"; fun dB_of_term (Bound i) = @{code dB.Var} (nat_of_integer i) | dB_of_term (t $ u) = @{code dB.App} (dB_of_term t, dB_of_term u) | dB_of_term (Abs (_, _, t)) = @{code dB.Abs} (dB_of_term t) | dB_of_term _ = error "dB_of_term: bad term"; fun term_of_dB Ts (Type ("fun", [T, U])) (@{code dB.Abs} dBt) = Abs ("x", T, term_of_dB (T :: Ts) U dBt) | term_of_dB Ts _ dBt = term_of_dB' Ts dBt and term_of_dB' Ts (@{code dB.Var} n) = Bound (@{code integer_of_nat} n) | term_of_dB' Ts (@{code dB.App} (dBt, dBu)) = let val t = term_of_dB' Ts dBt in case fastype_of1 (Ts, t) of Type ("fun", [T, _]) => t $ term_of_dB Ts T dBu | _ => error "term_of_dB: function type expected" end | term_of_dB' _ _ = error "term_of_dB: term not in normal form"; fun typing_of_term Ts e (Bound i) = @{code Var} (e, nat_of_integer i, dBtype_of_typ (nth Ts i)) | typing_of_term Ts e (t $ u) = (case fastype_of1 (Ts, t) of Type ("fun", [T, U]) => @{code App} (e, dB_of_term t, dBtype_of_typ T, dBtype_of_typ U, dB_of_term u, typing_of_term Ts e t, typing_of_term Ts e u) | _ => error "typing_of_term: function type expected") | typing_of_term Ts e (Abs (_, T, t)) = let val dBT = dBtype_of_typ T in @{code Abs} (e, dBT, dB_of_term t, dBtype_of_typ (fastype_of1 (T :: Ts, t)), typing_of_term (T :: Ts) (@{code shift} e @{code "0::nat"} dBT) t) end | typing_of_term _ _ _ = error "typing_of_term: bad term"; fun dummyf _ = error "dummy"; val ct1 = @{cterm "%f. ((%f x. f (f (f x))) ((%f x. f (f (f (f x)))) f))"}; val (dB1, _) = @{code type_NF} (typing_of_term [] dummyf (Thm.term_of ct1)); val ct1' = Thm.cterm_of @{context} (term_of_dB [] (Thm.typ_of_cterm ct1) dB1); val ct2 = @{cterm "%f x. (%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) x)))))"}; val (dB2, _) = @{code type_NF} (typing_of_term [] dummyf (Thm.term_of ct2)); val ct2' = Thm.cterm_of @{context} (term_of_dB [] (Thm.typ_of_cterm ct2) dB2); \ end diff --git a/src/HOL/SPARK/Manual/Example_Verification.thy b/src/HOL/SPARK/Manual/Example_Verification.thy --- a/src/HOL/SPARK/Manual/Example_Verification.thy +++ b/src/HOL/SPARK/Manual/Example_Verification.thy @@ -1,250 +1,250 @@ (*<*) theory Example_Verification imports "HOL-SPARK-Examples.Greatest_Common_Divisor" Simple_Greatest_Common_Divisor begin (*>*) chapter \Verifying an Example Program\ text \ \label{sec:example-verification} \begin{figure} \lstinputlisting{Gcd.ads} \lstinputlisting{Gcd.adb} \caption{\SPARK{} program for computing the greatest common divisor} \label{fig:gcd-prog} \end{figure} \begin{figure} \input{Greatest_Common_Divisor} \caption{Correctness proof for the greatest common divisor program} \label{fig:gcd-proof} \end{figure} We will now explain the usage of the \SPARK{} verification environment by proving the correctness of an example program. As an example, we use a program for computing the \emph{greatest common divisor} of two natural numbers shown in \figref{fig:gcd-prog}, -which has been taken from the book about \SPARK{} by Barnes @{cite \\S 11.6\ Barnes}. +which has been taken from the book about \SPARK{} by Barnes \<^cite>\\\S 11.6\ in Barnes\. \ section \Importing \SPARK{} VCs into Isabelle\ text \ In order to specify that the \SPARK{} procedure \texttt{G\_C\_D} behaves like its mathematical counterpart, Barnes introduces a \emph{proof function} \texttt{Gcd} in the package specification. Invoking the \SPARK{} Examiner and Simplifier on this program yields a file \texttt{g\_c\_d.siv} containing the simplified VCs, as well as files \texttt{g\_c\_d.fdl} and \texttt{g\_c\_d.rls}, containing FDL declarations and rules, respectively. The files generated by \SPARK{} are assumed to reside in the subdirectory \texttt{greatest\_common\_divisor}. For \texttt{G\_C\_D} the Examiner generates ten VCs, eight of which are proved automatically by the Simplifier. We now show how to prove the remaining two VCs interactively using HOL-\SPARK{}. For this purpose, we create a \emph{theory} \texttt{Greatest\_Common\_Divisor}, which is shown in \figref{fig:gcd-proof}. A theory file always starts with the keyword \isa{\isacommand{theory}} followed by the name of the theory, which must be the same as the file name. The theory name is followed by the keyword \isa{\isacommand{imports}} and a list of theories imported by the current theory. All theories using the HOL-\SPARK{} verification environment must import the theory \texttt{SPARK}. In addition, we also include the \texttt{GCD} theory. The list of imported theories is followed by the \isa{\isacommand{begin}} keyword. In order to interactively process the theory shown in \figref{fig:gcd-proof}, we start Isabelle with the command \begin{verbatim} isabelle jedit -l HOL-SPARK Greatest_Common_Divisor.thy \end{verbatim} The option ``\texttt{-l HOL-SPARK}'' instructs Isabelle to load the right object logic image containing the verification environment. Each proof function occurring in the specification of a \SPARK{} program must be linked with a corresponding Isabelle function. This is accomplished by the command \isa{\isacommand{spark\_proof\_functions}}, which expects a list of equations of the form \emph{name}\texttt{\ =\ }\emph{term}, where \emph{name} is the name of the proof function and \emph{term} is the corresponding Isabelle term. In the case of \texttt{gcd}, both the \SPARK{} proof function and its Isabelle counterpart happen to have the same name. Isabelle checks that the type of the term linked with a proof function agrees with the type of the function declared in the \texttt{*.fdl} file. It is worth noting that the \isa{\isacommand{spark\_proof\_functions}} command can be invoked both outside, i.e.\ before \isa{\isacommand{spark\_open}}, and inside the environment, i.e.\ after \isa{\isacommand{spark\_open}}, but before any \isa{\isacommand{spark\_vc}} command. The former variant is useful when having to declare proof functions that are shared by several procedures, whereas the latter has the advantage that the type of the proof function can be checked immediately, since the VCs, and hence also the declarations of proof functions in the \texttt{*.fdl} file have already been loaded. \begin{figure} \begin{flushleft} \tt Context: \\ \ \\ \begin{tabular}{ll} fixes & \m ::\\ "\int\" \\ and & \n ::\\ "\int\" \\ and & \c ::\\ "\int\" \\ and & \d ::\\ "\int\" \\ assumes & \g_c_d_rules1:\\ "\0 \ integer__size\" \\ and & \g_c_d_rules6:\\ "\0 \ natural__size\" \\ \multicolumn{2}{l}{notes definition} \\ \multicolumn{2}{l}{\hspace{2ex}\defns =\\ `\integer__first = - 2147483648\`} \\ \multicolumn{2}{l}{\hspace{4ex}`\integer__last = 2147483647\`} \\ \multicolumn{2}{l}{\hspace{4ex}\dots} \end{tabular}\ \\[1.5ex] \ \\ Definitions: \\ \ \\ \begin{tabular}{ll} \g_c_d_rules2:\ & \integer__first = - 2147483648\ \\ \g_c_d_rules3:\ & \integer__last = 2147483647\ \\ \dots \end{tabular}\ \\[1.5ex] \ \\ Verification conditions: \\ \ \\ path(s) from assertion of line 10 to assertion of line 10 \\ \ \\ \procedure_g_c_d_4\\ (unproved) \\ \ \ \begin{tabular}{ll} assumes & \H1:\\ "\0 \ c\" \\ and & \H2:\\ "\0 < d\" \\ and & \H3:\\ "\gcd c d = gcd m n\" \\ \dots \\ shows & "\0 < c - c sdiv d * d\" \\ and & "\gcd d (c - c sdiv d * d) = gcd m n\ \end{tabular}\ \\[1.5ex] \ \\ path(s) from assertion of line 10 to finish \\ \ \\ \procedure_g_c_d_11\\ (unproved) \\ \ \ \begin{tabular}{ll} assumes & \H1:\\ "\0 \ c\" \\ and & \H2:\\ "\0 < d\" \\ and & \H3:\\ "\gcd c d = gcd m n\" \\ \dots \\ shows & "\d = gcd m n\" \end{tabular} \end{flushleft} \caption{Output of \isa{\isacommand{spark\_status}} for \texttt{g\_c\_d.siv}} \label{fig:gcd-status} \end{figure} We now instruct Isabelle to open a new verification environment and load a set of VCs. This is done using the command \isa{\isacommand{spark\_open}}, which must be given the name of a \texttt{*.siv} file as an argument. Behind the scenes, Isabelle parses this file and the corresponding \texttt{*.fdl} and \texttt{*.rls} files, and converts the VCs to Isabelle terms. Using the command \isa{\isacommand{spark\_status}}, the user can display the current VCs together with their status (proved, unproved). The variants \isa{\isacommand{spark\_status}\ (proved)} and \isa{\isacommand{spark\_status}\ (unproved)} show only proved and unproved VCs, respectively. For \texttt{g\_c\_d.siv}, the output of \isa{\isacommand{spark\_status}} is shown in \figref{fig:gcd-status}. To minimize the number of assumptions, and hence the size of the VCs, FDL rules of the form ``\dots\ \texttt{may\_be\_replaced\_by}\ \dots'' are turned into native Isabelle definitions, whereas other rules are modelled as assumptions. \ section \Proving the VCs\ text \ \label{sec:proving-vcs} The two open VCs are \procedure_g_c_d_4\ and \procedure_g_c_d_11\, both of which contain the \gcd\ proof function that the \SPARK{} Simplifier does not know anything about. The proof of a particular VC can be started with the \isa{\isacommand{spark\_vc}} command, which is similar to the standard \isa{\isacommand{lemma}} and \isa{\isacommand{theorem}} commands, with the difference that it only takes a name of a VC but no formula as an argument. A VC can have several conclusions that can be referenced by the identifiers \?C1\, \?C2\, etc. If there is just one conclusion, it can also be referenced by \?thesis\. It is important to note that the \texttt{div} operator of FDL behaves differently from the \div\ operator of Isabelle/HOL on negative numbers. The former always truncates towards zero, whereas the latter truncates towards minus infinity. This is why the FDL \texttt{div} operator is mapped to the \sdiv\ operator in Isabelle/HOL, which is defined as @{thm [display] sdiv_def} For example, we have that @{lemma "-5 sdiv 4 = -1" by (simp add: sdiv_neg_pos)}, but @{lemma "(-5::int) div 4 = -2" by simp}. For non-negative dividend and divisor, \sdiv\ is equivalent to \div\, as witnessed by theorem \sdiv_pos_pos\: @{thm [display,mode=no_brackets] sdiv_pos_pos} In contrast, the behaviour of the FDL \texttt{mod} operator is equivalent to the one of Isabelle/HOL. Moreover, since FDL has no counterpart of the \SPARK{} operator \textbf{rem}, the \SPARK{} expression \texttt{c}\ \textbf{rem}\ \texttt{d} just becomes \c - c sdiv d * d\ in Isabelle. The first conclusion of \procedure_g_c_d_4\ requires us to prove that the remainder of \c\ and \d\ is greater than \0\. To do this, we use the theorem \minus_div_mult_eq_mod [symmetric]\ describing the correspondence between \div\ and \mod\ @{thm [display] minus_div_mult_eq_mod [symmetric]} together with the theorem \pos_mod_sign\ saying that the result of the \mod\ operator is non-negative when applied to a non-negative divisor: @{thm [display] pos_mod_sign} We will also need the aforementioned theorem \sdiv_pos_pos\ in order for the standard Isabelle/HOL theorems about \div\ to be applicable to the VC, which is formulated using \sdiv\ rather that \div\. Note that the proof uses \texttt{`\0 \ c\`} and \texttt{`\0 < d\`} rather than \H1\ and \H2\ to refer to the hypotheses of the current VC. While the latter variant seems more compact, it is not particularly robust, since the numbering of hypotheses can easily change if the corresponding program is modified, making the proof script hard to adjust when there are many hypotheses. Moreover, proof scripts using abbreviations like \H1\ and \H2\ are hard to read without assistance from Isabelle. The second conclusion of \procedure_g_c_d_4\ requires us to prove that the \gcd\ of \d\ and the remainder of \c\ and \d\ is equal to the \gcd\ of the original input values \m\ and \n\, which is the actual \emph{invariant} of the procedure. This is a consequence of theorem \gcd_non_0_int\ @{thm [display] gcd_non_0_int} Again, we also need theorems \minus_div_mult_eq_mod [symmetric]\ and \sdiv_pos_pos\ to justify that \SPARK{}'s \textbf{rem} operator is equivalent to Isabelle's \mod\ operator for non-negative operands. The VC \procedure_g_c_d_11\ says that if the loop invariant holds before the last iteration of the loop, the postcondition of the procedure will hold after execution of the loop body. To prove this, we observe that the remainder of \c\ and \d\, and hence \c mod d\ is \0\ when exiting the loop. This implies that \gcd c d = d\, since \c\ is divisible by \d\, so the conclusion follows using the assumption \gcd c d = gcd m n\. This concludes the proofs of the open VCs, and hence the \SPARK{} verification environment can be closed using the command \isa{\isacommand{spark\_end}}. This command checks that all VCs have been proved and issues an error message if there are remaining unproved VCs. Moreover, Isabelle checks that there is no open \SPARK{} verification environment when the final \isa{\isacommand{end}} command of a theory is encountered. \ section \Optimizing the proof\ text \ \begin{figure} \lstinputlisting{Simple_Gcd.adb} \input{Simple_Greatest_Common_Divisor} \caption{Simplified greatest common divisor program and proof} \label{fig:simple-gcd-proof} \end{figure} When looking at the program from \figref{fig:gcd-prog} once again, several optimizations come to mind. First of all, like the input parameters of the procedure, the local variables \texttt{C}, \texttt{D}, and \texttt{R} can be declared as \texttt{Natural} rather than \texttt{Integer}. Since natural numbers are non-negative by construction, the values computed by the algorithm are trivially proved to be non-negative. Since we are working with non-negative numbers, we can also just use \SPARK{}'s \textbf{mod} operator instead of \textbf{rem}, which spares us an application of theorems \minus_div_mult_eq_mod [symmetric]\ -and \sdiv_pos_pos\. Finally, as noted by Barnes @{cite \\S 11.5\ Barnes}, +and \sdiv_pos_pos\. Finally, as noted by Barnes \<^cite>\\\S 11.5\ in Barnes\, we can simplify matters by placing the \textbf{assert} statement between \textbf{while} and \textbf{loop} rather than directly after the \textbf{loop}. In the former case, the loop invariant has to be proved only once, whereas in the latter case, it has to be proved twice: since the \textbf{assert} occurs after the check of the exit condition, the invariant has to be proved for the path from the \textbf{assert} statement to the \textbf{assert} statement, and for the path from the \textbf{assert} statement to the postcondition. In the case of the \texttt{G\_C\_D} procedure, this might not seem particularly problematic, since the proof of the invariant is very simple, but it can unnecessarily complicate matters if the proof of the invariant is non-trivial. The simplified program for computing the greatest common divisor, together with its correctness proof, is shown in \figref{fig:simple-gcd-proof}. Since the package specification has not changed, we only show the body of the packages. The two VCs can now be proved by a single application of Isabelle's proof method \simp\. \ (*<*) end (*>*) diff --git a/src/HOL/SPARK/Manual/VC_Principles.thy b/src/HOL/SPARK/Manual/VC_Principles.thy --- a/src/HOL/SPARK/Manual/VC_Principles.thy +++ b/src/HOL/SPARK/Manual/VC_Principles.thy @@ -1,152 +1,152 @@ (*<*) theory VC_Principles imports Proc1 Proc2 begin (*>*) chapter \Principles of VC generation\ text \ \label{sec:vc-principles} In this section, we will discuss some aspects of VC generation that are useful for understanding and optimizing the VCs produced by the \SPARK{} Examiner. \begin{figure} \lstinputlisting{loop_invariant.ads} \lstinputlisting{loop_invariant.adb} \caption{Assertions in for-loops} \label{fig:loop-invariant-ex} \end{figure} \begin{figure} \begin{tikzpicture} \tikzstyle{box}=[draw, drop shadow, fill=white, rounded corners] \node[box] (pre) at (0,0) {precondition}; \node[box] (assn) at (0,-3) {assertion}; \node[box] (post) at (0,-6) {postcondition}; \draw[-latex] (pre) -- node [right] {\small$(1 - 1) * b \mod 2^{32} = 0$} (assn); \draw[-latex] (assn) .. controls (2.5,-4.5) and (2.5,-1.5) .. % node [right] {\small$\begin{array}{l} % (i - 1) * b \mod 2^{32} = c ~\longrightarrow \\ % (i + 1 - 1) * b \mod 2^{32} ~= \\ % (c + b) \mod 2^{32} % \end{array}$} (assn); \draw[-latex] (assn) -- node [right] {\small$\begin{array}{l} % (a - 1) * b \mod 2^{32} = c ~\longrightarrow \\ % a * b \mod 2^{32} = (c + b) \mod 2^{32} % \end{array}$} (post); \draw[-latex] (pre) .. controls (-2,-3) .. % node [left] {\small$\begin{array}{l} % \neg 1 \le a ~\longrightarrow \\ % a * b \mod 2^{32} = 0 % \end{array}$} (post); \end{tikzpicture} \caption{Control flow graph for procedure \texttt{Proc1}} \label{fig:proc1-graph} \end{figure} \begin{figure} \begin{tikzpicture} \tikzstyle{box}=[draw, drop shadow, fill=white, rounded corners] \node[box] (pre) at (0,0) {precondition}; \node[box] (assn1) at (2,-3) {assertion 1}; \node[box] (assn2) at (2,-6) {assertion 2}; \node[box] (post) at (0,-9) {postcondition}; \draw[-latex] (pre) -- node [right] {\small$(1 - 1) * b \mod 2^{32} = 0$} (assn1); \draw[-latex] (assn1) -- node [left] {\small$\begin{array}{l} % (i - 1) * b \mod 2^{32} = c \\ % \longrightarrow \\ % i * b \mod 2^{32} = \\ % (c + b) \mod 2^{32} % \end{array}$} (assn2); \draw[-latex] (assn2) .. controls (4.5,-7.5) and (4.5,-1.5) .. % node [right] {\small$\begin{array}{l} % i * b \mod 2^{32} = c ~\longrightarrow \\ % (i + 1 - 1) * b \mod 2^{32} = c % \end{array}$} (assn1); \draw[-latex] (assn2) -- node [right] {\small$\begin{array}{l} % a * b \mod 2^{32} = c ~\longrightarrow \\ % a * b \mod 2^{32} = c % \end{array}$} (post); \draw[-latex] (pre) .. controls (-3,-3) and (-3,-6) .. % node [left,very near start] {\small$\begin{array}{l} % \neg 1 \le a ~\longrightarrow \\ % a * b \mod 2^{32} = 0 % \end{array}$} (post); \end{tikzpicture} \caption{Control flow graph for procedure \texttt{Proc2}} \label{fig:proc2-graph} \end{figure} -As explained by Barnes @{cite \\S 11.5\ Barnes}, the \SPARK{} Examiner unfolds the loop +As explained by Barnes \<^cite>\\\S 11.5\ in Barnes\, the \SPARK{} Examiner unfolds the loop \begin{lstlisting} for I in T range L .. U loop --# assert P (I); S; end loop; \end{lstlisting} to \begin{lstlisting} if L <= U then I := L; loop --# assert P (I); S; exit when I = U; I := I + 1; end loop; end if; \end{lstlisting} Due to this treatment of for-loops, the user essentially has to prove twice that \texttt{S} preserves the invariant \textit{\texttt{P}}, namely for the path from the assertion to the assertion and from the assertion to the next cut point following the loop. The preservation of the invariant has to be proved even more often when the loop is followed by an if-statement. For trivial invariants, this might not seem like a big problem, but in realistic applications, where invariants are complex, this can be a major inconvenience. Often, the proofs of the invariant differ only in a few variables, so it is tempting to just copy and modify existing proof scripts, but this leads to proofs that are hard to maintain. The problem of having to prove the invariant several times can be avoided by rephrasing the above for-loop to \begin{lstlisting} for I in T range L .. U loop --# assert P (I); S; --# assert P (I + 1) end loop; \end{lstlisting} The VC for the path from the second assertion to the first assertion is trivial and can usually be proved automatically by the \SPARK{} Simplifier, whereas the VC for the path from the first assertion to the second assertion actually expresses the fact that \texttt{S} preserves the invariant. We illustrate this technique using the example package shown in \figref{fig:loop-invariant-ex}. It contains two procedures \texttt{Proc1} and \texttt{Proc2}, both of which implement multiplication via addition. The procedures have the same specification, but in \texttt{Proc1}, only one \textbf{assert} statement is placed at the beginning of the loop, whereas \texttt{Proc2} employs the trick explained above. After applying the \SPARK{} Simplifier to the VCs generated for \texttt{Proc1}, two very similar VCs @{thm [display] (concl) procedure_proc1_5 [simplified pow_2_32_simp]} and @{thm [display,margin=60] (concl) procedure_proc1_8 [simplified pow_2_32_simp]} remain, whereas for \texttt{Proc2}, only the first of the above VCs remains. Why placing \textbf{assert} statements both at the beginning and at the end of the loop body simplifies the proof of the invariant should become obvious when looking at \figref{fig:proc1-graph} and \figref{fig:proc2-graph} showing the \emph{control flow graphs} for \texttt{Proc1} and \texttt{Proc2}, respectively. The nodes in the graph correspond to cut points in the program, and the paths between the cut points are annotated with the corresponding VCs. To reduce the size of the graphs, we do not show nodes and edges corresponding to runtime checks. The VCs for the path bypassing the loop and for the path from the precondition to the (first) assertion are the same for both procedures. The graph for \texttt{Proc1} contains two VCs expressing that the invariant is preserved by the execution of the loop body: one for the path from the assertion to the assertion, and another one for the path from the assertion to the conclusion, which corresponds to the last iteration of the loop. The latter VC can be obtained from the former by simply replacing $i$ by $a$. In contrast, the graph for \texttt{Proc2} contains only one such VC for the path from assertion 1 to assertion 2. The VC for the path from assertion 2 to assertion 1 is trivial, and so is the VC for the path from assertion 2 to the postcondition, expressing that the loop invariant implies the postcondition when the loop has terminated. \ (*<*) end (*>*) diff --git a/src/HOL/Unix/Unix.thy b/src/HOL/Unix/Unix.thy --- a/src/HOL/Unix/Unix.thy +++ b/src/HOL/Unix/Unix.thy @@ -1,1047 +1,1044 @@ (* Title: HOL/Unix/Unix.thy Author: Markus Wenzel, TU Muenchen *) section \Unix file-systems \label{sec:unix-file-system}\ theory Unix imports Nested_Environment "HOL-Library.Sublist" begin text \ We give a simple mathematical model of the basic structures underlying the Unix file-system, together with a few fundamental operations that could be imagined to be performed internally by the Unix kernel. This forms the basis for the set of Unix system-calls to be introduced later (see \secref{sec:unix-trans}), which are the actual interface offered to processes running in user-space. \<^medskip> Basically, any Unix file is either a \<^emph>\plain file\ or a \<^emph>\directory\, consisting of some \<^emph>\content\ plus \<^emph>\attributes\. The content of a plain file is plain text. The content of a directory is a mapping from names to further files.\<^footnote>\In fact, this is the only way that names get associated with files. In Unix files do \<^emph>\not\ have a name in itself. Even more, any number of names may be associated with the very same file due to \<^emph>\hard links\ (although this is excluded from our model).\ Attributes include information to control various ways to access the file (read, write etc.). Our model will be quite liberal in omitting excessive detail that is easily seen to be ``irrelevant'' for the aspects of Unix file-systems to be discussed here. First of all, we ignore character and block special files, pipes, sockets, hard links, symbolic links, and mount points. \ subsection \Names\ text \ User ids and file name components shall be represented by natural numbers (without loss of generality). We do not bother about encoding of actual names (e.g.\ strings), nor a mapping between user names and user ids as would be present in a reality. \ type_synonym uid = nat type_synonym name = nat type_synonym path = "name list" subsection \Attributes\ text \ Unix file attributes mainly consist of \<^emph>\owner\ information and a number of \<^emph>\permission\ bits which control access for ``user'', ``group'', and ``others'' (see the Unix man pages \chmod(2)\ and \stat(2)\ for more details). \<^medskip> Our model of file permissions only considers the ``others'' part. The ``user'' field may be omitted without loss of overall generality, since the owner is usually able to change it anyway by performing \chmod\.\<^footnote>\The inclined Unix expert may try to figure out some exotic arrangements of a real-world Unix file-system such that the owner of a file is unable to apply the \chmod\ system call.\ We omit ``group'' permissions as a genuine simplification as we just do not intend to discuss a model of multiple groups and group membership, but pretend that everyone is member of a single global group.\<^footnote>\A general HOL model of user group structures and related - issues is given in @{cite "Naraschewski:2001"}.\ + issues is given in \<^cite>\"Naraschewski:2001"\.\ \ datatype perm = Readable | Writable | Executable \ \(ignored)\ type_synonym perms = "perm set" record att = owner :: uid others :: perms text \ For plain files \<^term>\Readable\ and \<^term>\Writable\ specify read and write access to the actual content, i.e.\ the string of text stored here. For directories \<^term>\Readable\ determines if the set of entry names may be accessed, and \<^term>\Writable\ controls the ability to create or delete any entries (both plain files or sub-directories). As another simplification, we ignore the \<^term>\Executable\ permission altogether. In reality it would indicate executable plain files (also known as ``binaries''), or control actual lookup of directory entries (recall that mere directory browsing is controlled via \<^term>\Readable\). Note that the latter means that in order to perform any file-system operation whatsoever, all directories encountered on the path would have to grant \<^term>\Executable\. We ignore this detail and pretend that all directories give \<^term>\Executable\ permission to anybody. \ subsection \Files\ text \ In order to model the general tree structure of a Unix file-system we use the arbitrarily branching datatype \<^typ>\('a, 'b, 'c) env\ from the - standard library of Isabelle/HOL @{cite "Bauer-et-al:2002:HOL-Library"}. + standard library of Isabelle/HOL \<^cite>\"Bauer-et-al:2002:HOL-Library"\. This type provides constructors \<^term>\Val\ and \<^term>\Env\ as follows: \<^medskip> {\def\isastyleminor{\isastyle} \begin{tabular}{l} @{term [source] "Val :: 'a \ ('a, 'b, 'c) env"} \\ @{term [source] "Env :: 'b \ ('c \ ('a, 'b, 'c) env option) \ ('a, 'b, 'c) env"} \\ \end{tabular}} \<^medskip> Here the parameter \<^typ>\'a\ refers to plain values occurring at leaf positions, parameter \<^typ>\'b\ to information kept with inner branch nodes, and parameter \<^typ>\'c\ to the branching type of the tree structure. For our purpose we use the type instance with \<^typ>\att \ string\ (representing plain files), \<^typ>\att\ (for attributes of directory nodes), and \<^typ>\name\ (for the index type of directory nodes). \ type_synonym "file" = "(att \ string, att, name) env" text \ \<^medskip> The HOL library also provides \<^term>\lookup\ and \<^term>\update\ operations for general tree structures with the subsequent primitive recursive characterizations. \<^medskip> {\def\isastyleminor{\isastyle} \begin{tabular}{l} @{term [source] "lookup :: ('a, 'b, 'c) env \ 'c list \ ('a, 'b, 'c) env option"} \\ @{term [source] "update :: 'c list \ ('a, 'b, 'c) env option \ ('a, 'b, 'c) env \ ('a, 'b, 'c) env"} \\ \end{tabular}} @{thm [display, indent = 2, eta_contract = false] lookup_eq [no_vars]} @{thm [display, indent = 2, eta_contract = false] update_eq [no_vars]} - Several further properties of these operations are proven in @{cite - "Bauer-et-al:2002:HOL-Library"}. These will be routinely used later on + Several further properties of these operations are proven in \<^cite>\"Bauer-et-al:2002:HOL-Library"\. These will be routinely used later on without further notice. \<^bigskip> Apparently, the elements of type \<^typ>\file\ contain an \<^typ>\att\ component in either case. We now define a few auxiliary operations to manipulate this field uniformly, following the conventions for record types - in Isabelle/HOL @{cite "Nipkow-et-al:2000:HOL"}. + in Isabelle/HOL \<^cite>\"Nipkow-et-al:2000:HOL"\. \ definition "attributes file = (case file of Val (att, text) \ att | Env att dir \ att)" definition "map_attributes f file = (case file of Val (att, text) \ Val (f att, text) | Env att dir \ Env (f att) dir)" lemma [simp]: "attributes (Val (att, text)) = att" by (simp add: attributes_def) lemma [simp]: "attributes (Env att dir) = att" by (simp add: attributes_def) lemma [simp]: "attributes (map_attributes f file) = f (attributes file)" by (cases "file") (simp_all add: attributes_def map_attributes_def split_tupled_all) lemma [simp]: "map_attributes f (Val (att, text)) = Val (f att, text)" by (simp add: map_attributes_def) lemma [simp]: "map_attributes f (Env att dir) = Env (f att) dir" by (simp add: map_attributes_def) subsection \Initial file-systems\ text \ Given a set of \<^emph>\known users\ a file-system shall be initialized by providing an empty home directory for each user, with read-only access for everyone else. (Note that we may directly use the user id as home directory name, since both types have been identified.) Certainly, the very root directory is owned by the super user (who has user id 0). \ definition "init users = Env \owner = 0, others = {Readable}\ (\u. if u \ users then Some (Env \owner = u, others = {Readable}\ Map.empty) else None)" subsection \Accessing file-systems\ text \ The main internal file-system operation is access of a file by a user, requesting a certain set of permissions. The resulting \<^typ>\file option\ indicates if a file had been present at the corresponding \<^typ>\path\ and if access was granted according to the permissions recorded within the file-system. - Note that by the rules of Unix file-system security (e.g.\ @{cite - "Tanenbaum:1992"}) both the super-user and owner may always access a file + Note that by the rules of Unix file-system security (e.g.\ \<^cite>\"Tanenbaum:1992"\) both the super-user and owner may always access a file unconditionally (in our simplified model). \ definition "access root path uid perms = (case lookup root path of None \ None | Some file \ if uid = 0 \ uid = owner (attributes file) \ perms \ others (attributes file) then Some file else None)" text \ \<^medskip> Successful access to a certain file is the main prerequisite for system-calls to be applicable (cf.\ \secref{sec:unix-trans}). Any modification of the file-system is then performed using the basic \<^term>\update\ operation. \<^medskip> We see that \<^term>\access\ is just a wrapper for the basic \<^term>\lookup\ function, with additional checking of attributes. Subsequently we establish a few auxiliary facts that stem from the primitive \<^term>\lookup\ used within \<^term>\access\. \ lemma access_empty_lookup: "access root path uid {} = lookup root path" by (simp add: access_def split: option.splits) lemma access_some_lookup: "access root path uid perms = Some file \ lookup root path = Some file" by (simp add: access_def split: option.splits if_splits) lemma access_update_other: assumes parallel: "path' \ path" shows "access (update path' opt root) path uid perms = access root path uid perms" proof - from parallel obtain y z xs ys zs where "y \ z" and "path' = xs @ y # ys" and "path = xs @ z # zs" by (blast dest: parallel_decomp) then have "lookup (update path' opt root) path = lookup root path" by (blast intro: lookup_update_other) then show ?thesis by (simp only: access_def) qed section \File-system transitions \label{sec:unix-trans}\ subsection \Unix system calls \label{sec:unix-syscall}\ text \ - According to established operating system design (cf.\ @{cite - "Tanenbaum:1992"}) user space processes may only initiate system operations + According to established operating system design (cf.\ \<^cite>\"Tanenbaum:1992"\) user space processes may only initiate system operations by a fixed set of \<^emph>\system-calls\. This enables the kernel to enforce certain security policies in the first place.\<^footnote>\Incidently, this is the very same principle employed by any ``LCF-style'' theorem proving system according to Milner's principle of ``correctness by construction'', such as Isabelle/HOL itself.\ \<^medskip> In our model of Unix we give a fixed datatype \operation\ for the syntax of system-calls, together with an inductive definition of file-system state transitions of the form \root \x\ root'\ for the operational semantics. \ datatype operation = Read uid string path | Write uid string path | Chmod uid perms path | Creat uid perms path | Unlink uid path | Mkdir uid perms path | Rmdir uid path | Readdir uid "name set" path text \ The \<^typ>\uid\ field of an operation corresponds to the \<^emph>\effective user id\ of the underlying process, although our model never mentions processes explicitly. The other parameters are provided as arguments by the caller; the \<^term>\path\ one is common to all kinds of system-calls. \ primrec uid_of :: "operation \ uid" where "uid_of (Read uid text path) = uid" | "uid_of (Write uid text path) = uid" | "uid_of (Chmod uid perms path) = uid" | "uid_of (Creat uid perms path) = uid" | "uid_of (Unlink uid path) = uid" | "uid_of (Mkdir uid path perms) = uid" | "uid_of (Rmdir uid path) = uid" | "uid_of (Readdir uid names path) = uid" primrec path_of :: "operation \ path" where "path_of (Read uid text path) = path" | "path_of (Write uid text path) = path" | "path_of (Chmod uid perms path) = path" | "path_of (Creat uid perms path) = path" | "path_of (Unlink uid path) = path" | "path_of (Mkdir uid perms path) = path" | "path_of (Rmdir uid path) = path" | "path_of (Readdir uid names path) = path" text \ \<^medskip> Note that we have omitted explicit \Open\ and \Close\ operations, pretending that \<^term>\Read\ and \<^term>\Write\ would already take care of this behind the scenes. Thus we have basically treated actual sequences of real system-calls \open\--\read\/\write\--\close\ as atomic. In principle, this could make big a difference in a model with explicit concurrent processes. On the other hand, even on a real Unix system the exact scheduling of concurrent \open\ and \close\ operations does \<^emph>\not\ directly affect the success of corresponding \read\ or \write\. Unix allows several processes to have files opened at the same time, even for writing! Certainly, the result from reading the contents later may be hard to predict, but the system-calls involved here will succeed in any case. \<^bigskip> The operational semantics of system calls is now specified via transitions of the file-system configuration. This is expressed as an inductive relation (although there is no actual recursion involved here). \ inductive transition :: "file \ operation \ file \ bool" ("_ \_\ _" [90, 1000, 90] 100) where read: "root \(Read uid text path)\ root" if "access root path uid {Readable} = Some (Val (att, text))" | "write": "root \(Write uid text path)\ update path (Some (Val (att, text))) root" if "access root path uid {Writable} = Some (Val (att, text'))" | chmod: "root \(Chmod uid perms path)\ update path (Some (map_attributes (others_update (\_. perms)) file)) root" if "access root path uid {} = Some file" and "uid = 0 \ uid = owner (attributes file)" | creat: "root \(Creat uid perms path)\ update path (Some (Val (\owner = uid, others = perms\, []))) root" if "path = parent_path @ [name]" and "access root parent_path uid {Writable} = Some (Env att parent)" and "access root path uid {} = None" | unlink: "root \(Unlink uid path)\ update path None root" if "path = parent_path @ [name]" and "access root parent_path uid {Writable} = Some (Env att parent)" and "access root path uid {} = Some (Val plain)" | mkdir: "root \(Mkdir uid perms path)\ update path (Some (Env \owner = uid, others = perms\ Map.empty)) root" if "path = parent_path @ [name]" and "access root parent_path uid {Writable} = Some (Env att parent)" and "access root path uid {} = None" | rmdir: "root \(Rmdir uid path)\ update path None root" if "path = parent_path @ [name]" and "access root parent_path uid {Writable} = Some (Env att parent)" and "access root path uid {} = Some (Env att' Map.empty)" | readdir: "root \(Readdir uid names path)\ root" if "access root path uid {Readable} = Some (Env att dir)" and "names = dom dir" text \ \<^medskip> Certainly, the above specification is central to the whole formal development. Any of the results to be established later on are only meaningful to the outside world if this transition system provides an adequate model of real Unix systems. This kind of ``reality-check'' of a formal model is the well-known problem of \<^emph>\validation\. If in doubt, one may consider to compare our definition with the informal specifications given the corresponding Unix man pages, or even peek at an - actual implementation such as @{cite "Torvalds-et-al:Linux"}. Another common + actual implementation such as \<^cite>\"Torvalds-et-al:Linux"\. Another common way to gain confidence into the formal model is to run simple simulations (see \secref{sec:unix-examples}), and check the results with that of experiments performed on a real Unix system. \ subsection \Basic properties of single transitions \label{sec:unix-single-trans}\ text \ The transition system \root \x\ root'\ defined above determines a unique result \<^term>\root'\ from given \<^term>\root\ and \<^term>\x\ (this is holds rather trivially, since there is even only one clause for each operation). This uniqueness statement will simplify our subsequent development to some extent, since we only have to reason about a partial function rather than a general relation. \ theorem transition_uniq: assumes root': "root \x\ root'" and root'': "root \x\ root''" shows "root' = root''" using root'' proof cases case read with root' show ?thesis by cases auto next case "write" with root' show ?thesis by cases auto next case chmod with root' show ?thesis by cases auto next case creat with root' show ?thesis by cases auto next case unlink with root' show ?thesis by cases auto next case mkdir with root' show ?thesis by cases auto next case rmdir with root' show ?thesis by cases auto next case readdir with root' show ?thesis by cases blast+ qed text \ \<^medskip> Apparently, file-system transitions are \<^emph>\type-safe\ in the sense that the result of transforming an actual directory yields again a directory. \ theorem transition_type_safe: assumes tr: "root \x\ root'" and inv: "\att dir. root = Env att dir" shows "\att dir. root' = Env att dir" proof (cases "path_of x") case Nil with tr inv show ?thesis by cases (auto simp add: access_def split: if_splits) next case Cons from tr obtain opt where "root' = root \ root' = update (path_of x) opt root" by cases auto with inv Cons show ?thesis by (auto simp add: update_eq split: list.splits) qed text \ The previous result may be seen as the most basic invariant on the file-system state that is enforced by any proper kernel implementation. So user processes --- being bound to the system-call interface --- may never mess up a file-system such that the root becomes a plain file instead of a directory, which would be a strange situation indeed. \ subsection \Iterated transitions\ text \ Iterated system transitions via finite sequences of system operations are modeled inductively as follows. In a sense, this relation describes the cumulative effect of the sequence of system-calls issued by a number of running processes over a finite amount of time. \ inductive transitions :: "file \ operation list \ file \ bool" ("_ \_\ _" [90, 1000, 90] 100) where nil: "root \[]\ root" | cons: "root \(x # xs)\ root''" if "root \x\ root'" and "root' \xs\ root''" text \ \<^medskip> We establish a few basic facts relating iterated transitions with single ones, according to the recursive structure of lists. \ lemma transitions_nil_eq: "root \[]\ root' \ root = root'" proof assume "root \[]\ root'" then show "root = root'" by cases simp_all next assume "root = root'" then show "root \[]\ root'" by (simp only: transitions.nil) qed lemma transitions_cons_eq: "root \(x # xs)\ root'' \ (\root'. root \x\ root' \ root' \xs\ root'')" proof assume "root \(x # xs)\ root''" then show "\root'. root \x\ root' \ root' \xs\ root''" by cases auto next assume "\root'. root \x\ root' \ root' \xs\ root''" then show "root \(x # xs)\ root''" by (blast intro: transitions.cons) qed text \ The next two rules show how to ``destruct'' known transition sequences. Note that the second one actually relies on the uniqueness property of the basic transition system (see \secref{sec:unix-single-trans}). \ lemma transitions_nilD: "root \[]\ root' \ root' = root" by (simp add: transitions_nil_eq) lemma transitions_consD: assumes list: "root \(x # xs)\ root''" and hd: "root \x\ root'" shows "root' \xs\ root''" proof - from list obtain r' where r': "root \x\ r'" and root'': "r' \xs\ root''" by cases simp_all from r' hd have "r' = root'" by (rule transition_uniq) with root'' show "root' \xs\ root''" by simp qed text \ \<^medskip> The following fact shows how an invariant \<^term>\Q\ of single transitions with property \<^term>\P\ may be transferred to iterated transitions. The proof is rather obvious by rule induction over the definition of \<^term>\root \xs\ root'\. \ lemma transitions_invariant: assumes r: "\r x r'. r \x\ r' \ Q r \ P x \ Q r'" and trans: "root \xs\ root'" shows "Q root \ \x \ set xs. P x \ Q root'" using trans proof induct case nil show ?case by (rule nil.prems) next case (cons root x root' xs root'') note P = \\x \ set (x # xs). P x\ then have "P x" by simp with \root \x\ root'\ and \Q root\ have Q': "Q root'" by (rule r) from P have "\x \ set xs. P x" by simp with Q' show "Q root''" by (rule cons.hyps) qed text \ As an example of applying the previous result, we transfer the basic type-safety property (see \secref{sec:unix-single-trans}) from single transitions to iterated ones, which is a rather obvious result indeed. \ theorem transitions_type_safe: assumes "root \xs\ root'" and "\att dir. root = Env att dir" shows "\att dir. root' = Env att dir" using transition_type_safe and assms proof (rule transitions_invariant) show "\x \ set xs. True" by blast qed section \Executable sequences\ text \ An inductively defined relation such as the one of \root \x\ root'\ (see \secref{sec:unix-syscall}) has two main aspects. First of all, the resulting system admits a certain set of transition rules (introductions) as given in the specification. Furthermore, there is an explicit least-fixed-point construction involved, which results in induction (and case-analysis) rules to eliminate known transitions in an exhaustive manner. \<^medskip> Subsequently, we explore our transition system in an experimental style, mainly using the introduction rules with basic algebraic properties of the underlying structures. The technique closely resembles that of Prolog combined with functional evaluation in a very simple manner. Just as the ``closed-world assumption'' is left implicit in Prolog, we do not refer to induction over the whole transition system here. So this is still purely positive reasoning about possible executions; exhaustive reasoning will be employed only later on (see \secref{sec:unix-bogosity}), when we shall demonstrate that certain behavior is \<^emph>\not\ possible. \ subsection \Possible transitions\ text \ Rather obviously, a list of system operations can be executed within a certain state if there is a result state reached by an iterated transition. \ definition "can_exec root xs \ (\root'. root \xs\ root')" lemma can_exec_nil: "can_exec root []" unfolding can_exec_def by (blast intro: transitions.intros) lemma can_exec_cons: "root \x\ root' \ can_exec root' xs \ can_exec root (x # xs)" unfolding can_exec_def by (blast intro: transitions.intros) text \ \<^medskip> In case that we already know that a certain sequence can be executed we may destruct it backwardly into individual transitions. \ lemma can_exec_snocD: "can_exec root (xs @ [y]) \ \root' root''. root \xs\ root' \ root' \y\ root''" proof (induct xs arbitrary: root) case Nil then show ?case by (simp add: can_exec_def transitions_nil_eq transitions_cons_eq) next case (Cons x xs) from \can_exec root ((x # xs) @ [y])\ obtain r root'' where x: "root \x\ r" and xs_y: "r \(xs @ [y])\ root''" by (auto simp add: can_exec_def transitions_nil_eq transitions_cons_eq) from xs_y Cons.hyps obtain root' r' where xs: "r \xs\ root'" and y: "root' \y\ r'" unfolding can_exec_def by blast from x xs have "root \(x # xs)\ root'" by (rule transitions.cons) with y show ?case by blast qed subsection \Example executions \label{sec:unix-examples}\ text \ We are now ready to perform a few experiments within our formal model of Unix system-calls. The common technique is to alternate introduction rules of the transition system (see \secref{sec:unix-trans}), and steps to solve any emerging side conditions using algebraic properties of the underlying file-system structures (see \secref{sec:unix-file-system}). \ lemmas eval = access_def init_def theorem "u \ users \ can_exec (init users) [Mkdir u perms [u, name]]" apply (rule can_exec_cons) \ \back-chain \<^term>\can_exec\ (of @{term [source] Cons})\ apply (rule mkdir) \ \back-chain \<^term>\Mkdir\\ apply (force simp add: eval)+ \ \solve preconditions of \<^term>\Mkdir\\ apply (simp add: eval) \ \peek at resulting dir (optional)\ txt \@{subgoals [display]}\ apply (rule can_exec_nil) \ \back-chain \<^term>\can_exec\ (of @{term [source] Nil})\ done text \ By inspecting the result shown just before the final step above, we may gain confidence that our specification of Unix system-calls actually makes sense. Further common errors are usually exhibited when preconditions of transition rules fail unexpectedly. \<^medskip> Here are a few further experiments, using the same techniques as before. \ theorem "u \ users \ can_exec (init users) [Creat u perms [u, name], Unlink u [u, name]]" apply (rule can_exec_cons) apply (rule creat) apply (force simp add: eval)+ apply (simp add: eval) apply (rule can_exec_cons) apply (rule unlink) apply (force simp add: eval)+ apply (simp add: eval) txt \peek at result: @{subgoals [display]}\ apply (rule can_exec_nil) done theorem "u \ users \ Writable \ perms\<^sub>1 \ Readable \ perms\<^sub>2 \ name\<^sub>3 \ name\<^sub>4 \ can_exec (init users) [Mkdir u perms\<^sub>1 [u, name\<^sub>1], Mkdir u' perms\<^sub>2 [u, name\<^sub>1, name\<^sub>2], Creat u' perms\<^sub>3 [u, name\<^sub>1, name\<^sub>2, name\<^sub>3], Creat u' perms\<^sub>3 [u, name\<^sub>1, name\<^sub>2, name\<^sub>4], Readdir u {name\<^sub>3, name\<^sub>4} [u, name\<^sub>1, name\<^sub>2]]" apply (rule can_exec_cons, rule transition.intros, (force simp add: eval)+, (simp add: eval)?)+ txt \peek at result: @{subgoals [display]}\ apply (rule can_exec_nil) done theorem "u \ users \ Writable \ perms\<^sub>1 \ Readable \ perms\<^sub>3 \ can_exec (init users) [Mkdir u perms\<^sub>1 [u, name\<^sub>1], Mkdir u' perms\<^sub>2 [u, name\<^sub>1, name\<^sub>2], Creat u' perms\<^sub>3 [u, name\<^sub>1, name\<^sub>2, name\<^sub>3], Write u' ''foo'' [u, name\<^sub>1, name\<^sub>2, name\<^sub>3], Read u ''foo'' [u, name\<^sub>1, name\<^sub>2, name\<^sub>3]]" apply (rule can_exec_cons, rule transition.intros, (force simp add: eval)+, (simp add: eval)?)+ txt \peek at result: @{subgoals [display]}\ apply (rule can_exec_nil) done section \Odd effects --- treated formally \label{sec:unix-bogosity}\ text \ We are now ready to give a completely formal treatment of the slightly odd situation discussed in the introduction (see \secref{sec:unix-intro}): the file-system can easily reach a state where a user is unable to remove his very own directory, because it is still populated by items placed there by another user in an uncouth manner. \ subsection \The general procedure \label{sec:unix-inv-procedure}\ text \ The following theorem expresses the general procedure we are following to achieve the main result. \ theorem general_procedure: assumes cannot_y: "\r r'. Q r \ r \y\ r' \ False" and init_inv: "\root. init users \bs\ root \ Q root" and preserve_inv: "\r x r'. r \x\ r' \ Q r \ P x \ Q r'" and init_result: "init users \bs\ root" shows "\ (\xs. (\x \ set xs. P x) \ can_exec root (xs @ [y]))" proof - { fix xs assume Ps: "\x \ set xs. P x" assume can_exec: "can_exec root (xs @ [y])" then obtain root' root'' where xs: "root \xs\ root'" and y: "root' \y\ root''" by (blast dest: can_exec_snocD) from init_result have "Q root" by (rule init_inv) from preserve_inv xs this Ps have "Q root'" by (rule transitions_invariant) from this y have False by (rule cannot_y) } then show ?thesis by blast qed text \ Here \<^prop>\P x\ refers to the restriction on file-system operations that are admitted after having reached the critical configuration; according to the problem specification this will become \<^prop>\uid_of x = user\<^sub>1\ later on. Furthermore, \<^term>\y\ refers to the operations we claim to be impossible to perform afterwards, we will take \<^term>\Rmdir\ later. Moreover \<^term>\Q\ is a suitable (auxiliary) invariant over the file-system; choosing \<^term>\Q\ adequately is very important to make the proof work (see \secref{sec:unix-inv-lemmas}). \ subsection \The particular situation\ text \ We introduce a few global declarations and axioms to describe our particular situation considered here. Thus we avoid excessive use of local parameters in the subsequent development. \ context fixes users :: "uid set" and user\<^sub>1 :: uid and user\<^sub>2 :: uid and name\<^sub>1 :: name and name\<^sub>2 :: name and name\<^sub>3 :: name and perms\<^sub>1 :: perms and perms\<^sub>2 :: perms assumes user\<^sub>1_known: "user\<^sub>1 \ users" and user\<^sub>1_not_root: "user\<^sub>1 \ 0" and users_neq: "user\<^sub>1 \ user\<^sub>2" and perms\<^sub>1_writable: "Writable \ perms\<^sub>1" and perms\<^sub>2_not_writable: "Writable \ perms\<^sub>2" notes facts = user\<^sub>1_known user\<^sub>1_not_root users_neq perms\<^sub>1_writable perms\<^sub>2_not_writable begin definition "bogus = [Mkdir user\<^sub>1 perms\<^sub>1 [user\<^sub>1, name\<^sub>1], Mkdir user\<^sub>2 perms\<^sub>2 [user\<^sub>1, name\<^sub>1, name\<^sub>2], Creat user\<^sub>2 perms\<^sub>2 [user\<^sub>1, name\<^sub>1, name\<^sub>2, name\<^sub>3]]" definition "bogus_path = [user\<^sub>1, name\<^sub>1, name\<^sub>2]" text \ The \<^term>\bogus\ operations are the ones that lead into the uncouth situation; \<^term>\bogus_path\ is the key position within the file-system where things go awry. \ subsection \Invariance lemmas \label{sec:unix-inv-lemmas}\ text \ The following invariant over the root file-system describes the bogus situation in an abstract manner: located at a certain \<^term>\path\ within the file-system is a non-empty directory that is neither owned nor writable by \<^term>\user\<^sub>1\. \ definition "invariant root path \ (\att dir. access root path user\<^sub>1 {} = Some (Env att dir) \ dir \ Map.empty \ user\<^sub>1 \ owner att \ access root path user\<^sub>1 {Writable} = None)" text \ Following the general procedure (see \secref{sec:unix-inv-procedure}) we will now establish the three key lemmas required to yield the final result. \<^enum> The invariant is sufficiently strong to entail the pathological case that \<^term>\user\<^sub>1\ is unable to remove the (owned) directory at \<^term>\[user\<^sub>1, name\<^sub>1]\. \<^enum> The invariant does hold after having executed the \<^term>\bogus\ list of operations (starting with an initial file-system configuration). \<^enum> The invariant is preserved by any file-system operation performed by \<^term>\user\<^sub>1\ alone, without any help by other users. As the invariant appears both as assumptions and conclusions in the course of proof, its formulation is rather critical for the whole development to work out properly. In particular, the third step is very sensitive to the invariant being either too strong or too weak. Moreover the invariant has to be sufficiently abstract, lest the proof become cluttered by confusing detail. \<^medskip> The first two lemmas are technically straight forward --- we just have to inspect rather special cases. \ lemma cannot_rmdir: assumes inv: "invariant root bogus_path" and rmdir: "root \(Rmdir user\<^sub>1 [user\<^sub>1, name\<^sub>1])\ root'" shows False proof - from inv obtain "file" where "access root bogus_path user\<^sub>1 {} = Some file" unfolding invariant_def by blast moreover from rmdir obtain att where "access root [user\<^sub>1, name\<^sub>1] user\<^sub>1 {} = Some (Env att Map.empty)" by cases auto then have "access root ([user\<^sub>1, name\<^sub>1] @ [name\<^sub>2]) user\<^sub>1 {} = Map.empty name\<^sub>2" by (simp only: access_empty_lookup lookup_append_some) simp ultimately show False by (simp add: bogus_path_def) qed lemma assumes init: "init users \bogus\ root" shows init_invariant: "invariant root bogus_path" supply eval = facts access_def init_def using init apply (unfold bogus_def bogus_path_def) apply (drule transitions_consD, rule transition.intros, (force simp add: eval)+, (simp add: eval)?)+ \ \evaluate all operations\ apply (drule transitions_nilD) \ \reach final result\ apply (simp add: invariant_def eval) \ \check the invariant\ done text \ \<^medskip> At last we are left with the main effort to show that the ``bogosity'' invariant is preserved by any file-system operation performed by \<^term>\user\<^sub>1\ alone. Note that this holds for any \<^term>\path\ given, the particular \<^term>\bogus_path\ is not required here. \ lemma preserve_invariant: assumes tr: "root \x\ root'" and inv: "invariant root path" and uid: "uid_of x = user\<^sub>1" shows "invariant root' path" proof - from inv obtain att dir where inv1: "access root path user\<^sub>1 {} = Some (Env att dir)" and inv2: "dir \ Map.empty" and inv3: "user\<^sub>1 \ owner att" and inv4: "access root path user\<^sub>1 {Writable} = None" by (auto simp add: invariant_def) from inv1 have lookup: "lookup root path = Some (Env att dir)" by (simp only: access_empty_lookup) from inv1 inv3 inv4 and user\<^sub>1_not_root have not_writable: "Writable \ others att" by (auto simp add: access_def split: option.splits) show ?thesis proof cases assume "root' = root" with inv show "invariant root' path" by (simp only:) next assume changed: "root' \ root" with tr obtain opt where root': "root' = update (path_of x) opt root" by cases auto show ?thesis proof (rule prefix_cases) assume "path_of x \ path" with inv root' have "\perms. access root' path user\<^sub>1 perms = access root path user\<^sub>1 perms" by (simp only: access_update_other) with inv show "invariant root' path" by (simp only: invariant_def) next assume "prefix (path_of x) path" then obtain ys where path: "path = path_of x @ ys" .. show ?thesis proof (cases ys) assume "ys = []" with tr uid inv2 inv3 lookup changed path and user\<^sub>1_not_root have False by cases (auto simp add: access_empty_lookup dest: access_some_lookup) then show ?thesis .. next fix z zs assume ys: "ys = z # zs" have "lookup root' path = lookup root path" proof - from inv2 lookup path ys have look: "lookup root (path_of x @ z # zs) = Some (Env att dir)" by (simp only:) then obtain att' dir' file' where look': "lookup root (path_of x) = Some (Env att' dir')" and dir': "dir' z = Some file'" and file': "lookup file' zs = Some (Env att dir)" by (blast dest: lookup_some_upper) from tr uid changed look' dir' obtain att'' where look'': "lookup root' (path_of x) = Some (Env att'' dir')" by cases (auto simp add: access_empty_lookup lookup_update_some dest: access_some_lookup) with dir' file' have "lookup root' (path_of x @ z # zs) = Some (Env att dir)" by (simp add: lookup_append_some) with look path ys show ?thesis by simp qed with inv show "invariant root' path" by (simp only: invariant_def access_def) qed next assume "strict_prefix path (path_of x)" then obtain y ys where path: "path_of x = path @ y # ys" .. obtain dir' where lookup': "lookup root' path = Some (Env att dir')" and inv2': "dir' \ Map.empty" proof (cases ys) assume "ys = []" with path have parent: "path_of x = path @ [y]" by simp with tr uid inv4 changed obtain "file" where "root' = update (path_of x) (Some file) root" by cases auto with lookup parent have "lookup root' path = Some (Env att (dir(y\file)))" by (simp only: update_append_some update_cons_nil_env) moreover have "dir(y\file) \ Map.empty" by simp ultimately show ?thesis .. next fix z zs assume ys: "ys = z # zs" with lookup root' path have "lookup root' path = Some (update (y # ys) opt (Env att dir))" by (simp only: update_append_some) also obtain file' where "update (y # ys) opt (Env att dir) = Env att (dir(y\file'))" proof - have "dir y \ None" proof - have "dir y = lookup (Env att dir) [y]" by (simp split: option.splits) also from lookup have "\ = lookup root (path @ [y])" by (simp only: lookup_append_some) also have "\ \ None" proof - from ys obtain us u where rev_ys: "ys = us @ [u]" by (cases ys rule: rev_cases) simp with tr path have "lookup root ((path @ [y]) @ (us @ [u])) \ None \ lookup root ((path @ [y]) @ us) \ None" by cases (auto dest: access_some_lookup) then show ?thesis by (fastforce dest!: lookup_some_append) qed finally show ?thesis . qed with ys show ?thesis using that by auto qed also have "dir(y\file') \ Map.empty" by simp ultimately show ?thesis .. qed from lookup' have inv1': "access root' path user\<^sub>1 {} = Some (Env att dir')" by (simp only: access_empty_lookup) from inv3 lookup' and not_writable user\<^sub>1_not_root have "access root' path user\<^sub>1 {Writable} = None" by (simp add: access_def) with inv1' inv2' inv3 show ?thesis unfolding invariant_def by blast qed qed qed subsection \Putting it all together \label{sec:unix-main-result}\ text \ The main result is now imminent, just by composing the three invariance lemmas (see \secref{sec:unix-inv-lemmas}) according the the overall procedure (see \secref{sec:unix-inv-procedure}). \ corollary assumes bogus: "init users \bogus\ root" shows "\ (\xs. (\x \ set xs. uid_of x = user\<^sub>1) \ can_exec root (xs @ [Rmdir user\<^sub>1 [user\<^sub>1, name\<^sub>1]]))" proof - from cannot_rmdir init_invariant preserve_invariant and bogus show ?thesis by (rule general_procedure) qed end end diff --git a/src/HOL/ex/CTL.thy b/src/HOL/ex/CTL.thy --- a/src/HOL/ex/CTL.thy +++ b/src/HOL/ex/CTL.thy @@ -1,312 +1,311 @@ (* Title: HOL/ex/CTL.thy Author: Gertrud Bauer *) section \CTL formulae\ theory CTL imports Main begin text \ - We formalize basic concepts of Computational Tree Logic (CTL) @{cite - "McMillan-PhDThesis" and "McMillan-LectureNotes"} within the simply-typed + We formalize basic concepts of Computational Tree Logic (CTL) \<^cite>\"McMillan-PhDThesis" and "McMillan-LectureNotes"\ within the simply-typed set theory of HOL. By using the common technique of ``shallow embedding'', a CTL formula is identified with the corresponding set of states where it holds. Consequently, CTL operations such as negation, conjunction, disjunction simply become complement, intersection, union of sets. We only require a separate operation for implication, as point-wise inclusion is usually not encountered in plain set-theory. \ lemmas [intro!] = Int_greatest Un_upper2 Un_upper1 Int_lower1 Int_lower2 type_synonym 'a ctl = "'a set" definition imp :: "'a ctl \ 'a ctl \ 'a ctl" (infixr "\" 75) where "p \ q = - p \ q" lemma [intro!]: "p \ p \ q \ q" unfolding imp_def by auto lemma [intro!]: "p \ (q \ p)" unfolding imp_def by rule text \ \<^smallskip> The CTL path operators are more interesting; they are based on an arbitrary, but fixed model \\\, which is simply a transition relation over states \<^typ>\'a\. \ axiomatization \ :: "('a \ 'a) set" text \ The operators \\<^bold>E\<^bold>X\, \\<^bold>E\<^bold>F\, \\<^bold>E\<^bold>G\ are taken as primitives, while \\<^bold>A\<^bold>X\, \\<^bold>A\<^bold>F\, \\<^bold>A\<^bold>G\ are defined as derived ones. The formula \\<^bold>E\<^bold>X p\ holds in a state \s\, iff there is a successor state \s'\ (with respect to the model \\\), such that \p\ holds in \s'\. The formula \\<^bold>E\<^bold>F p\ holds in a state \s\, iff there is a path in \\\, starting from \s\, such that there exists a state \s'\ on the path, such that \p\ holds in \s'\. The formula \\<^bold>E\<^bold>G p\ holds in a state \s\, iff there is a path, starting from \s\, such that for all states \s'\ on the path, \p\ holds in \s'\. It is easy to see that \\<^bold>E\<^bold>F p\ and \\<^bold>E\<^bold>G p\ may be expressed using least and greatest fixed points - @{cite "McMillan-PhDThesis"}. + \<^cite>\"McMillan-PhDThesis"\. \ definition EX ("\<^bold>E\<^bold>X _" [80] 90) where [simp]: "\<^bold>E\<^bold>X p = {s. \s'. (s, s') \ \ \ s' \ p}" definition EF ("\<^bold>E\<^bold>F _" [80] 90) where [simp]: "\<^bold>E\<^bold>F p = lfp (\s. p \ \<^bold>E\<^bold>X s)" definition EG ("\<^bold>E\<^bold>G _" [80] 90) where [simp]: "\<^bold>E\<^bold>G p = gfp (\s. p \ \<^bold>E\<^bold>X s)" text \ \\<^bold>A\<^bold>X\, \\<^bold>A\<^bold>F\ and \\<^bold>A\<^bold>G\ are now defined dually in terms of \\<^bold>E\<^bold>X\, \\<^bold>E\<^bold>F\ and \\<^bold>E\<^bold>G\. \ definition AX ("\<^bold>A\<^bold>X _" [80] 90) where [simp]: "\<^bold>A\<^bold>X p = - \<^bold>E\<^bold>X - p" definition AF ("\<^bold>A\<^bold>F _" [80] 90) where [simp]: "\<^bold>A\<^bold>F p = - \<^bold>E\<^bold>G - p" definition AG ("\<^bold>A\<^bold>G _" [80] 90) where [simp]: "\<^bold>A\<^bold>G p = - \<^bold>E\<^bold>F - p" subsection \Basic fixed point properties\ text \ First of all, we use the de-Morgan property of fixed points. \ lemma lfp_gfp: "lfp f = - gfp (\s::'a set. - (f (- s)))" proof show "lfp f \ - gfp (\s. - f (- s))" proof show "x \ - gfp (\s. - f (- s))" if l: "x \ lfp f" for x proof assume "x \ gfp (\s. - f (- s))" then obtain u where "x \ u" and "u \ - f (- u)" by (auto simp add: gfp_def) then have "f (- u) \ - u" by auto then have "lfp f \ - u" by (rule lfp_lowerbound) from l and this have "x \ u" by auto with \x \ u\ show False by contradiction qed qed show "- gfp (\s. - f (- s)) \ lfp f" proof (rule lfp_greatest) fix u assume "f u \ u" then have "- u \ - f u" by auto then have "- u \ - f (- (- u))" by simp then have "- u \ gfp (\s. - f (- s))" by (rule gfp_upperbound) then show "- gfp (\s. - f (- s)) \ u" by auto qed qed lemma lfp_gfp': "- lfp f = gfp (\s::'a set. - (f (- s)))" by (simp add: lfp_gfp) lemma gfp_lfp': "- gfp f = lfp (\s::'a set. - (f (- s)))" by (simp add: lfp_gfp) text \ In order to give dual fixed point representations of \<^term>\\<^bold>A\<^bold>F p\ and \<^term>\\<^bold>A\<^bold>G p\: \ lemma AF_lfp: "\<^bold>A\<^bold>F p = lfp (\s. p \ \<^bold>A\<^bold>X s)" by (simp add: lfp_gfp) lemma AG_gfp: "\<^bold>A\<^bold>G p = gfp (\s. p \ \<^bold>A\<^bold>X s)" by (simp add: lfp_gfp) lemma EF_fp: "\<^bold>E\<^bold>F p = p \ \<^bold>E\<^bold>X \<^bold>E\<^bold>F p" proof - have "mono (\s. p \ \<^bold>E\<^bold>X s)" by rule auto then show ?thesis by (simp only: EF_def) (rule lfp_unfold) qed lemma AF_fp: "\<^bold>A\<^bold>F p = p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>F p" proof - have "mono (\s. p \ \<^bold>A\<^bold>X s)" by rule auto then show ?thesis by (simp only: AF_lfp) (rule lfp_unfold) qed lemma EG_fp: "\<^bold>E\<^bold>G p = p \ \<^bold>E\<^bold>X \<^bold>E\<^bold>G p" proof - have "mono (\s. p \ \<^bold>E\<^bold>X s)" by rule auto then show ?thesis by (simp only: EG_def) (rule gfp_unfold) qed text \ From the greatest fixed point definition of \<^term>\\<^bold>A\<^bold>G p\, we derive as a consequence of the Knaster-Tarski theorem on the one hand that \<^term>\\<^bold>A\<^bold>G p\ is a fixed point of the monotonic function \<^term>\\s. p \ \<^bold>A\<^bold>X s\. \ lemma AG_fp: "\<^bold>A\<^bold>G p = p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" proof - have "mono (\s. p \ \<^bold>A\<^bold>X s)" by rule auto then show ?thesis by (simp only: AG_gfp) (rule gfp_unfold) qed text \ This fact may be split up into two inequalities (merely using transitivity of \\\, which is an instance of the overloaded \\\ in Isabelle/HOL). \ lemma AG_fp_1: "\<^bold>A\<^bold>G p \ p" proof - note AG_fp also have "p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G p \ p" by auto finally show ?thesis . qed lemma AG_fp_2: "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" proof - note AG_fp also have "p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by auto finally show ?thesis . qed text \ On the other hand, we have from the Knaster-Tarski fixed point theorem that any other post-fixed point of \<^term>\\s. p \ \<^bold>A\<^bold>X s\ is smaller than \<^term>\\<^bold>A\<^bold>G p\. A post-fixed point is a set of states \q\ such that \<^term>\q \ p \ \<^bold>A\<^bold>X q\. This leads to the following co-induction principle for \<^term>\\<^bold>A\<^bold>G p\. \ lemma AG_I: "q \ p \ \<^bold>A\<^bold>X q \ q \ \<^bold>A\<^bold>G p" by (simp only: AG_gfp) (rule gfp_upperbound) subsection \The tree induction principle \label{sec:calc-ctl-tree-induct}\ text \ With the most basic facts available, we are now able to establish a few more interesting results, leading to the \<^emph>\tree induction\ principle for \\<^bold>A\<^bold>G\ (see below). We will use some elementary monotonicity and distributivity rules. \ lemma AX_int: "\<^bold>A\<^bold>X (p \ q) = \<^bold>A\<^bold>X p \ \<^bold>A\<^bold>X q" by auto lemma AX_mono: "p \ q \ \<^bold>A\<^bold>X p \ \<^bold>A\<^bold>X q" by auto lemma AG_mono: "p \ q \ \<^bold>A\<^bold>G p \ \<^bold>A\<^bold>G q" by (simp only: AG_gfp, rule gfp_mono) auto text \ The formula \<^term>\AG p\ implies \<^term>\AX p\ (we use substitution of \\\ with monotonicity). \ lemma AG_AX: "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>X p" proof - have "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by (rule AG_fp_2) also have "\<^bold>A\<^bold>G p \ p" by (rule AG_fp_1) moreover note AX_mono finally show ?thesis . qed text \ Furthermore we show idempotency of the \\<^bold>A\<^bold>G\ operator. The proof is a good example of how accumulated facts may get used to feed a single rule step. \ lemma AG_AG: "\<^bold>A\<^bold>G \<^bold>A\<^bold>G p = \<^bold>A\<^bold>G p" proof show "\<^bold>A\<^bold>G \<^bold>A\<^bold>G p \ \<^bold>A\<^bold>G p" by (rule AG_fp_1) next show "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>G \<^bold>A\<^bold>G p" proof (rule AG_I) have "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>G p" .. moreover have "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" by (rule AG_fp_2) ultimately show "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>G p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" .. qed qed text \ \<^smallskip> We now give an alternative characterization of the \\<^bold>A\<^bold>G\ operator, which describes the \\<^bold>A\<^bold>G\ operator in an ``operational'' way by tree induction: In a state holds \<^term>\AG p\ iff in that state holds \p\, and in all reachable states \s\ follows from the fact that \p\ holds in \s\, that \p\ also holds in all successor states of \s\. We use the co-induction principle @{thm [source] AG_I} to establish this in a purely algebraic manner. \ theorem AG_induct: "p \ \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p) = \<^bold>A\<^bold>G p" proof show "p \ \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p) \ \<^bold>A\<^bold>G p" (is "?lhs \ _") proof (rule AG_I) show "?lhs \ p \ \<^bold>A\<^bold>X ?lhs" proof show "?lhs \ p" .. show "?lhs \ \<^bold>A\<^bold>X ?lhs" proof - { have "\<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p) \ p \ \<^bold>A\<^bold>X p" by (rule AG_fp_1) also have "p \ p \ \<^bold>A\<^bold>X p \ \<^bold>A\<^bold>X p" .. finally have "?lhs \ \<^bold>A\<^bold>X p" by auto } moreover { have "p \ \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p) \ \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p)" .. also have "\ \ \<^bold>A\<^bold>X \" by (rule AG_fp_2) finally have "?lhs \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p)" . } ultimately have "?lhs \ \<^bold>A\<^bold>X p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p)" .. also have "\ = \<^bold>A\<^bold>X ?lhs" by (simp only: AX_int) finally show ?thesis . qed qed qed next show "\<^bold>A\<^bold>G p \ p \ \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p)" proof show "\<^bold>A\<^bold>G p \ p" by (rule AG_fp_1) show "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p)" proof - have "\<^bold>A\<^bold>G p = \<^bold>A\<^bold>G \<^bold>A\<^bold>G p" by (simp only: AG_AG) also have "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>X p" by (rule AG_AX) moreover note AG_mono also have "\<^bold>A\<^bold>X p \ (p \ \<^bold>A\<^bold>X p)" .. moreover note AG_mono finally show ?thesis . qed qed qed subsection \An application of tree induction \label{sec:calc-ctl-commute}\ text \ Further interesting properties of CTL expressions may be demonstrated with the help of tree induction; here we show that \\<^bold>A\<^bold>X\ and \\<^bold>A\<^bold>G\ commute. \ theorem AG_AX_commute: "\<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>X \<^bold>A\<^bold>G p" proof - have "\<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>X p \ \<^bold>A\<^bold>X \<^bold>A\<^bold>G \<^bold>A\<^bold>X p" by (rule AG_fp) also have "\ = \<^bold>A\<^bold>X (p \ \<^bold>A\<^bold>G \<^bold>A\<^bold>X p)" by (simp only: AX_int) also have "p \ \<^bold>A\<^bold>G \<^bold>A\<^bold>X p = \<^bold>A\<^bold>G p" (is "?lhs = _") proof have "\<^bold>A\<^bold>X p \ p \ \<^bold>A\<^bold>X p" .. also have "p \ \<^bold>A\<^bold>G (p \ \<^bold>A\<^bold>X p) = \<^bold>A\<^bold>G p" by (rule AG_induct) also note Int_mono AG_mono ultimately show "?lhs \ \<^bold>A\<^bold>G p" by fast next have "\<^bold>A\<^bold>G p \ p" by (rule AG_fp_1) moreover { have "\<^bold>A\<^bold>G p = \<^bold>A\<^bold>G \<^bold>A\<^bold>G p" by (simp only: AG_AG) also have "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>X p" by (rule AG_AX) also note AG_mono ultimately have "\<^bold>A\<^bold>G p \ \<^bold>A\<^bold>G \<^bold>A\<^bold>X p" . } ultimately show "\<^bold>A\<^bold>G p \ ?lhs" .. qed finally show ?thesis . qed end diff --git a/src/Pure/Examples/Higher_Order_Logic.thy b/src/Pure/Examples/Higher_Order_Logic.thy --- a/src/Pure/Examples/Higher_Order_Logic.thy +++ b/src/Pure/Examples/Higher_Order_Logic.thy @@ -1,520 +1,519 @@ (* Title: Pure/Examples/Higher_Order_Logic.thy Author: Makarius *) section \Foundations of HOL\ theory Higher_Order_Logic imports Pure begin text \ The following theory development illustrates the foundations of Higher-Order - Logic. The ``HOL'' logic that is given here resembles @{cite - "Gordon:1985:HOL"} and its predecessor @{cite "church40"}, but the order of + Logic. The ``HOL'' logic that is given here resembles \<^cite>\"Gordon:1985:HOL"\ and its predecessor \<^cite>\"church40"\, but the order of axiomatizations and defined connectives has be adapted to modern presentations of \\\-calculus and Constructive Type Theory. Thus it fits nicely to the underlying Natural Deduction framework of Isabelle/Pure and Isabelle/Isar. \ section \HOL syntax within Pure\ class type default_sort type typedecl o instance o :: type .. instance "fun" :: (type, type) type .. judgment Trueprop :: "o \ prop" ("_" 5) section \Minimal logic (axiomatization)\ axiomatization imp :: "o \ o \ o" (infixr "\" 25) where impI [intro]: "(A \ B) \ A \ B" and impE [dest, trans]: "A \ B \ A \ B" axiomatization All :: "('a \ o) \ o" (binder "\" 10) where allI [intro]: "(\x. P x) \ \x. P x" and allE [dest]: "\x. P x \ P a" lemma atomize_imp [atomize]: "(A \ B) \ Trueprop (A \ B)" by standard (fact impI, fact impE) lemma atomize_all [atomize]: "(\x. P x) \ Trueprop (\x. P x)" by standard (fact allI, fact allE) subsubsection \Derived connectives\ definition False :: o where "False \ \A. A" lemma FalseE [elim]: assumes "False" shows A proof - from \False\ have "\A. A" by (simp only: False_def) then show A .. qed definition True :: o where "True \ False \ False" lemma TrueI [intro]: True unfolding True_def .. definition not :: "o \ o" ("\ _" [40] 40) where "not \ \A. A \ False" lemma notI [intro]: assumes "A \ False" shows "\ A" using assms unfolding not_def .. lemma notE [elim]: assumes "\ A" and A shows B proof - from \\ A\ have "A \ False" by (simp only: not_def) from this and \A\ have "False" .. then show B .. qed lemma notE': "A \ \ A \ B" by (rule notE) lemmas contradiction = notE notE' \ \proof by contradiction in any order\ definition conj :: "o \ o \ o" (infixr "\" 35) where "A \ B \ \C. (A \ B \ C) \ C" lemma conjI [intro]: assumes A and B shows "A \ B" unfolding conj_def proof fix C show "(A \ B \ C) \ C" proof assume "A \ B \ C" also note \A\ also note \B\ finally show C . qed qed lemma conjE [elim]: assumes "A \ B" obtains A and B proof from \A \ B\ have *: "(A \ B \ C) \ C" for C unfolding conj_def .. show A proof - note * [of A] also have "A \ B \ A" proof assume A then show "B \ A" .. qed finally show ?thesis . qed show B proof - note * [of B] also have "A \ B \ B" proof show "B \ B" .. qed finally show ?thesis . qed qed definition disj :: "o \ o \ o" (infixr "\" 30) where "A \ B \ \C. (A \ C) \ (B \ C) \ C" lemma disjI1 [intro]: assumes A shows "A \ B" unfolding disj_def proof fix C show "(A \ C) \ (B \ C) \ C" proof assume "A \ C" from this and \A\ have C .. then show "(B \ C) \ C" .. qed qed lemma disjI2 [intro]: assumes B shows "A \ B" unfolding disj_def proof fix C show "(A \ C) \ (B \ C) \ C" proof show "(B \ C) \ C" proof assume "B \ C" from this and \B\ show C .. qed qed qed lemma disjE [elim]: assumes "A \ B" obtains (a) A | (b) B proof - from \A \ B\ have "(A \ thesis) \ (B \ thesis) \ thesis" unfolding disj_def .. also have "A \ thesis" proof assume A then show thesis by (rule a) qed also have "B \ thesis" proof assume B then show thesis by (rule b) qed finally show thesis . qed definition Ex :: "('a \ o) \ o" (binder "\" 10) where "\x. P x \ \C. (\x. P x \ C) \ C" lemma exI [intro]: "P a \ \x. P x" unfolding Ex_def proof fix C assume "P a" show "(\x. P x \ C) \ C" proof assume "\x. P x \ C" then have "P a \ C" .. from this and \P a\ show C .. qed qed lemma exE [elim]: assumes "\x. P x" obtains (that) x where "P x" proof - from \\x. P x\ have "(\x. P x \ thesis) \ thesis" unfolding Ex_def .. also have "\x. P x \ thesis" proof fix x show "P x \ thesis" proof assume "P x" then show thesis by (rule that) qed qed finally show thesis . qed subsubsection \Extensional equality\ axiomatization equal :: "'a \ 'a \ o" (infixl "=" 50) where refl [intro]: "x = x" and subst: "x = y \ P x \ P y" abbreviation not_equal :: "'a \ 'a \ o" (infixl "\" 50) where "x \ y \ \ (x = y)" abbreviation iff :: "o \ o \ o" (infixr "\" 25) where "A \ B \ A = B" axiomatization where ext [intro]: "(\x. f x = g x) \ f = g" and iff [intro]: "(A \ B) \ (B \ A) \ A \ B" for f g :: "'a \ 'b" lemma sym [sym]: "y = x" if "x = y" using that by (rule subst) (rule refl) lemma [trans]: "x = y \ P y \ P x" by (rule subst) (rule sym) lemma [trans]: "P x \ x = y \ P y" by (rule subst) lemma arg_cong: "f x = f y" if "x = y" using that by (rule subst) (rule refl) lemma fun_cong: "f x = g x" if "f = g" using that by (rule subst) (rule refl) lemma trans [trans]: "x = y \ y = z \ x = z" by (rule subst) lemma iff1 [elim]: "A \ B \ A \ B" by (rule subst) lemma iff2 [elim]: "A \ B \ B \ A" by (rule subst) (rule sym) subsection \Cantor's Theorem\ text \ Cantor's Theorem states that there is no surjection from a set to its powerset. The subsequent formulation uses elementary \\\-calculus and predicate logic, with standard introduction and elimination rules. \ lemma iff_contradiction: assumes *: "\ A \ A" shows C proof (rule notE) show "\ A" proof assume A with * have "\ A" .. from this and \A\ show False .. qed with * show A .. qed theorem Cantor: "\ (\f :: 'a \ 'a \ o. \A. \x. A = f x)" proof assume "\f :: 'a \ 'a \ o. \A. \x. A = f x" then obtain f :: "'a \ 'a \ o" where *: "\A. \x. A = f x" .. let ?D = "\x. \ f x x" from * have "\x. ?D = f x" .. then obtain a where "?D = f a" .. then have "?D a \ f a a" using refl by (rule subst) then have "\ f a a \ f a a" . then show False by (rule iff_contradiction) qed subsection \Characterization of Classical Logic\ text \ The subsequent rules of classical reasoning are all equivalent. \ locale classical = assumes classical: "(\ A \ A) \ A" \ \predicate definition and hypothetical context\ begin lemma classical_contradiction: assumes "\ A \ False" shows A proof (rule classical) assume "\ A" then have False by (rule assms) then show A .. qed lemma double_negation: assumes "\ \ A" shows A proof (rule classical_contradiction) assume "\ A" with \\ \ A\ show False by (rule contradiction) qed lemma tertium_non_datur: "A \ \ A" proof (rule double_negation) show "\ \ (A \ \ A)" proof assume "\ (A \ \ A)" have "\ A" proof assume A then have "A \ \ A" .. with \\ (A \ \ A)\ show False by (rule contradiction) qed then have "A \ \ A" .. with \\ (A \ \ A)\ show False by (rule contradiction) qed qed lemma classical_cases: obtains A | "\ A" using tertium_non_datur proof assume A then show thesis .. next assume "\ A" then show thesis .. qed end lemma classical_if_cases: classical if cases: "\A C. (A \ C) \ (\ A \ C) \ C" proof fix A assume *: "\ A \ A" show A proof (rule cases) assume A then show A . next assume "\ A" then show A by (rule *) qed qed section \Peirce's Law\ text \ Peirce's Law is another characterization of classical reasoning. Its statement only requires implication. \ theorem (in classical) Peirce's_Law: "((A \ B) \ A) \ A" proof assume *: "(A \ B) \ A" show A proof (rule classical) assume "\ A" have "A \ B" proof assume A with \\ A\ show B by (rule contradiction) qed with * show A .. qed qed section \Hilbert's choice operator (axiomatization)\ axiomatization Eps :: "('a \ o) \ 'a" where someI: "P x \ P (Eps P)" syntax "_Eps" :: "pttrn \ o \ 'a" ("(3SOME _./ _)" [0, 10] 10) translations "SOME x. P" \ "CONST Eps (\x. P)" text \ \<^medskip> It follows a derivation of the classical law of tertium-non-datur by means of Hilbert's choice operator (due to Berghofer, Beeson, Harrison, based on a proof by Diaconescu). \<^medskip> \ theorem Diaconescu: "A \ \ A" proof - let ?P = "\x. (A \ x) \ \ x" let ?Q = "\x. (A \ \ x) \ x" have a: "?P (Eps ?P)" proof (rule someI) have "\ False" .. then show "?P False" .. qed have b: "?Q (Eps ?Q)" proof (rule someI) have True .. then show "?Q True" .. qed from a show ?thesis proof assume "A \ Eps ?P" then have A .. then show ?thesis .. next assume "\ Eps ?P" from b show ?thesis proof assume "A \ \ Eps ?Q" then have A .. then show ?thesis .. next assume "Eps ?Q" have neq: "?P \ ?Q" proof assume "?P = ?Q" then have "Eps ?P \ Eps ?Q" by (rule arg_cong) also note \Eps ?Q\ finally have "Eps ?P" . with \\ Eps ?P\ show False by (rule contradiction) qed have "\ A" proof assume A have "?P = ?Q" proof (rule ext) show "?P x \ ?Q x" for x proof assume "?P x" then show "?Q x" proof assume "\ x" with \A\ have "A \ \ x" .. then show ?thesis .. next assume "A \ x" then have x .. then show ?thesis .. qed next assume "?Q x" then show "?P x" proof assume "A \ \ x" then have "\ x" .. then show ?thesis .. next assume x with \A\ have "A \ x" .. then show ?thesis .. qed qed qed with neq show False by (rule contradiction) qed then show ?thesis .. qed qed qed text \ This means, the hypothetical predicate \<^const>\classical\ always holds unconditionally (with all consequences). \ interpretation classical proof (rule classical_if_cases) fix A C assume *: "A \ C" and **: "\ A \ C" from Diaconescu [of A] show C proof assume A then show C by (rule *) next assume "\ A" then show C by (rule **) qed qed thm classical classical_contradiction double_negation tertium_non_datur classical_cases Peirce's_Law end diff --git a/src/ZF/Induct/Acc.thy b/src/ZF/Induct/Acc.thy --- a/src/ZF/Induct/Acc.thy +++ b/src/ZF/Induct/Acc.thy @@ -1,61 +1,61 @@ (* Title: ZF/Induct/Acc.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \The accessible part of a relation\ theory Acc imports ZF begin text \ - Inductive definition of \acc(r)\; see @{cite "paulin-tlca"}. + Inductive definition of \acc(r)\; see \<^cite>\"paulin-tlca"\. \ consts acc :: "i \ i" inductive domains "acc(r)" \ "field(r)" intros vimage: "\r-``{a}: Pow(acc(r)); a \ field(r)\ \ a \ acc(r)" monos Pow_mono text \ The introduction rule must require \<^prop>\a \ field(r)\, otherwise \acc(r)\ would be a proper class! \medskip The intended introduction rule: \ lemma accI: "\\b. \b,a\:r \ b \ acc(r); a \ field(r)\ \ a \ acc(r)" by (blast intro: acc.intros) lemma acc_downward: "\b \ acc(r); \a,b\: r\ \ a \ acc(r)" by (erule acc.cases) blast lemma acc_induct [consumes 1, case_names vimage, induct set: acc]: "\a \ acc(r); \x. \x \ acc(r); \y. \y,x\:r \ P(y)\ \ P(x) \ \ P(a)" by (erule acc.induct) (blast intro: acc.intros) lemma wf_on_acc: "wf[acc(r)](r)" apply (rule wf_onI2) apply (erule acc_induct) apply fast done lemma acc_wfI: "field(r) \ acc(r) \ wf(r)" by (erule wf_on_acc [THEN wf_on_subset_A, THEN wf_on_field_imp_wf]) lemma acc_wfD: "wf(r) \ field(r) \ acc(r)" apply (rule subsetI) apply (erule wf_induct2, assumption) apply (blast intro: accI)+ done lemma wf_acc_iff: "wf(r) \ field(r) \ acc(r)" by (rule iffI, erule acc_wfD, erule acc_wfI) end diff --git a/src/ZF/Induct/Comb.thy b/src/ZF/Induct/Comb.thy --- a/src/ZF/Induct/Comb.thy +++ b/src/ZF/Induct/Comb.thy @@ -1,273 +1,273 @@ (* Title: ZF/Induct/Comb.thy Author: Lawrence C Paulson Copyright 1994 University of Cambridge *) section \Combinatory Logic example: the Church-Rosser Theorem\ theory Comb imports ZF begin text \ Curiously, combinators do not include free variables. - Example taken from @{cite camilleri92}. + Example taken from \<^cite>\camilleri92\. \ subsection \Definitions\ text \Datatype definition of combinators \S\ and \K\.\ consts comb :: i datatype comb = K | S | app ("p \ comb", "q \ comb") (infixl \\\ 90) text \ Inductive definition of contractions, \\\<^sup>1\ and (multi-step) reductions, \\\. \ consts contract :: i abbreviation contract_syntax :: "[i,i] \ o" (infixl \\\<^sup>1\ 50) where "p \\<^sup>1 q \ \p,q\ \ contract" abbreviation contract_multi :: "[i,i] \ o" (infixl \\\ 50) where "p \ q \ \p,q\ \ contract^*" inductive domains "contract" \ "comb \ comb" intros K: "\p \ comb; q \ comb\ \ K\p\q \\<^sup>1 p" S: "\p \ comb; q \ comb; r \ comb\ \ S\p\q\r \\<^sup>1 (p\r)\(q\r)" Ap1: "\p\\<^sup>1q; r \ comb\ \ p\r \\<^sup>1 q\r" Ap2: "\p\\<^sup>1q; r \ comb\ \ r\p \\<^sup>1 r\q" type_intros comb.intros text \ Inductive definition of parallel contractions, \\\<^sup>1\ and (multi-step) parallel reductions, \\\. \ consts parcontract :: i abbreviation parcontract_syntax :: "[i,i] \ o" (infixl \\\<^sup>1\ 50) where "p \\<^sup>1 q \ \p,q\ \ parcontract" abbreviation parcontract_multi :: "[i,i] \ o" (infixl \\\ 50) where "p \ q \ \p,q\ \ parcontract^+" inductive domains "parcontract" \ "comb \ comb" intros refl: "\p \ comb\ \ p \\<^sup>1 p" K: "\p \ comb; q \ comb\ \ K\p\q \\<^sup>1 p" S: "\p \ comb; q \ comb; r \ comb\ \ S\p\q\r \\<^sup>1 (p\r)\(q\r)" Ap: "\p\\<^sup>1q; r\\<^sup>1s\ \ p\r \\<^sup>1 q\s" type_intros comb.intros text \ Misc definitions. \ definition I :: i where "I \ S\K\K" definition diamond :: "i \ o" where "diamond(r) \ \x y. \x,y\\r \ (\y'. \r \ (\z. \y,z\\r \ \ r))" subsection \Transitive closure preserves the Church-Rosser property\ lemma diamond_strip_lemmaD [rule_format]: "\diamond(r); \x,y\:r^+\ \ \y'. :r \ (\z. : r^+ \ \y,z\: r)" unfolding diamond_def apply (erule trancl_induct) apply (blast intro: r_into_trancl) apply clarify apply (drule spec [THEN mp], assumption) apply (blast intro: r_into_trancl trans_trancl [THEN transD]) done lemma diamond_trancl: "diamond(r) \ diamond(r^+)" apply (simp (no_asm_simp) add: diamond_def) apply (rule impI [THEN allI, THEN allI]) apply (erule trancl_induct) apply auto apply (best intro: r_into_trancl trans_trancl [THEN transD] dest: diamond_strip_lemmaD)+ done inductive_cases Ap_E [elim!]: "p\q \ comb" subsection \Results about Contraction\ text \ For type checking: replaces \<^term>\a \\<^sup>1 b\ by \a, b \ comb\. \ lemmas contract_combE2 = contract.dom_subset [THEN subsetD, THEN SigmaE2] and contract_combD1 = contract.dom_subset [THEN subsetD, THEN SigmaD1] and contract_combD2 = contract.dom_subset [THEN subsetD, THEN SigmaD2] lemma field_contract_eq: "field(contract) = comb" by (blast intro: contract.K elim!: contract_combE2) lemmas reduction_refl = field_contract_eq [THEN equalityD2, THEN subsetD, THEN rtrancl_refl] lemmas rtrancl_into_rtrancl2 = r_into_rtrancl [THEN trans_rtrancl [THEN transD]] declare reduction_refl [intro!] contract.K [intro!] contract.S [intro!] lemmas reduction_rls = contract.K [THEN rtrancl_into_rtrancl2] contract.S [THEN rtrancl_into_rtrancl2] contract.Ap1 [THEN rtrancl_into_rtrancl2] contract.Ap2 [THEN rtrancl_into_rtrancl2] lemma "p \ comb \ I\p \ p" \ \Example only: not used\ unfolding I_def by (blast intro: reduction_rls) lemma comb_I: "I \ comb" unfolding I_def by blast subsection \Non-contraction results\ text \Derive a case for each combinator constructor.\ inductive_cases K_contractE [elim!]: "K \\<^sup>1 r" and S_contractE [elim!]: "S \\<^sup>1 r" and Ap_contractE [elim!]: "p\q \\<^sup>1 r" lemma I_contract_E: "I \\<^sup>1 r \ P" by (auto simp add: I_def) lemma K1_contractD: "K\p \\<^sup>1 r \ (\q. r = K\q \ p \\<^sup>1 q)" by auto lemma Ap_reduce1: "\p \ q; r \ comb\ \ p\r \ q\r" apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) apply (erule rtrancl_induct) apply (blast intro: reduction_rls) apply (erule trans_rtrancl [THEN transD]) apply (blast intro: contract_combD2 reduction_rls) done lemma Ap_reduce2: "\p \ q; r \ comb\ \ r\p \ r\q" apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) apply (erule rtrancl_induct) apply (blast intro: reduction_rls) apply (blast intro: trans_rtrancl [THEN transD] contract_combD2 reduction_rls) done text \Counterexample to the diamond property for \\\<^sup>1\.\ lemma KIII_contract1: "K\I\(I\I) \\<^sup>1 I" by (blast intro: comb_I) lemma KIII_contract2: "K\I\(I\I) \\<^sup>1 K\I\((K\I)\(K\I))" by (unfold I_def) (blast intro: contract.intros) lemma KIII_contract3: "K\I\((K\I)\(K\I)) \\<^sup>1 I" by (blast intro: comb_I) lemma not_diamond_contract: "\ diamond(contract)" unfolding diamond_def apply (blast intro: KIII_contract1 KIII_contract2 KIII_contract3 elim!: I_contract_E) done subsection \Results about Parallel Contraction\ text \For type checking: replaces \a \\<^sup>1 b\ by \a, b \ comb\\ lemmas parcontract_combE2 = parcontract.dom_subset [THEN subsetD, THEN SigmaE2] and parcontract_combD1 = parcontract.dom_subset [THEN subsetD, THEN SigmaD1] and parcontract_combD2 = parcontract.dom_subset [THEN subsetD, THEN SigmaD2] lemma field_parcontract_eq: "field(parcontract) = comb" by (blast intro: parcontract.K elim!: parcontract_combE2) text \Derive a case for each combinator constructor.\ inductive_cases K_parcontractE [elim!]: "K \\<^sup>1 r" and S_parcontractE [elim!]: "S \\<^sup>1 r" and Ap_parcontractE [elim!]: "p\q \\<^sup>1 r" declare parcontract.intros [intro] subsection \Basic properties of parallel contraction\ lemma K1_parcontractD [dest!]: "K\p \\<^sup>1 r \ (\p'. r = K\p' \ p \\<^sup>1 p')" by auto lemma S1_parcontractD [dest!]: "S\p \\<^sup>1 r \ (\p'. r = S\p' \ p \\<^sup>1 p')" by auto lemma S2_parcontractD [dest!]: "S\p\q \\<^sup>1 r \ (\p' q'. r = S\p'\q' \ p \\<^sup>1 p' \ q \\<^sup>1 q')" by auto lemma diamond_parcontract: "diamond(parcontract)" \ \Church-Rosser property for parallel contraction\ unfolding diamond_def apply (rule impI [THEN allI, THEN allI]) apply (erule parcontract.induct) apply (blast elim!: comb.free_elims intro: parcontract_combD2)+ done text \ \medskip Equivalence of \<^prop>\p \ q\ and \<^prop>\p \ q\. \ lemma contract_imp_parcontract: "p\\<^sup>1q \ p\\<^sup>1q" by (induct set: contract) auto lemma reduce_imp_parreduce: "p\q \ p\q" apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) apply (erule rtrancl_induct) apply (blast intro: r_into_trancl) apply (blast intro: contract_imp_parcontract r_into_trancl trans_trancl [THEN transD]) done lemma parcontract_imp_reduce: "p\\<^sup>1q \ p\q" apply (induct set: parcontract) apply (blast intro: reduction_rls) apply (blast intro: reduction_rls) apply (blast intro: reduction_rls) apply (blast intro: trans_rtrancl [THEN transD] Ap_reduce1 Ap_reduce2 parcontract_combD1 parcontract_combD2) done lemma parreduce_imp_reduce: "p\q \ p\q" apply (frule trancl_type [THEN subsetD, THEN SigmaD1]) apply (drule field_parcontract_eq [THEN equalityD1, THEN subsetD]) apply (erule trancl_induct, erule parcontract_imp_reduce) apply (erule trans_rtrancl [THEN transD]) apply (erule parcontract_imp_reduce) done lemma parreduce_iff_reduce: "p\q \ p\q" by (blast intro: parreduce_imp_reduce reduce_imp_parreduce) end diff --git a/src/ZF/Induct/ListN.thy b/src/ZF/Induct/ListN.thy --- a/src/ZF/Induct/ListN.thy +++ b/src/ZF/Induct/ListN.thy @@ -1,61 +1,61 @@ (* Title: ZF/Induct/ListN.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \Lists of n elements\ theory ListN imports ZF begin text \ Inductive definition of lists of \n\ elements; see - @{cite "paulin-tlca"}. + \<^cite>\"paulin-tlca"\. \ consts listn :: "i\i" inductive domains "listn(A)" \ "nat \ list(A)" intros NilI: "\0,Nil\ \ listn(A)" ConsI: "\a \ A; \n,l\ \ listn(A)\ \ \ listn(A)" type_intros nat_typechecks list.intros lemma list_into_listn: "l \ list(A) \ \ listn(A)" by (induct set: list) (simp_all add: listn.intros) lemma listn_iff: "\n,l\ \ listn(A) \ l \ list(A) \ length(l)=n" apply (rule iffI) apply (erule listn.induct) apply auto apply (blast intro: list_into_listn) done lemma listn_image_eq: "listn(A)``{n} = {l \ list(A). length(l)=n}" apply (rule equality_iffI) apply (simp add: listn_iff separation image_singleton_iff) done lemma listn_mono: "A \ B \ listn(A) \ listn(B)" unfolding listn.defs apply (rule lfp_mono) apply (rule listn.bnd_mono)+ apply (assumption | rule univ_mono Sigma_mono list_mono basic_monos)+ done lemma listn_append: "\\n,l\ \ listn(A); \ listn(A)\ \ \ listn(A)" apply (erule listn.induct) apply (frule listn.dom_subset [THEN subsetD]) apply (simp_all add: listn.intros) done inductive_cases Nil_listn_case: "\i,Nil\ \ listn(A)" and Cons_listn_case: " \ listn(A)" inductive_cases zero_listn_case: "\0,l\ \ listn(A)" and succ_listn_case: " \ listn(A)" end diff --git a/src/ZF/Induct/Primrec.thy b/src/ZF/Induct/Primrec.thy --- a/src/ZF/Induct/Primrec.thy +++ b/src/ZF/Induct/Primrec.thy @@ -1,374 +1,374 @@ (* Title: ZF/Induct/Primrec.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) section \Primitive Recursive Functions: the inductive definition\ theory Primrec imports ZF begin text \ - Proof adopted from @{cite szasz93}. + Proof adopted from \<^cite>\szasz93\. - See also @{cite \page 250, exercise 11\ mendelson}. + See also \<^cite>\\page 250, exercise 11\ in mendelson\. \ subsection \Basic definitions\ definition SC :: "i" where "SC \ \l \ list(nat). list_case(0, \x xs. succ(x), l)" definition CONSTANT :: "i\i" where "CONSTANT(k) \ \l \ list(nat). k" definition PROJ :: "i\i" where "PROJ(i) \ \l \ list(nat). list_case(0, \x xs. x, drop(i,l))" definition COMP :: "[i,i]\i" where "COMP(g,fs) \ \l \ list(nat). g ` map(\f. f`l, fs)" definition PREC :: "[i,i]\i" where "PREC(f,g) \ \l \ list(nat). list_case(0, \x xs. rec(x, f`xs, \y r. g ` Cons(r, Cons(y, xs))), l)" \ \Note that \g\ is applied first to \<^term>\PREC(f,g)`y\ and then to \y\!\ consts ACK :: "i\i" primrec "ACK(0) = SC" "ACK(succ(i)) = PREC (CONSTANT (ACK(i) ` [1]), COMP(ACK(i), [PROJ(0)]))" abbreviation ack :: "[i,i]\i" where "ack(x,y) \ ACK(x) ` [y]" text \ \medskip Useful special cases of evaluation. \ lemma SC: "\x \ nat; l \ list(nat)\ \ SC ` (Cons(x,l)) = succ(x)" by (simp add: SC_def) lemma CONSTANT: "l \ list(nat) \ CONSTANT(k) ` l = k" by (simp add: CONSTANT_def) lemma PROJ_0: "\x \ nat; l \ list(nat)\ \ PROJ(0) ` (Cons(x,l)) = x" by (simp add: PROJ_def) lemma COMP_1: "l \ list(nat) \ COMP(g,[f]) ` l = g` [f`l]" by (simp add: COMP_def) lemma PREC_0: "l \ list(nat) \ PREC(f,g) ` (Cons(0,l)) = f`l" by (simp add: PREC_def) lemma PREC_succ: "\x \ nat; l \ list(nat)\ \ PREC(f,g) ` (Cons(succ(x),l)) = g ` Cons(PREC(f,g)`(Cons(x,l)), Cons(x,l))" by (simp add: PREC_def) subsection \Inductive definition of the PR functions\ consts prim_rec :: i inductive domains prim_rec \ "list(nat)->nat" intros "SC \ prim_rec" "k \ nat \ CONSTANT(k) \ prim_rec" "i \ nat \ PROJ(i) \ prim_rec" "\g \ prim_rec; fs\list(prim_rec)\ \ COMP(g,fs) \ prim_rec" "\f \ prim_rec; g \ prim_rec\ \ PREC(f,g) \ prim_rec" monos list_mono con_defs SC_def CONSTANT_def PROJ_def COMP_def PREC_def type_intros nat_typechecks list.intros lam_type list_case_type drop_type map_type apply_type rec_type lemma prim_rec_into_fun [TC]: "c \ prim_rec \ c \ list(nat) -> nat" by (erule subsetD [OF prim_rec.dom_subset]) lemmas [TC] = apply_type [OF prim_rec_into_fun] declare prim_rec.intros [TC] declare nat_into_Ord [TC] declare rec_type [TC] lemma ACK_in_prim_rec [TC]: "i \ nat \ ACK(i) \ prim_rec" by (induct set: nat) simp_all lemma ack_type [TC]: "\i \ nat; j \ nat\ \ ack(i,j) \ nat" by auto subsection \Ackermann's function cases\ lemma ack_0: "j \ nat \ ack(0,j) = succ(j)" \ \PROPERTY A 1\ by (simp add: SC) lemma ack_succ_0: "ack(succ(i), 0) = ack(i,1)" \ \PROPERTY A 2\ by (simp add: CONSTANT PREC_0) lemma ack_succ_succ: "\i\nat; j\nat\ \ ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))" \ \PROPERTY A 3\ by (simp add: CONSTANT PREC_succ COMP_1 PROJ_0) lemmas [simp] = ack_0 ack_succ_0 ack_succ_succ ack_type and [simp del] = ACK.simps lemma lt_ack2: "i \ nat \ j \ nat \ j < ack(i,j)" \ \PROPERTY A 4\ apply (induct i arbitrary: j set: nat) apply simp apply (induct_tac j) apply (erule_tac [2] succ_leI [THEN lt_trans1]) apply (rule nat_0I [THEN nat_0_le, THEN lt_trans]) apply auto done lemma ack_lt_ack_succ2: "\i\nat; j\nat\ \ ack(i,j) < ack(i, succ(j))" \ \PROPERTY A 5-, the single-step lemma\ by (induct set: nat) (simp_all add: lt_ack2) lemma ack_lt_mono2: "\j nat; k \ nat\ \ ack(i,j) < ack(i,k)" \ \PROPERTY A 5, monotonicity for \<\\ apply (frule lt_nat_in_nat, assumption) apply (erule succ_lt_induct) apply assumption apply (rule_tac [2] lt_trans) apply (auto intro: ack_lt_ack_succ2) done lemma ack_le_mono2: "\j\k; i\nat; k\nat\ \ ack(i,j) \ ack(i,k)" \ \PROPERTY A 5', monotonicity for \\\\ apply (rule_tac f = "\j. ack (i,j) " in Ord_lt_mono_imp_le_mono) apply (assumption | rule ack_lt_mono2 ack_type [THEN nat_into_Ord])+ done lemma ack2_le_ack1: "\i\nat; j\nat\ \ ack(i, succ(j)) \ ack(succ(i), j)" \ \PROPERTY A 6\ apply (induct_tac j) apply simp_all apply (rule ack_le_mono2) apply (rule lt_ack2 [THEN succ_leI, THEN le_trans]) apply auto done lemma ack_lt_ack_succ1: "\i \ nat; j \ nat\ \ ack(i,j) < ack(succ(i),j)" \ \PROPERTY A 7-, the single-step lemma\ apply (rule ack_lt_mono2 [THEN lt_trans2]) apply (rule_tac [4] ack2_le_ack1) apply auto done lemma ack_lt_mono1: "\i nat; k \ nat\ \ ack(i,k) < ack(j,k)" \ \PROPERTY A 7, monotonicity for \<\\ apply (frule lt_nat_in_nat, assumption) apply (erule succ_lt_induct) apply assumption apply (rule_tac [2] lt_trans) apply (auto intro: ack_lt_ack_succ1) done lemma ack_le_mono1: "\i\j; j \ nat; k \ nat\ \ ack(i,k) \ ack(j,k)" \ \PROPERTY A 7', monotonicity for \\\\ apply (rule_tac f = "\j. ack (j,k) " in Ord_lt_mono_imp_le_mono) apply (assumption | rule ack_lt_mono1 ack_type [THEN nat_into_Ord])+ done lemma ack_1: "j \ nat \ ack(1,j) = succ(succ(j))" \ \PROPERTY A 8\ by (induct set: nat) simp_all lemma ack_2: "j \ nat \ ack(succ(1),j) = succ(succ(succ(j#+j)))" \ \PROPERTY A 9\ by (induct set: nat) (simp_all add: ack_1) lemma ack_nest_bound: "\i1 \ nat; i2 \ nat; j \ nat\ \ ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)" \ \PROPERTY A 10\ apply (rule lt_trans2 [OF _ ack2_le_ack1]) apply simp apply (rule add_le_self [THEN ack_le_mono1, THEN lt_trans1]) apply auto apply (force intro: add_le_self2 [THEN ack_lt_mono1, THEN ack_lt_mono2]) done lemma ack_add_bound: "\i1 \ nat; i2 \ nat; j \ nat\ \ ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)" \ \PROPERTY A 11\ apply (rule_tac j = "ack (succ (1), ack (i1 #+ i2, j))" in lt_trans) apply (simp add: ack_2) apply (rule_tac [2] ack_nest_bound [THEN lt_trans2]) apply (rule add_le_mono [THEN leI, THEN leI]) apply (auto intro: add_le_self add_le_self2 ack_le_mono1) done lemma ack_add_bound2: "\i < ack(k,j); j \ nat; k \ nat\ \ i#+j < ack(succ(succ(succ(succ(k)))), j)" \ \PROPERTY A 12.\ \ \Article uses existential quantifier but the ALF proof used \<^term>\k#+#4\.\ \ \Quantified version must be nested \\k'. \i,j \\.\ apply (rule_tac j = "ack (k,j) #+ ack (0,j) " in lt_trans) apply (rule_tac [2] ack_add_bound [THEN lt_trans2]) apply (rule add_lt_mono) apply auto done subsection \Main result\ declare list_add_type [simp] lemma SC_case: "l \ list(nat) \ SC ` l < ack(1, list_add(l))" unfolding SC_def apply (erule list.cases) apply (simp add: succ_iff) apply (simp add: ack_1 add_le_self) done lemma lt_ack1: "\i \ nat; j \ nat\ \ i < ack(i,j)" \ \PROPERTY A 4'? Extra lemma needed for \CONSTANT\ case, constant functions.\ apply (induct_tac i) apply (simp add: nat_0_le) apply (erule lt_trans1 [OF succ_leI ack_lt_ack_succ1]) apply auto done lemma CONSTANT_case: "\l \ list(nat); k \ nat\ \ CONSTANT(k) ` l < ack(k, list_add(l))" by (simp add: CONSTANT_def lt_ack1) lemma PROJ_case [rule_format]: "l \ list(nat) \ \i \ nat. PROJ(i) ` l < ack(0, list_add(l))" unfolding PROJ_def apply simp apply (erule list.induct) apply (simp add: nat_0_le) apply simp apply (rule ballI) apply (erule_tac n = i in natE) apply (simp add: add_le_self) apply simp apply (erule bspec [THEN lt_trans2]) apply (rule_tac [2] add_le_self2 [THEN succ_leI]) apply auto done text \ \medskip \COMP\ case. \ lemma COMP_map_lemma: "fs \ list({f \ prim_rec. \kf \ nat. \l \ list(nat). f`l < ack(kf, list_add(l))}) \ \k \ nat. \l \ list(nat). list_add(map(\f. f ` l, fs)) < ack(k, list_add(l))" apply (induct set: list) apply (rule_tac x = 0 in bexI) apply (simp_all add: lt_ack1 nat_0_le) apply clarify apply (rule ballI [THEN bexI]) apply (rule add_lt_mono [THEN lt_trans]) apply (rule_tac [5] ack_add_bound) apply blast apply auto done lemma COMP_case: "\kg\nat; \l \ list(nat). g`l < ack(kg, list_add(l)); fs \ list({f \ prim_rec . \kf \ nat. \l \ list(nat). f`l < ack(kf, list_add(l))})\ \ \k \ nat. \l \ list(nat). COMP(g,fs)`l < ack(k, list_add(l))" apply (simp add: COMP_def) apply (frule list_CollectD) apply (erule COMP_map_lemma [THEN bexE]) apply (rule ballI [THEN bexI]) apply (erule bspec [THEN lt_trans]) apply (rule_tac [2] lt_trans) apply (rule_tac [3] ack_nest_bound) apply (erule_tac [2] bspec [THEN ack_lt_mono2]) apply auto done text \ \medskip \PREC\ case. \ lemma PREC_case_lemma: "\\l \ list(nat). f`l #+ list_add(l) < ack(kf, list_add(l)); \l \ list(nat). g`l #+ list_add(l) < ack(kg, list_add(l)); f \ prim_rec; kf\nat; g \ prim_rec; kg\nat; l \ list(nat)\ \ PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))" unfolding PREC_def apply (erule list.cases) apply (simp add: lt_trans [OF nat_le_refl lt_ack2]) apply simp apply (erule ssubst) \ \get rid of the needless assumption\ apply (induct_tac a) apply simp_all txt \base case\ apply (rule lt_trans, erule bspec, assumption) apply (simp add: add_le_self [THEN ack_lt_mono1]) txt \ind step\ apply (rule succ_leI [THEN lt_trans1]) apply (rule_tac j = "g ` ll #+ mm" for ll mm in lt_trans1) apply (erule_tac [2] bspec) apply (rule nat_le_refl [THEN add_le_mono]) apply typecheck apply (simp add: add_le_self2) txt \final part of the simplification\ apply simp apply (rule add_le_self2 [THEN ack_le_mono1, THEN lt_trans1]) apply (erule_tac [4] ack_lt_mono2) apply auto done lemma PREC_case: "\f \ prim_rec; kf\nat; g \ prim_rec; kg\nat; \l \ list(nat). f`l < ack(kf, list_add(l)); \l \ list(nat). g`l < ack(kg, list_add(l))\ \ \k \ nat. \l \ list(nat). PREC(f,g)`l< ack(k, list_add(l))" apply (rule ballI [THEN bexI]) apply (rule lt_trans1 [OF add_le_self PREC_case_lemma]) apply typecheck apply (blast intro: ack_add_bound2 list_add_type)+ done lemma ack_bounds_prim_rec: "f \ prim_rec \ \k \ nat. \l \ list(nat). f`l < ack(k, list_add(l))" apply (induct set: prim_rec) apply (auto intro: SC_case CONSTANT_case PROJ_case COMP_case PREC_case) done theorem ack_not_prim_rec: "(\l \ list(nat). list_case(0, \x xs. ack(x,x), l)) \ prim_rec" apply (rule notI) apply (drule ack_bounds_prim_rec) apply force done end