diff --git a/doc-src/Classes/Thy/Classes.thy b/doc-src/Classes/Thy/Classes.thy --- a/doc-src/Classes/Thy/Classes.thy +++ b/doc-src/Classes/Thy/Classes.thy @@ -1,634 +1,634 @@ theory Classes imports Main Setup begin section {* Introduction *} text {* Type classes were introduces 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 @{text "eq \ \ \ \ \ bool"} which is overloaded on different types for @{text "\"}, which is achieved by splitting introduction of the @{text eq} function from its overloaded definitions by means of @{text class} and @{text instance} declarations: \footnote{syntax here is a kind of isabellized Haskell} \begin{quote} \noindent@{text "class eq where"} \\ \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} \medskip\noindent@{text "instance nat \ eq where"} \\ \hspace*{2ex}@{text "eq 0 0 = True"} \\ \hspace*{2ex}@{text "eq 0 _ = False"} \\ \hspace*{2ex}@{text "eq _ 0 = False"} \\ \hspace*{2ex}@{text "eq (Suc n) (Suc m) = eq n m"} \medskip\noindent@{text "instance (\\eq, \\eq) pair \ eq where"} \\ \hspace*{2ex}@{text "eq (x1, y1) (x2, y2) = eq x1 x2 \ eq y1 y2"} \medskip\noindent@{text "class ord extends eq where"} \\ \hspace*{2ex}@{text "less_eq \ \ \ \ \ bool"} \\ \hspace*{2ex}@{text "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-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. From a software engeneering 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 @{text "class eq"} above could be given the following specification, demanding that @{text "class eq"} is an equivalence relation obeying reflexivity, symmetry and transitivity: \begin{quote} \noindent@{text "class eq where"} \\ \hspace*{2ex}@{text "eq \ \ \ \ \ bool"} \\ @{text "satisfying"} \\ \hspace*{2ex}@{text "refl: eq x x"} \\ \hspace*{2ex}@{text "sym: eq x y \ eq x y"} \\ \hspace*{2ex}@{text "trans: eq x y \ eq y z \ eq x z"} \end{quote} \noindent From a theoretic 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 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 (aka locales \cite{kammueller-locales}). \end{enumerate} \noindent Isar type classes also directly support code generation in a Haskell like fashion. 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 which some familiarity is assumed. Here we merely present the look-and-feel for end users. Internally, those are mapped to more primitive Isabelle concepts. See \cite{Haftmann-Wenzel:2006:classes} for more detail. *} section {* A simple algebra example \label{sec:example} *} subsection {* Class definition *} text {* Depending on an arbitrary type @{text "\"}, class @{text "semigroup"} introduces a binary operator @{text "(\)"} 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 @{text "(\)"} 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} allows to define 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 default} method, which for such instance proofs maps to the @{method intro_classes} method. This boils down an instance judgement to the relevant primitive proof goals and should conveniently always be 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} are 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 occurence of the name @{text mult_nat} in the primrec declaration; by default, the local name of a class operation @{text f} to instantiate on type constructor @{text \} are mangled as @{text f_\}. In case of uncertainty, these names may be inspected using the @{command "print_context"} command or the corresponding ProofGeneral button. *} subsection {* Lifting and parametric types *} text {* Overloaded definitions giving on 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 * :: (semigroup, semigroup) semigroup begin definition %quote mult_prod_def: "p\<^isub>1 \ p\<^isub>2 = (fst p\<^isub>1 \ fst p\<^isub>2, snd p\<^isub>1 \ snd p\<^isub>2)" instance %quote proof fix p\<^isub>1 p\<^isub>2 p\<^isub>3 :: "\\semigroup \ \\semigroup" show "p\<^isub>1 \ p\<^isub>2 \ p\<^isub>3 = p\<^isub>1 \ (p\<^isub>2 \ p\<^isub>3)" unfolding mult_prod_def by (simp add: assoc) qed end %quote text {* \noindent Associativity from product semigroups is established using the definition of @{text "(\)"} on products and the hypothetical associativity of the type components; these hypotheses are facts 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 @{text monoidl} (a semigroup with a left-hand neutral) by extending @{class semigroup} with one additional parameter @{text neutral} together with its 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 * :: (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 * :: (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 @{text 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 scene *} 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 else 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 axclass *} (*<*)setup %invisible {* Sign.add_path "foo" *} (*>*) axclass %quote idem < type idem: "f (f x) = f x" (*<*)setup %invisible {* Sign.parent_path *}(*>*) text {* \noindent together with a corresponding interpretation: *} interpretation %quote idem_class: idem "f \ (\\idem) \ \" proof qed (rule idem) text {* \noindent This gives you at hand the full power of the Isabelle module system; conclusions in locale @{text idem} are implicitly propagated to class @{text 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 @{text "left_cancel"} lemma for groups, which states that the function @{text "(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 @{text "int"} has been made an instance of @{text "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 support a concept of local definitions in locales: *} 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 @{text 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 (e.g. \cite{krauss2006}). *} subsection {* A functor analogy *} text {* We introduced Isar classes by analogy to type classes 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 which have a canonical interpretation as type classes. Anyway, there is also the possibility of other interpretations. For example, also @{text list}s form a monoid with @{text 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 simply can do a particular interpretation of monoids for lists: *} -interpretation %quote list_monoid!: monoid append "[]" +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 "[]" where +interpretation %quote list_monoid: monoid append "[]" where "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 subsection {* Additional subclass relations *} text {* Any @{text "group"} is also a @{text "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 @{text group} an instance of @{text monoid} by adding an additional edge to the graph of subclass relations (cf.\ \figref{fig:subclass}). \begin{figure}[htbp] \begin{center} \small \unitlength 0.6mm \begin{picture}(40,60)(0,0) \put(20,60){\makebox(0,0){@{text semigroup}}} \put(20,40){\makebox(0,0){@{text monoidl}}} \put(00,20){\makebox(0,0){@{text monoid}}} \put(40,00){\makebox(0,0){@{text 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){@{text semigroup}}} \put(20,40){\makebox(0,0){@{text monoidl}}} \put(00,20){\makebox(0,0){@{text monoid}}} \put(40,00){\makebox(0,0){@{text 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 @{text "group \ monoid"}; transitive edges are left out.} \label{fig:subclass} \end{center} \end{figure} For illustration, a derived definition in @{text group} which uses @{text 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 commodity, class context syntax allows to refer 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 @{text "mult [\]"}, whereas in example 2 the type constraint enforces the global class operation @{text "mult [nat]"}. In the global context in example 3, the reference is to the polymorphic global class operation @{text "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. Concerning target languages lacking type classes (e.g.~SML), type classes are implemented by 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: *} text %quote {*@{code_stmts example (Haskell)}*} text {* \noindent The whole code in SML with explicit dictionary passing: *} text %quote {*@{code_stmts example (SML)}*} 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. \end{description} *} end diff --git a/doc-src/Classes/Thy/document/Classes.tex b/doc-src/Classes/Thy/document/Classes.tex --- a/doc-src/Classes/Thy/document/Classes.tex +++ b/doc-src/Classes/Thy/document/Classes.tex @@ -1,1335 +1,1335 @@ % \begin{isabellebody}% \def\isabellecontext{Classes}% % \isadelimtheory % \endisadelimtheory % \isatagtheory \isacommand{theory}\isamarkupfalse% \ Classes\isanewline \isakeyword{imports}\ Main\ Setup\isanewline \isakeyword{begin}% \endisatagtheory {\isafoldtheory}% % \isadelimtheory % \endisadelimtheory % \isamarkupsection{Introduction% } \isamarkuptrue% % \begin{isamarkuptext}% Type classes were introduces 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 \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on different types for \isa{{\isasymalpha}}, which is achieved by splitting introduction of the \isa{eq} function from its overloaded definitions by means of \isa{class} and \isa{instance} declarations: \footnote{syntax here is a kind of isabellized Haskell} \begin{quote} \noindent\isa{class\ eq\ where} \\ \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \medskip\noindent\isa{instance\ nat\ {\isasymColon}\ eq\ where} \\ \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isadigit{0}}\ {\isacharequal}\ True} \\ \hspace*{2ex}\isa{eq\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ False} \\ \hspace*{2ex}\isa{eq\ {\isacharunderscore}\ {\isadigit{0}}\ {\isacharequal}\ False} \\ \hspace*{2ex}\isa{eq\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharparenleft}Suc\ m{\isacharparenright}\ {\isacharequal}\ eq\ n\ m} \medskip\noindent\isa{instance\ {\isacharparenleft}{\isasymalpha}{\isasymColon}eq{\isacharcomma}\ {\isasymbeta}{\isasymColon}eq{\isacharparenright}\ pair\ {\isasymColon}\ eq\ where} \\ \hspace*{2ex}\isa{eq\ {\isacharparenleft}x{\isadigit{1}}{\isacharcomma}\ y{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ eq\ x{\isadigit{1}}\ x{\isadigit{2}}\ {\isasymand}\ eq\ y{\isadigit{1}}\ y{\isadigit{2}}} \medskip\noindent\isa{class\ ord\ extends\ eq\ where} \\ \hspace*{2ex}\isa{less{\isacharunderscore}eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ \hspace*{2ex}\isa{less\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ 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-Prehofer:1993,nipkow-sorts93,Wenzel:1997:TPHOL}. From a software engeneering 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 \isa{class\ eq} above could be given the following specification, demanding that \isa{class\ eq} is an equivalence relation obeying reflexivity, symmetry and transitivity: \begin{quote} \noindent\isa{class\ eq\ where} \\ \hspace*{2ex}\isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} \\ \isa{satisfying} \\ \hspace*{2ex}\isa{refl{\isacharcolon}\ eq\ x\ x} \\ \hspace*{2ex}\isa{sym{\isacharcolon}\ eq\ x\ y\ {\isasymlongleftrightarrow}\ eq\ x\ y} \\ \hspace*{2ex}\isa{trans{\isacharcolon}\ eq\ x\ y\ {\isasymand}\ eq\ y\ z\ {\isasymlongrightarrow}\ eq\ x\ z} \end{quote} \noindent From a theoretic 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 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 (aka locales \cite{kammueller-locales}). \end{enumerate} \noindent Isar type classes also directly support code generation in a Haskell like fashion. 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 which some familiarity is assumed. Here we merely present the look-and-feel for end users. Internally, those are mapped to more primitive Isabelle concepts. See \cite{Haftmann-Wenzel:2006:classes} for more detail.% \end{isamarkuptext}% \isamarkuptrue% % \isamarkupsection{A simple algebra example \label{sec:example}% } \isamarkuptrue% % \isamarkupsubsection{Class definition% } \isamarkuptrue% % \begin{isamarkuptext}% Depending on an arbitrary type \isa{{\isasymalpha}}, class \isa{semigroup} introduces a binary operator \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} that is assumed to be associative:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{class}\isamarkupfalse% \ semigroup\ {\isacharequal}\isanewline \ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline \ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two parts: the \qn{operational} part names the class parameter (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}). The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel, yielding the global parameter \isa{{\isachardoublequote}mult\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the global theorem \hyperlink{fact.semigroup.assoc:}{\mbox{\isa{semigroup{\isachardot}assoc{\isacharcolon}}}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.% \end{isamarkuptext}% \isamarkuptrue% % \isamarkupsubsection{Class instantiation \label{sec:class_inst}% } \isamarkuptrue% % \begin{isamarkuptext}% The concrete type \isa{int} is made a \isa{semigroup} instance by providing a suitable definition for the class parameter \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}. This is accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{instantiation}\isamarkupfalse% \ int\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% \isanewline \ \ mult{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i\ {\isasymotimes}\ j\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ i\ j\ k\ {\isacharcolon}{\isacharcolon}\ int\ \isacommand{have}\isamarkupfalse% \ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isacharplus}\ j{\isacharparenright}\ {\isacharplus}\ k\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j\ {\isacharplus}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \ \ \isacommand{then}\isamarkupfalse% \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymotimes}\ j{\isacharparenright}\ {\isasymotimes}\ k\ {\isacharequal}\ i\ {\isasymotimes}\ {\isacharparenleft}j\ {\isasymotimes}\ k{\isacharparenright}{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% \isanewline \isacommand{qed}\isamarkupfalse% \isanewline \isanewline \isacommand{end}\isamarkupfalse% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} allows to define class parameters at a particular instance using common specification tools (here, \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}). The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} opens a proof that the given parameters actually conform to the class specification. Note that the first proof step is the \hyperlink{method.default}{\mbox{\isa{default}}} method, which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method. This boils down an instance judgement to the relevant primitive proof goals and should conveniently always be the first method applied in an instantiation proof. From now on, the type-checker will consider \isa{int} as a \isa{semigroup} automatically, i.e.\ any general results are immediately available on concrete instances. \medskip Another instance of \isa{semigroup} are the natural numbers:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{instantiation}\isamarkupfalse% \ nat\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{primrec}\isamarkupfalse% \ mult{\isacharunderscore}nat\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}{\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline \ \ {\isacharbar}\ {\isachardoublequoteopen}Suc\ m\ {\isasymotimes}\ n\ {\isacharequal}\ Suc\ {\isacharparenleft}m\ {\isasymotimes}\ n{\isacharparenright}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\ \isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{by}\isamarkupfalse% \ {\isacharparenleft}induct\ m{\isacharparenright}\ auto\isanewline \isacommand{qed}\isamarkupfalse% \isanewline \isanewline \isacommand{end}\isamarkupfalse% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat} in the primrec declaration; by default, the local name of a class operation \isa{f} to instantiate on type constructor \isa{{\isasymkappa}} are mangled as \isa{f{\isacharunderscore}{\isasymkappa}}. In case of uncertainty, these names may be inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command or the corresponding ProofGeneral button.% \end{isamarkuptext}% \isamarkuptrue% % \isamarkupsubsection{Lifting and parametric types% } \isamarkuptrue% % \begin{isamarkuptext}% Overloaded definitions giving on class instantiation may include recursion over the syntactic structure of types. As a canonical example, we model product semigroups using our simple algebra:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{instantiation}\isamarkupfalse% \ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% \isanewline \ \ mult{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isacharequal}\ {\isacharparenleft}fst\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ fst\ p\isactrlisub {\isadigit{2}}{\isacharcomma}\ snd\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ snd\ p\isactrlisub {\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ p\isactrlisub {\isadigit{1}}\ p\isactrlisub {\isadigit{2}}\ p\isactrlisub {\isadigit{3}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}semigroup\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}semigroup{\isachardoublequoteclose}\isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}\ {\isacharequal}\ p\isactrlisub {\isadigit{1}}\ {\isasymotimes}\ {\isacharparenleft}p\isactrlisub {\isadigit{2}}\ {\isasymotimes}\ p\isactrlisub {\isadigit{3}}{\isacharparenright}{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% \ {\isacharparenleft}simp\ add{\isacharcolon}\ assoc{\isacharparenright}\isanewline \isacommand{qed}\isamarkupfalse% \ \ \ \ \ \ \isanewline \isanewline \isacommand{end}\isamarkupfalse% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent Associativity from product semigroups is established using the definition of \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} on products and the hypothetical associativity of the type components; these hypotheses are facts due to the \isa{semigroup} constraints imposed on the type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition. Indeed, this pattern often occurs with parametric types and type classes.% \end{isamarkuptext}% \isamarkuptrue% % \isamarkupsubsection{Subclassing% } \isamarkuptrue% % \begin{isamarkuptext}% We define a subclass \isa{monoidl} (a semigroup with a left-hand neutral) by extending \isa{semigroup} with one additional parameter \isa{neutral} together with its property:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{class}\isamarkupfalse% \ monoidl\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline \ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}{\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline \ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \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:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{instantiation}\isamarkupfalse% \ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% \isanewline \ \ neutral{\isacharunderscore}nat{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}nat{\isacharparenright}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% \isanewline \ \ neutral{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{0}}{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \isacommand{next}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ k\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \isacommand{qed}\isamarkupfalse% \isanewline \isanewline \isacommand{end}\isamarkupfalse% \isanewline \isanewline \isacommand{instantiation}\isamarkupfalse% \ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoidl{\isacharcomma}\ monoidl{\isacharparenright}\ monoidl\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% \isanewline \ \ neutral{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymone}\ {\isacharequal}\ {\isacharparenleft}{\isasymone}{\isacharcomma}\ {\isasymone}{\isacharparenright}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoidl\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoidl{\isachardoublequoteclose}\isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ p\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% \ {\isacharparenleft}simp\ add{\isacharcolon}\ neutl{\isacharparenright}\isanewline \isacommand{qed}\isamarkupfalse% \isanewline \isanewline \isacommand{end}\isamarkupfalse% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent Fully-fledged monoids are modelled by another subclass which does not add new parameters but tightens the specification:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{class}\isamarkupfalse% \ monoid\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline \ \ \isakeyword{assumes}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instantiation}\isamarkupfalse% \ nat\ \isakeyword{and}\ int\ {\isacharcolon}{\isacharcolon}\ monoid\ \isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}n\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ n{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ neutral{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% \ {\isacharparenleft}induct\ n{\isacharparenright}\ simp{\isacharunderscore}all\isanewline \isacommand{next}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}k\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ k{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \isacommand{qed}\isamarkupfalse% \isanewline \isanewline \isacommand{end}\isamarkupfalse% \isanewline \isanewline \isacommand{instantiation}\isamarkupfalse% \ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoid{\isacharcomma}\ monoid{\isacharparenright}\ monoid\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \ \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ p\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isasymColon}monoid\ {\isasymtimes}\ {\isasymbeta}{\isasymColon}monoid{\isachardoublequoteclose}\isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}p\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ p{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ neutral{\isacharunderscore}prod{\isacharunderscore}def\ mult{\isacharunderscore}prod{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% \ {\isacharparenleft}simp\ add{\isacharcolon}\ neutr{\isacharparenright}\isanewline \isacommand{qed}\isamarkupfalse% \isanewline \isanewline \isacommand{end}\isamarkupfalse% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent To finish our small algebra example, we add a \isa{group} class with a corresponding instance:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{class}\isamarkupfalse% \ group\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline \ \ \isakeyword{fixes}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline \ \ \isakeyword{assumes}\ invl{\isacharcolon}\ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instantiation}\isamarkupfalse% \ int\ {\isacharcolon}{\isacharcolon}\ group\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% \isanewline \ \ inverse{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isacharequal}\ {\isacharminus}\ {\isacharparenleft}i{\isasymColon}int{\isacharparenright}{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ i\ {\isacharcolon}{\isacharcolon}\ int\isanewline \ \ \isacommand{have}\isamarkupfalse% \ {\isachardoublequoteopen}{\isacharminus}i\ {\isacharplus}\ i\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \ \ \isacommand{then}\isamarkupfalse% \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymotimes}\ i\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline \ \ \ \ \isacommand{unfolding}\isamarkupfalse% \ mult{\isacharunderscore}int{\isacharunderscore}def\ neutral{\isacharunderscore}int{\isacharunderscore}def\ inverse{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% \isanewline \isacommand{qed}\isamarkupfalse% \isanewline \isanewline \isacommand{end}\isamarkupfalse% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \isamarkupsection{Type classes as locales% } \isamarkuptrue% % \isamarkupsubsection{A look behind the scene% } \isamarkuptrue% % \begin{isamarkuptext}% 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 else than a locale:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{class}\isamarkupfalse% \ idem\ {\isacharequal}\isanewline \ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline \ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent essentially introduces the locale% \end{isamarkuptext}% \isamarkuptrue% \ % \isadeliminvisible % \endisadeliminvisible % \isataginvisible % \endisataginvisible {\isafoldinvisible}% % \isadeliminvisible % \endisadeliminvisible % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{locale}\isamarkupfalse% \ idem\ {\isacharequal}\isanewline \ \ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline \ \ \isakeyword{assumes}\ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent together with corresponding constant(s):% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{consts}\isamarkupfalse% \ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent The connection to the type system is done by means of a primitive axclass% \end{isamarkuptext}% \isamarkuptrue% \ % \isadeliminvisible % \endisadeliminvisible % \isataginvisible % \endisataginvisible {\isafoldinvisible}% % \isadeliminvisible % \endisadeliminvisible % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{axclass}\isamarkupfalse% \ idem\ {\isacharless}\ type\isanewline \ \ idem{\isacharcolon}\ {\isachardoublequoteopen}f\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharequal}\ f\ x{\isachardoublequoteclose}\ % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \isadeliminvisible % \endisadeliminvisible % \isataginvisible % \endisataginvisible {\isafoldinvisible}% % \isadeliminvisible % \endisadeliminvisible % \begin{isamarkuptext}% \noindent together with a corresponding interpretation:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{interpretation}\isamarkupfalse% \ idem{\isacharunderscore}class{\isacharcolon}\isanewline \ \ idem\ {\isachardoublequoteopen}f\ {\isasymColon}\ {\isacharparenleft}{\isasymalpha}{\isasymColon}idem{\isacharparenright}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\isanewline \isacommand{proof}\isamarkupfalse% \ \isacommand{qed}\isamarkupfalse% \ {\isacharparenleft}rule\ idem{\isacharparenright}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent This gives you at hand the full power of the Isabelle module system; conclusions in locale \isa{idem} are implicitly propagated to class \isa{idem}.% \end{isamarkuptext}% \isamarkuptrue% \ % \isadeliminvisible % \endisadeliminvisible % \isataginvisible % \endisataginvisible {\isafoldinvisible}% % \isadeliminvisible % \endisadeliminvisible % \isamarkupsubsection{Abstract reasoning% } \isamarkuptrue% % \begin{isamarkuptext}% Isabelle locales enable reasoning at a general level, while results are implicitly transferred to all instances. For example, we can now establish the \isa{left{\isacharunderscore}cancel} lemma for groups, which states that the function \isa{{\isacharparenleft}x\ {\isasymotimes}{\isacharparenright}} is injective:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{lemma}\isamarkupfalse% \ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ left{\isacharunderscore}cancel{\isacharcolon}\ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{assume}\isamarkupfalse% \ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\isanewline \ \ \isacommand{then}\isamarkupfalse% \ \isacommand{have}\isamarkupfalse% \ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \ \ \isacommand{then}\isamarkupfalse% \ \isacommand{have}\isamarkupfalse% \ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ y\ {\isacharequal}\ {\isacharparenleft}x{\isasymdiv}\ {\isasymotimes}\ x{\isacharparenright}\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% \ assoc\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \ \ \isacommand{then}\isamarkupfalse% \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% \ neutl\ \isakeyword{and}\ invl\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \isacommand{next}\isamarkupfalse% \isanewline \ \ \isacommand{assume}\isamarkupfalse% \ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline \ \ \isacommand{then}\isamarkupfalse% \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \isacommand{qed}\isamarkupfalse% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{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 \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}. Since type \isa{int} has been made an instance of \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.% \end{isamarkuptext}% \isamarkuptrue% % \isamarkupsubsection{Derived definitions% } \isamarkuptrue% % \begin{isamarkuptext}% Isabelle locales support a concept of local definitions in locales:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{primrec}\isamarkupfalse% \ {\isacharparenleft}\isakeyword{in}\ monoid{\isacharparenright}\ pow{\isacharunderscore}nat\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline \ \ {\isacharbar}\ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent If the locale \isa{group} is also a class, this local definition is propagated onto a global definition of \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}} with corresponding theorems \isa{pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}\isasep\isanewline% pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x}. \noindent As you can see from this example, for local definitions you may use any specification tool which works together with locales (e.g. \cite{krauss2006}).% \end{isamarkuptext}% \isamarkuptrue% % \isamarkupsubsection{A functor analogy% } \isamarkuptrue% % \begin{isamarkuptext}% We introduced Isar classes by analogy to type classes 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 which have a canonical interpretation as type classes. Anyway, there is also the possibility of other interpretations. For example, also \isa{list}s form a monoid with \isa{append} and \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it seems inappropriate to apply to lists the same operations as for genuinely algebraic types. In such a case, we simply can do a particular interpretation of monoids for lists:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{interpretation}\isamarkupfalse% -\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +\ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline \ \ \isacommand{proof}\isamarkupfalse% \ \isacommand{qed}\isamarkupfalse% \ auto% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent This enables us to apply facts on monoids to lists, e.g. \isa{{\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ x\ {\isacharequal}\ x}. When using this interpretation pattern, it may also be appropriate to map derived definitions accordingly:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{primrec}\isamarkupfalse% \ replicate\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isasymalpha}\ list\ {\isasymRightarrow}\ {\isasymalpha}\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}replicate\ {\isadigit{0}}\ {\isacharunderscore}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline \ \ {\isacharbar}\ {\isachardoublequoteopen}replicate\ {\isacharparenleft}Suc\ n{\isacharparenright}\ xs\ {\isacharequal}\ xs\ {\isacharat}\ replicate\ n\ xs{\isachardoublequoteclose}\isanewline \isanewline \isacommand{interpretation}\isamarkupfalse% -\ list{\isacharunderscore}monoid{\isacharbang}{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ list{\isacharunderscore}monoid{\isacharcolon}\ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline \isacommand{proof}\isamarkupfalse% \ {\isacharminus}\isanewline \ \ \isacommand{interpret}\isamarkupfalse% \ monoid\ append\ {\isachardoublequoteopen}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse% \isanewline \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ replicate{\isachardoublequoteclose}\isanewline \ \ \isacommand{proof}\isamarkupfalse% \isanewline \ \ \ \ \isacommand{fix}\isamarkupfalse% \ n\isanewline \ \ \ \ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}monoid{\isachardot}pow{\isacharunderscore}nat\ append\ {\isacharbrackleft}{\isacharbrackright}\ n\ {\isacharequal}\ replicate\ n{\isachardoublequoteclose}\isanewline \ \ \ \ \ \ \isacommand{by}\isamarkupfalse% \ {\isacharparenleft}induct\ n{\isacharparenright}\ auto\isanewline \ \ \isacommand{qed}\isamarkupfalse% \isanewline \isacommand{qed}\isamarkupfalse% \ intro{\isacharunderscore}locales% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \isamarkupsubsection{Additional subclass relations% } \isamarkuptrue% % \begin{isamarkuptext}% Any \isa{group} is also a \isa{monoid}; this can be made explicit by claiming an additional subclass relation, together with a proof of the logical difference:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{subclass}\isamarkupfalse% \ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ monoid\isanewline \isacommand{proof}\isamarkupfalse% \isanewline \ \ \isacommand{fix}\isamarkupfalse% \ x\isanewline \ \ \isacommand{from}\isamarkupfalse% \ invl\ \isacommand{have}\isamarkupfalse% \ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \ \ \isacommand{with}\isamarkupfalse% \ assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}\ neutl\ invl\ \isacommand{have}\isamarkupfalse% \ {\isachardoublequoteopen}x{\isasymdiv}\ {\isasymotimes}\ {\isacharparenleft}x\ {\isasymotimes}\ {\isasymone}{\isacharparenright}\ {\isacharequal}\ x{\isasymdiv}\ {\isasymotimes}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \ \ \isacommand{with}\isamarkupfalse% \ left{\isacharunderscore}cancel\ \isacommand{show}\isamarkupfalse% \ {\isachardoublequoteopen}x\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% \ simp\isanewline \isacommand{qed}\isamarkupfalse% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% The logical proof is carried out on the locale level. Afterwards it is propagated to the type system, making \isa{group} an instance of \isa{monoid} by adding an additional edge to the graph of subclass relations (cf.\ \figref{fig:subclass}). \begin{figure}[htbp] \begin{center} \small \unitlength 0.6mm \begin{picture}(40,60)(0,0) \put(20,60){\makebox(0,0){\isa{semigroup}}} \put(20,40){\makebox(0,0){\isa{monoidl}}} \put(00,20){\makebox(0,0){\isa{monoid}}} \put(40,00){\makebox(0,0){\isa{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){\isa{semigroup}}} \put(20,40){\makebox(0,0){\isa{monoidl}}} \put(00,20){\makebox(0,0){\isa{monoid}}} \put(40,00){\makebox(0,0){\isa{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 \isa{group\ {\isasymsubseteq}\ monoid}; transitive edges are left out.} \label{fig:subclass} \end{center} \end{figure} For illustration, a derived definition in \isa{group} which uses \isa{pow{\isacharunderscore}nat}:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{definition}\isamarkupfalse% \ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ pow{\isacharunderscore}int\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}int\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ k\ {\isachargreater}{\isacharequal}\ {\isadigit{0}}\isanewline \ \ \ \ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\isanewline \ \ \ \ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent yields the global definition of \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}} with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.% \end{isamarkuptext}% \isamarkuptrue% % \isamarkupsubsection{A note on syntax% } \isamarkuptrue% % \begin{isamarkuptext}% As a commodity, class context syntax allows to refer to local class operations and their global counterparts uniformly; type inference resolves ambiguities. For example:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{context}\isamarkupfalse% \ semigroup\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{term}\isamarkupfalse% \ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % \isamarkupcmt{example 1% } \isanewline \isacommand{term}\isamarkupfalse% \ {\isachardoublequoteopen}{\isacharparenleft}x{\isasymColon}nat{\isacharparenright}\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % \isamarkupcmt{example 2% } \isanewline \isanewline \isacommand{end}\isamarkupfalse% \isanewline \isanewline \isacommand{term}\isamarkupfalse% \ {\isachardoublequoteopen}x\ {\isasymotimes}\ y{\isachardoublequoteclose}\ % \isamarkupcmt{example 3% } % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent Here in example 1, the term refers to the local class operation \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type constraint enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}. In the global context in example 3, the reference is to the polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.% \end{isamarkuptext}% \isamarkuptrue% % \isamarkupsection{Further issues% } \isamarkuptrue% % \isamarkupsubsection{Type classes and code generation% } \isamarkuptrue% % \begin{isamarkuptext}% Turning back to the first motivation for type classes, namely overloading, it is obvious that overloading stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} targets naturally maps to Haskell type classes. The code generator framework \cite{isabelle-codegen} takes this into account. Concerning target languages lacking type classes (e.g.~SML), type classes are implemented by explicit dictionary construction. As example, let's go back to the power function:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote \isacommand{definition}\isamarkupfalse% \ example\ {\isacharcolon}{\isacharcolon}\ int\ \isakeyword{where}\isanewline \ \ {\isachardoublequoteopen}example\ {\isacharequal}\ pow{\isacharunderscore}int\ {\isadigit{1}}{\isadigit{0}}\ {\isacharparenleft}{\isacharminus}{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent This maps to Haskell as:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote % \begin{isamarkuptext}% \isatypewriter% \noindent% \hspace*{0pt}module Example where {\char123}\\ \hspace*{0pt}\\ \hspace*{0pt}\\ \hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ \hspace*{0pt}\\ \hspace*{0pt}nat{\char95}aux ::~Integer -> Nat -> Nat;\\ \hspace*{0pt}nat{\char95}aux i n = (if i <= 0 then n else nat{\char95}aux (i - 1) (Suc n));\\ \hspace*{0pt}\\ \hspace*{0pt}nat ::~Integer -> Nat;\\ \hspace*{0pt}nat i = nat{\char95}aux i Zero{\char95}nat;\\ \hspace*{0pt}\\ \hspace*{0pt}class Semigroup a where {\char123}\\ \hspace*{0pt} ~mult ::~a -> a -> a;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ \hspace*{0pt}class (Semigroup a) => Monoidl a where {\char123}\\ \hspace*{0pt} ~neutral ::~a;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ \hspace*{0pt}class (Monoidl a) => Monoid a where {\char123}\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ \hspace*{0pt}class (Monoid a) => Group a where {\char123}\\ \hspace*{0pt} ~inverse ::~a -> a;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ \hspace*{0pt}inverse{\char95}int ::~Integer -> Integer;\\ \hspace*{0pt}inverse{\char95}int i = negate i;\\ \hspace*{0pt}\\ \hspace*{0pt}neutral{\char95}int ::~Integer;\\ \hspace*{0pt}neutral{\char95}int = 0;\\ \hspace*{0pt}\\ \hspace*{0pt}mult{\char95}int ::~Integer -> Integer -> Integer;\\ \hspace*{0pt}mult{\char95}int i j = i + j;\\ \hspace*{0pt}\\ \hspace*{0pt}instance Semigroup Integer where {\char123}\\ \hspace*{0pt} ~mult = mult{\char95}int;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ \hspace*{0pt}instance Monoidl Integer where {\char123}\\ \hspace*{0pt} ~neutral = neutral{\char95}int;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ \hspace*{0pt}instance Monoid Integer where {\char123}\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ \hspace*{0pt}instance Group Integer where {\char123}\\ \hspace*{0pt} ~inverse = inverse{\char95}int;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ \hspace*{0pt}pow{\char95}nat ::~forall a.~(Monoid a) => Nat -> a -> a;\\ \hspace*{0pt}pow{\char95}nat Zero{\char95}nat x = neutral;\\ -\hspace*{0pt}pow{\char95}nat (Suc n) x = mult x (pow{\char95}nat n x);\\ +\hspace*{0pt}pow{\char95}nat (Suc n) xa = mult xa (pow{\char95}nat n xa);\\ \hspace*{0pt}\\ \hspace*{0pt}pow{\char95}int ::~forall a.~(Group a) => Integer -> a -> a;\\ \hspace*{0pt}pow{\char95}int k x =\\ \hspace*{0pt} ~(if 0 <= k then pow{\char95}nat (nat k) x\\ \hspace*{0pt} ~~~else inverse (pow{\char95}nat (nat (negate k)) x));\\ \hspace*{0pt}\\ \hspace*{0pt}example ::~Integer;\\ \hspace*{0pt}example = pow{\char95}int 10 (-2);\\ \hspace*{0pt}\\ \hspace*{0pt}{\char125}% \end{isamarkuptext}% \isamarkuptrue% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \begin{isamarkuptext}% \noindent The whole code in SML with explicit dictionary passing:% \end{isamarkuptext}% \isamarkuptrue% % \isadelimquote % \endisadelimquote % \isatagquote % \begin{isamarkuptext}% \isatypewriter% \noindent% \hspace*{0pt}structure Example = \\ \hspace*{0pt}struct\\ \hspace*{0pt}\\ \hspace*{0pt}datatype nat = Zero{\char95}nat | Suc of nat;\\ \hspace*{0pt}\\ \hspace*{0pt}fun nat{\char95}aux i n =\\ \hspace*{0pt} ~(if IntInf.<= (i,~(0 :~IntInf.int)) then n\\ \hspace*{0pt} ~~~else nat{\char95}aux (IntInf.- (i,~(1 :~IntInf.int))) (Suc n));\\ \hspace*{0pt}\\ \hspace*{0pt}fun nat i = nat{\char95}aux i Zero{\char95}nat;\\ \hspace*{0pt}\\ \hspace*{0pt}type 'a semigroup = {\char123}mult :~'a -> 'a -> 'a{\char125};\\ \hspace*{0pt}fun mult (A{\char95}:'a semigroup) = {\char35}mult A{\char95};\\ \hspace*{0pt}\\ \hspace*{0pt}type 'a monoidl =\\ \hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl :~'a semigroup,~neutral :~'a{\char125};\\ \hspace*{0pt}fun semigroup{\char95}monoidl (A{\char95}:'a monoidl) = {\char35}Classes{\char95}{\char95}semigroup{\char95}monoidl A{\char95};\\ \hspace*{0pt}fun neutral (A{\char95}:'a monoidl) = {\char35}neutral A{\char95};\\ \hspace*{0pt}\\ \hspace*{0pt}type 'a monoid = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid :~'a monoidl{\char125};\\ \hspace*{0pt}fun monoidl{\char95}monoid (A{\char95}:'a monoid) = {\char35}Classes{\char95}{\char95}monoidl{\char95}monoid A{\char95};\\ \hspace*{0pt}\\ \hspace*{0pt}type 'a group = {\char123}Classes{\char95}{\char95}monoid{\char95}group :~'a monoid,~inverse :~'a -> 'a{\char125};\\ \hspace*{0pt}fun monoid{\char95}group (A{\char95}:'a group) = {\char35}Classes{\char95}{\char95}monoid{\char95}group A{\char95};\\ \hspace*{0pt}fun inverse (A{\char95}:'a group) = {\char35}inverse A{\char95};\\ \hspace*{0pt}\\ \hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\ \hspace*{0pt}\\ \hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int)\\ \hspace*{0pt}\\ \hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\ \hspace*{0pt}\\ \hspace*{0pt}val semigroup{\char95}int = {\char123}mult = mult{\char95}int{\char125}~:~IntInf.int semigroup;\\ \hspace*{0pt}\\ \hspace*{0pt}val monoidl{\char95}int =\\ \hspace*{0pt} ~{\char123}Classes{\char95}{\char95}semigroup{\char95}monoidl = semigroup{\char95}int,~neutral = neutral{\char95}int{\char125}~:\\ \hspace*{0pt} ~IntInf.int monoidl;\\ \hspace*{0pt}\\ \hspace*{0pt}val monoid{\char95}int = {\char123}Classes{\char95}{\char95}monoidl{\char95}monoid = monoidl{\char95}int{\char125}~:\\ \hspace*{0pt} ~IntInf.int monoid;\\ \hspace*{0pt}\\ \hspace*{0pt}val group{\char95}int =\\ \hspace*{0pt} ~{\char123}Classes{\char95}{\char95}monoid{\char95}group = monoid{\char95}int,~inverse = inverse{\char95}int{\char125}~:\\ \hspace*{0pt} ~IntInf.int group;\\ \hspace*{0pt}\\ \hspace*{0pt}fun pow{\char95}nat A{\char95}~Zero{\char95}nat x = neutral (monoidl{\char95}monoid A{\char95})\\ -\hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) x =\\ -\hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) x (pow{\char95}nat A{\char95}~n x);\\ +\hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) xa =\\ +\hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) xa (pow{\char95}nat A{\char95}~n xa);\\ \hspace*{0pt}\\ \hspace*{0pt}fun pow{\char95}int A{\char95}~k x =\\ \hspace*{0pt} ~(if IntInf.<= ((0 :~IntInf.int),~k)\\ \hspace*{0pt} ~~~then pow{\char95}nat (monoid{\char95}group A{\char95}) (nat k) x\\ \hspace*{0pt} ~~~else inverse A{\char95}~(pow{\char95}nat (monoid{\char95}group A{\char95}) (nat (IntInf.{\char126}~k)) x));\\ \hspace*{0pt}\\ \hspace*{0pt}val example :~IntInf.int =\\ \hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int)\\ \hspace*{0pt}\\ \hspace*{0pt}end;~(*struct Example*)% \end{isamarkuptext}% \isamarkuptrue% % \endisatagquote {\isafoldquote}% % \isadelimquote % \endisadelimquote % \isamarkupsubsection{Inspecting the type class universe% } \isamarkuptrue% % \begin{isamarkuptext}% To facilitate orientation in complex subclass structures, two diagnostics commands are provided: \begin{description} \item[\hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isacharunderscore}classes}}}}] print a list of all classes together with associated operations etc. \item[\hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isacharunderscore}deps}}}}] visualizes the subclass relation between all classes as a Hasse diagram. \end{description}% \end{isamarkuptext}% \isamarkuptrue% % \isadelimtheory % \endisadelimtheory % \isatagtheory \isacommand{end}\isamarkupfalse% % \endisatagtheory {\isafoldtheory}% % \isadelimtheory % \endisadelimtheory \isanewline \end{isabellebody}% %%% Local Variables: %%% mode: latex %%% TeX-master: "root" %%% End: diff --git a/src/FOL/ex/LocaleTest.thy b/src/FOL/ex/LocaleTest.thy --- a/src/FOL/ex/LocaleTest.thy +++ b/src/FOL/ex/LocaleTest.thy @@ -1,483 +1,483 @@ (* Title: FOL/ex/NewLocaleTest.thy Author: Clemens Ballarin, TU Muenchen Testing environment for locale expressions. *) theory LocaleTest imports FOL begin typedecl int arities int :: "term" consts plus :: "int => int => int" (infixl "+" 60) zero :: int ("0") minus :: "int => int" ("- _") axioms int_assoc: "(x + y::int) + z = x + (y + z)" int_zero: "0 + x = x" int_minus: "(-x) + x = 0" int_minus2: "-(-x) = x" section {* Inference of parameter types *} locale param1 = fixes p print_locale! param1 locale param2 = fixes p :: 'b print_locale! param2 (* locale param_top = param2 r for r :: "'b :: {}" Fails, cannot generalise parameter. *) locale param3 = fixes p (infix ".." 50) print_locale! param3 locale param4 = fixes p :: "'a => 'a => 'a" (infix ".." 50) print_locale! param4 subsection {* Incremental type constraints *} locale constraint1 = fixes prod (infixl "**" 65) assumes l_id: "x ** y = x" assumes assoc: "(x ** y) ** z = x ** (y ** z)" print_locale! constraint1 locale constraint2 = fixes p and q assumes "p = q" print_locale! constraint2 section {* Inheritance *} locale semi = fixes prod (infixl "**" 65) assumes assoc: "(x ** y) ** z = x ** (y ** z)" print_locale! semi thm semi_def locale lgrp = semi + fixes one and inv assumes lone: "one ** x = x" and linv: "inv(x) ** x = one" print_locale! lgrp thm lgrp_def lgrp_axioms_def locale add_lgrp = semi "op ++" for sum (infixl "++" 60) + fixes zero and neg assumes lzero: "zero ++ x = x" and lneg: "neg(x) ++ x = zero" print_locale! add_lgrp thm add_lgrp_def add_lgrp_axioms_def locale rev_lgrp = semi "%x y. y ++ x" for sum (infixl "++" 60) print_locale! rev_lgrp thm rev_lgrp_def locale hom = f: semi f + g: semi g for f and g print_locale! hom thm hom_def locale perturbation = semi + d: semi "%x y. delta(x) ** delta(y)" for delta print_locale! perturbation thm perturbation_def locale pert_hom = d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2 print_locale! pert_hom thm pert_hom_def text {* Alternative expression, obtaining nicer names in @{text "semi f"}. *} locale pert_hom' = semi f + d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2 print_locale! pert_hom' thm pert_hom'_def section {* Syntax declarations *} locale logic = fixes land (infixl "&&" 55) and lnot ("-- _" [60] 60) assumes assoc: "(x && y) && z = x && (y && z)" and notnot: "-- (-- x) = x" begin definition lor (infixl "||" 50) where "x || y = --(-- x && -- y)" end print_locale! logic locale use_decl = logic + semi "op ||" print_locale! use_decl thm use_decl_def locale extra_type = fixes a :: 'a and P :: "'a => 'b => o" begin definition test :: "'a => o" where "test(x) <-> (ALL b. P(x, b))" end term extra_type.test thm extra_type.test_def -interpretation var: extra_type "0" "%x y. x = 0" . +interpretation var?: extra_type "0" "%x y. x = 0" . thm var.test_def section {* Foundational versions of theorems *} thm logic.assoc thm logic.lor_def section {* Defines *} locale logic_def = fixes land (infixl "&&" 55) and lor (infixl "||" 50) and lnot ("-- _" [60] 60) assumes assoc: "(x && y) && z = x && (y && z)" and notnot: "-- (-- x) = x" defines "x || y == --(-- x && -- y)" begin thm lor_def (* Can we get rid the the additional hypothesis, caused by LocalTheory.notes? *) lemma "x || y = --(-- x && --y)" by (unfold lor_def) (rule refl) end (* Inheritance of defines *) locale logic_def2 = logic_def begin lemma "x || y = --(-- x && --y)" by (unfold lor_def) (rule refl) end section {* Notes *} (* A somewhat arcane homomorphism example *) definition semi_hom where "semi_hom(prod, sum, h) <-> (ALL x y. h(prod(x, y)) = sum(h(x), h(y)))" lemma semi_hom_mult: "semi_hom(prod, sum, h) ==> h(prod(x, y)) = sum(h(x), h(y))" by (simp add: semi_hom_def) locale semi_hom_loc = prod: semi prod + sum: semi sum for prod and sum and h + assumes semi_homh: "semi_hom(prod, sum, h)" notes semi_hom_mult = semi_hom_mult [OF semi_homh] thm semi_hom_loc.semi_hom_mult (* unspecified, attribute not applied in backgroud theory !!! *) lemma (in semi_hom_loc) "h(prod(x, y)) = sum(h(x), h(y))" by (rule semi_hom_mult) (* Referring to facts from within a context specification *) lemma assumes x: "P <-> P" notes y = x shows True .. section {* Theorem statements *} lemma (in lgrp) lcancel: "x ** y = x ** z <-> y = z" proof assume "x ** y = x ** z" then have "inv(x) ** x ** y = inv(x) ** x ** z" by (simp add: assoc) then show "y = z" by (simp add: lone linv) qed simp print_locale! lgrp locale rgrp = semi + fixes one and inv assumes rone: "x ** one = x" and rinv: "x ** inv(x) = one" begin lemma rcancel: "y ** x = z ** x <-> y = z" proof assume "y ** x = z ** x" then have "y ** (x ** inv(x)) = z ** (x ** inv(x))" by (simp add: assoc [symmetric]) then show "y = z" by (simp add: rone rinv) qed simp end print_locale! rgrp subsection {* Patterns *} lemma (in rgrp) assumes "y ** x = z ** x" (is ?a) shows "y = z" (is ?t) proof - txt {* Weird proof involving patterns from context element and conclusion. *} { assume ?a then have "y ** (x ** inv(x)) = z ** (x ** inv(x))" by (simp add: assoc [symmetric]) then have ?t by (simp add: rone rinv) } note x = this show ?t by (rule x [OF `?a`]) qed section {* Interpretation between locales: sublocales *} sublocale lgrp < right: rgrp print_facts proof unfold_locales { fix x have "inv(x) ** x ** one = inv(x) ** x" by (simp add: linv lone) then show "x ** one = x" by (simp add: assoc lcancel) } note rone = this { fix x have "inv(x) ** x ** inv(x) = inv(x) ** one" by (simp add: linv lone rone) then show "x ** inv(x) = one" by (simp add: assoc lcancel) } qed (* effect on printed locale *) print_locale! lgrp (* use of derived theorem *) lemma (in lgrp) "y ** x = z ** x <-> y = z" apply (rule rcancel) done (* circular interpretation *) sublocale rgrp < left: lgrp proof unfold_locales { fix x have "one ** (x ** inv(x)) = x ** inv(x)" by (simp add: rinv rone) then show "one ** x = x" by (simp add: assoc [symmetric] rcancel) } note lone = this { fix x have "inv(x) ** (x ** inv(x)) = one ** inv(x)" by (simp add: rinv lone rone) then show "inv(x) ** x = one" by (simp add: assoc [symmetric] rcancel) } qed (* effect on printed locale *) print_locale! rgrp print_locale! lgrp (* Duality *) locale order = fixes less :: "'a => 'a => o" (infix "<<" 50) assumes refl: "x << x" and trans: "[| x << y; y << z |] ==> x << z" sublocale order < dual: order "%x y. y << x" apply unfold_locales apply (rule refl) apply (blast intro: trans) done print_locale! order (* Only two instances of order. *) locale order' = fixes less :: "'a => 'a => o" (infix "<<" 50) assumes refl: "x << x" and trans: "[| x << y; y << z |] ==> x << z" locale order_with_def = order' begin definition greater :: "'a => 'a => o" (infix ">>" 50) where "x >> y <-> y << x" end sublocale order_with_def < dual: order' "op >>" apply unfold_locales unfolding greater_def apply (rule refl) apply (blast intro: trans) done print_locale! order_with_def (* Note that decls come after theorems that make use of them. *) (* locale with many parameters --- interpretations generate alternating group A5 *) locale A5 = fixes A and B and C and D and E assumes eq: "A <-> B <-> C <-> D <-> E" sublocale A5 < 1: A5 _ _ D E C print_facts using eq apply (blast intro: A5.intro) done sublocale A5 < 2: A5 C _ E _ A print_facts using eq apply (blast intro: A5.intro) done sublocale A5 < 3: A5 B C A _ _ print_facts using eq apply (blast intro: A5.intro) done (* Any even permutation of parameters is subsumed by the above. *) print_locale! A5 (* Free arguments of instance *) locale trivial = fixes P and Q :: o assumes Q: "P <-> P <-> Q" begin lemma Q_triv: "Q" using Q by fast end sublocale trivial < x: trivial x _ apply unfold_locales using Q by fast print_locale! trivial context trivial begin thm x.Q [where ?x = True] end sublocale trivial < y: trivial Q Q by unfold_locales (* Succeeds since previous interpretation is more general. *) print_locale! trivial (* No instance for y created (subsumed). *) subsection {* Sublocale, then interpretation in theory *} -interpretation int: lgrp "op +" "0" "minus" +interpretation int?: lgrp "op +" "0" "minus" proof unfold_locales qed (rule int_assoc int_zero int_minus)+ thm int.assoc int.semi_axioms -interpretation int2: semi "op +" +interpretation int2?: semi "op +" by unfold_locales (* subsumed, thm int2.assoc not generated *) thm int.lone int.right.rone (* the latter comes through the sublocale relation *) subsection {* Interpretation in theory, then sublocale *} interpretation (* fol: *) logic "op +" "minus" (* FIXME declaration of qualified names *) by unfold_locales (rule int_assoc int_minus2)+ locale logic2 = fixes land (infixl "&&" 55) and lnot ("-- _" [60] 60) assumes assoc: "(x && y) && z = x && (y && z)" and notnot: "-- (-- x) = x" begin (* FIXME interpretation below fails definition lor (infixl "||" 50) where "x || y = --(-- x && -- y)" *) end sublocale logic < two: logic2 by unfold_locales (rule assoc notnot)+ thm two.assoc subsection {* Declarations and sublocale *} locale logic_a = logic locale logic_b = logic sublocale logic_a < logic_b by unfold_locales subsection {* Equations *} locale logic_o = fixes land (infixl "&&" 55) and lnot ("-- _" [60] 60) assumes assoc_o: "(x && y) && z <-> x && (y && z)" and notnot_o: "-- (-- x) <-> x" begin definition lor_o (infixl "||" 50) where "x || y <-> --(-- x && -- y)" end -interpretation x!: logic_o "op &" "Not" +interpretation x: logic_o "op &" "Not" where bool_logic_o: "logic_o.lor_o(op &, Not, x, y) <-> x | y" proof - show bool_logic_o: "PROP logic_o(op &, Not)" by unfold_locales fast+ show "logic_o.lor_o(op &, Not, x, y) <-> x | y" by (unfold logic_o.lor_o_def [OF bool_logic_o]) fast qed thm x.lor_o_def bool_logic_o lemma lor_triv: "z <-> z" .. lemma (in logic_o) lor_triv: "x || y <-> x || y" by fast thm lor_triv [where z = True] (* Check strict prefix. *) x.lor_triv subsection {* Interpretation in proofs *} lemma True proof interpret "local": lgrp "op +" "0" "minus" by unfold_locales (* subsumed *) { fix zero :: int assume "!!x. zero + x = x" "!!x. (-x) + x = zero" then interpret local_fixed: lgrp "op +" zero "minus" by unfold_locales thm local_fixed.lone } assume "!!x zero. zero + x = x" "!!x zero. (-x) + x = zero" then interpret local_free: lgrp "op +" zero "minus" for zero by unfold_locales thm local_free.lone [where ?zero = 0] qed end diff --git a/src/HOL/Algebra/AbelCoset.thy b/src/HOL/Algebra/AbelCoset.thy --- a/src/HOL/Algebra/AbelCoset.thy +++ b/src/HOL/Algebra/AbelCoset.thy @@ -1,749 +1,749 @@ (* Title: HOL/Algebra/AbelCoset.thy Author: Stephan Hohe, TU Muenchen *) theory AbelCoset imports Coset Ring begin subsection {* More Lifting from Groups to Abelian Groups *} subsubsection {* Definitions *} text {* Hiding @{text "<+>"} from @{theory Sum_Type} until I come up with better syntax here *} no_notation Plus (infixr "<+>" 65) constdefs (structure G) a_r_coset :: "[_, 'a set, 'a] \ 'a set" (infixl "+>\" 60) "a_r_coset G \ r_coset \carrier = carrier G, mult = add G, one = zero G\" a_l_coset :: "[_, 'a, 'a set] \ 'a set" (infixl "<+\" 60) "a_l_coset G \ l_coset \carrier = carrier G, mult = add G, one = zero G\" A_RCOSETS :: "[_, 'a set] \ ('a set)set" ("a'_rcosets\ _" [81] 80) "A_RCOSETS G H \ RCOSETS \carrier = carrier G, mult = add G, one = zero G\ H" set_add :: "[_, 'a set ,'a set] \ 'a set" (infixl "<+>\" 60) "set_add G \ set_mult \carrier = carrier G, mult = add G, one = zero G\" A_SET_INV :: "[_,'a set] \ 'a set" ("a'_set'_inv\ _" [81] 80) "A_SET_INV G H \ SET_INV \carrier = carrier G, mult = add G, one = zero G\ H" constdefs (structure G) a_r_congruent :: "[('a,'b)ring_scheme, 'a set] \ ('a*'a)set" ("racong\ _") "a_r_congruent G \ r_congruent \carrier = carrier G, mult = add G, one = zero G\" constdefs A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \ ('a set) monoid" (infixl "A'_Mod" 65) --{*Actually defined for groups rather than monoids*} "A_FactGroup G H \ FactGroup \carrier = carrier G, mult = add G, one = zero G\ H" constdefs a_kernel :: "('a, 'm) ring_scheme \ ('b, 'n) ring_scheme \ ('a \ 'b) \ 'a set" --{*the kernel of a homomorphism (additive)*} "a_kernel G H h \ kernel \carrier = carrier G, mult = add G, one = zero G\ \carrier = carrier H, mult = add H, one = zero H\ h" locale abelian_group_hom = G: abelian_group G + H: abelian_group H for G (structure) and H (structure) + fixes h assumes a_group_hom: "group_hom (| carrier = carrier G, mult = add G, one = zero G |) (| carrier = carrier H, mult = add H, one = zero H |) h" lemmas a_r_coset_defs = a_r_coset_def r_coset_def lemma a_r_coset_def': fixes G (structure) shows "H +> a \ \h\H. {h \ a}" unfolding a_r_coset_defs by simp lemmas a_l_coset_defs = a_l_coset_def l_coset_def lemma a_l_coset_def': fixes G (structure) shows "a <+ H \ \h\H. {a \ h}" unfolding a_l_coset_defs by simp lemmas A_RCOSETS_defs = A_RCOSETS_def RCOSETS_def lemma A_RCOSETS_def': fixes G (structure) shows "a_rcosets H \ \a\carrier G. {H +> a}" unfolding A_RCOSETS_defs by (fold a_r_coset_def, simp) lemmas set_add_defs = set_add_def set_mult_def lemma set_add_def': fixes G (structure) shows "H <+> K \ \h\H. \k\K. {h \ k}" unfolding set_add_defs by simp lemmas A_SET_INV_defs = A_SET_INV_def SET_INV_def lemma A_SET_INV_def': fixes G (structure) shows "a_set_inv H \ \h\H. {\ h}" unfolding A_SET_INV_defs by (fold a_inv_def) subsubsection {* Cosets *} lemma (in abelian_group) a_coset_add_assoc: "[| M \ carrier G; g \ carrier G; h \ carrier G |] ==> (M +> g) +> h = M +> (g \ h)" by (rule group.coset_mult_assoc [OF a_group, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_coset_add_zero [simp]: "M \ carrier G ==> M +> \ = M" by (rule group.coset_mult_one [OF a_group, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_coset_add_inv1: "[| M +> (x \ (\ y)) = M; x \ carrier G ; y \ carrier G; M \ carrier G |] ==> M +> x = M +> y" by (rule group.coset_mult_inv1 [OF a_group, folded a_r_coset_def a_inv_def, simplified monoid_record_simps]) lemma (in abelian_group) a_coset_add_inv2: "[| M +> x = M +> y; x \ carrier G; y \ carrier G; M \ carrier G |] ==> M +> (x \ (\ y)) = M" by (rule group.coset_mult_inv2 [OF a_group, folded a_r_coset_def a_inv_def, simplified monoid_record_simps]) lemma (in abelian_group) a_coset_join1: "[| H +> x = H; x \ carrier G; subgroup H (|carrier = carrier G, mult = add G, one = zero G|) |] ==> x \ H" by (rule group.coset_join1 [OF a_group, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_solve_equation: "\subgroup H (|carrier = carrier G, mult = add G, one = zero G|); x \ H; y \ H\ \ \h\H. y = h \ x" by (rule group.solve_equation [OF a_group, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_repr_independence: "\y \ H +> x; x \ carrier G; subgroup H \carrier = carrier G, mult = add G, one = zero G\ \ \ H +> x = H +> y" by (rule group.repr_independence [OF a_group, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_coset_join2: "\x \ carrier G; subgroup H \carrier = carrier G, mult = add G, one = zero G\; x\H\ \ H +> x = H" by (rule group.coset_join2 [OF a_group, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_monoid) a_r_coset_subset_G: "[| H \ carrier G; x \ carrier G |] ==> H +> x \ carrier G" by (rule monoid.r_coset_subset_G [OF a_monoid, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_rcosI: "[| h \ H; H \ carrier G; x \ carrier G|] ==> h \ x \ H +> x" by (rule group.rcosI [OF a_group, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_rcosetsI: "\H \ carrier G; x \ carrier G\ \ H +> x \ a_rcosets H" by (rule group.rcosetsI [OF a_group, folded a_r_coset_def A_RCOSETS_def, simplified monoid_record_simps]) text{*Really needed?*} lemma (in abelian_group) a_transpose_inv: "[| x \ y = z; x \ carrier G; y \ carrier G; z \ carrier G |] ==> (\ x) \ z = y" by (rule group.transpose_inv [OF a_group, folded a_r_coset_def a_inv_def, simplified monoid_record_simps]) (* --"duplicate" lemma (in abelian_group) a_rcos_self: "[| x \ carrier G; subgroup H \carrier = carrier G, mult = add G, one = zero G\ |] ==> x \ H +> x" by (rule group.rcos_self [OF a_group, folded a_r_coset_def, simplified monoid_record_simps]) *) subsubsection {* Subgroups *} locale additive_subgroup = fixes H and G (structure) assumes a_subgroup: "subgroup H \carrier = carrier G, mult = add G, one = zero G\" lemma (in additive_subgroup) is_additive_subgroup: shows "additive_subgroup H G" by (rule additive_subgroup_axioms) lemma additive_subgroupI: fixes G (structure) assumes a_subgroup: "subgroup H \carrier = carrier G, mult = add G, one = zero G\" shows "additive_subgroup H G" by (rule additive_subgroup.intro) (rule a_subgroup) lemma (in additive_subgroup) a_subset: "H \ carrier G" by (rule subgroup.subset[OF a_subgroup, simplified monoid_record_simps]) lemma (in additive_subgroup) a_closed [intro, simp]: "\x \ H; y \ H\ \ x \ y \ H" by (rule subgroup.m_closed[OF a_subgroup, simplified monoid_record_simps]) lemma (in additive_subgroup) zero_closed [simp]: "\ \ H" by (rule subgroup.one_closed[OF a_subgroup, simplified monoid_record_simps]) lemma (in additive_subgroup) a_inv_closed [intro,simp]: "x \ H \ \ x \ H" by (rule subgroup.m_inv_closed[OF a_subgroup, folded a_inv_def, simplified monoid_record_simps]) subsubsection {* Additive subgroups are normal *} text {* Every subgroup of an @{text "abelian_group"} is normal *} locale abelian_subgroup = additive_subgroup + abelian_group G + assumes a_normal: "normal H \carrier = carrier G, mult = add G, one = zero G\" lemma (in abelian_subgroup) is_abelian_subgroup: shows "abelian_subgroup H G" by (rule abelian_subgroup_axioms) lemma abelian_subgroupI: assumes a_normal: "normal H \carrier = carrier G, mult = add G, one = zero G\" and a_comm: "!!x y. [| x \ carrier G; y \ carrier G |] ==> x \\<^bsub>G\<^esub> y = y \\<^bsub>G\<^esub> x" shows "abelian_subgroup H G" proof - interpret normal "H" "\carrier = carrier G, mult = add G, one = zero G\" by (rule a_normal) show "abelian_subgroup H G" proof qed (simp add: a_comm) qed lemma abelian_subgroupI2: fixes G (structure) assumes a_comm_group: "comm_group \carrier = carrier G, mult = add G, one = zero G\" and a_subgroup: "subgroup H \carrier = carrier G, mult = add G, one = zero G\" shows "abelian_subgroup H G" proof - interpret comm_group "\carrier = carrier G, mult = add G, one = zero G\" by (rule a_comm_group) interpret subgroup "H" "\carrier = carrier G, mult = add G, one = zero G\" by (rule a_subgroup) show "abelian_subgroup H G" apply unfold_locales proof (simp add: r_coset_def l_coset_def, clarsimp) fix x assume xcarr: "x \ carrier G" from a_subgroup have Hcarr: "H \ carrier G" by (unfold subgroup_def, simp) from xcarr Hcarr show "(\h\H. {h \\<^bsub>G\<^esub> x}) = (\h\H. {x \\<^bsub>G\<^esub> h})" using m_comm[simplified] by fast qed qed lemma abelian_subgroupI3: fixes G (structure) assumes asg: "additive_subgroup H G" and ag: "abelian_group G" shows "abelian_subgroup H G" apply (rule abelian_subgroupI2) apply (rule abelian_group.a_comm_group[OF ag]) apply (rule additive_subgroup.a_subgroup[OF asg]) done lemma (in abelian_subgroup) a_coset_eq: "(\x \ carrier G. H +> x = x <+ H)" by (rule normal.coset_eq[OF a_normal, folded a_r_coset_def a_l_coset_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_inv_op_closed1: shows "\x \ carrier G; h \ H\ \ (\ x) \ h \ x \ H" by (rule normal.inv_op_closed1 [OF a_normal, folded a_inv_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_inv_op_closed2: shows "\x \ carrier G; h \ H\ \ x \ h \ (\ x) \ H" by (rule normal.inv_op_closed2 [OF a_normal, folded a_inv_def, simplified monoid_record_simps]) text{*Alternative characterization of normal subgroups*} lemma (in abelian_group) a_normal_inv_iff: "(N \ \carrier = carrier G, mult = add G, one = zero G\) = (subgroup N \carrier = carrier G, mult = add G, one = zero G\ & (\x \ carrier G. \h \ N. x \ h \ (\ x) \ N))" (is "_ = ?rhs") by (rule group.normal_inv_iff [OF a_group, folded a_inv_def, simplified monoid_record_simps]) lemma (in abelian_group) a_lcos_m_assoc: "[| M \ carrier G; g \ carrier G; h \ carrier G |] ==> g <+ (h <+ M) = (g \ h) <+ M" by (rule group.lcos_m_assoc [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_lcos_mult_one: "M \ carrier G ==> \ <+ M = M" by (rule group.lcos_mult_one [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_l_coset_subset_G: "[| H \ carrier G; x \ carrier G |] ==> x <+ H \ carrier G" by (rule group.l_coset_subset_G [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_l_coset_swap: "\y \ x <+ H; x \ carrier G; subgroup H \carrier = carrier G, mult = add G, one = zero G\\ \ x \ y <+ H" by (rule group.l_coset_swap [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_l_coset_carrier: "[| y \ x <+ H; x \ carrier G; subgroup H \carrier = carrier G, mult = add G, one = zero G\ |] ==> y \ carrier G" by (rule group.l_coset_carrier [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_l_repr_imp_subset: assumes y: "y \ x <+ H" and x: "x \ carrier G" and sb: "subgroup H \carrier = carrier G, mult = add G, one = zero G\" shows "y <+ H \ x <+ H" apply (rule group.l_repr_imp_subset [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) apply (rule y) apply (rule x) apply (rule sb) done lemma (in abelian_group) a_l_repr_independence: assumes y: "y \ x <+ H" and x: "x \ carrier G" and sb: "subgroup H \carrier = carrier G, mult = add G, one = zero G\" shows "x <+ H = y <+ H" apply (rule group.l_repr_independence [OF a_group, folded a_l_coset_def, simplified monoid_record_simps]) apply (rule y) apply (rule x) apply (rule sb) done lemma (in abelian_group) setadd_subset_G: "\H \ carrier G; K \ carrier G\ \ H <+> K \ carrier G" by (rule group.setmult_subset_G [OF a_group, folded set_add_def, simplified monoid_record_simps]) lemma (in abelian_group) subgroup_add_id: "subgroup H \carrier = carrier G, mult = add G, one = zero G\ \ H <+> H = H" by (rule group.subgroup_mult_id [OF a_group, folded set_add_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_rcos_inv: assumes x: "x \ carrier G" shows "a_set_inv (H +> x) = H +> (\ x)" by (rule normal.rcos_inv [OF a_normal, folded a_r_coset_def a_inv_def A_SET_INV_def, simplified monoid_record_simps]) (rule x) lemma (in abelian_group) a_setmult_rcos_assoc: "\H \ carrier G; K \ carrier G; x \ carrier G\ \ H <+> (K +> x) = (H <+> K) +> x" by (rule group.setmult_rcos_assoc [OF a_group, folded set_add_def a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_group) a_rcos_assoc_lcos: "\H \ carrier G; K \ carrier G; x \ carrier G\ \ (H +> x) <+> K = H <+> (x <+ K)" by (rule group.rcos_assoc_lcos [OF a_group, folded set_add_def a_r_coset_def a_l_coset_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_rcos_sum: "\x \ carrier G; y \ carrier G\ \ (H +> x) <+> (H +> y) = H +> (x \ y)" by (rule normal.rcos_sum [OF a_normal, folded set_add_def a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) rcosets_add_eq: "M \ a_rcosets H \ H <+> M = M" -- {* generalizes @{text subgroup_mult_id} *} by (rule normal.rcosets_mult_eq [OF a_normal, folded set_add_def A_RCOSETS_def, simplified monoid_record_simps]) subsubsection {* Congruence Relation *} lemma (in abelian_subgroup) a_equiv_rcong: shows "equiv (carrier G) (racong H)" by (rule subgroup.equiv_rcong [OF a_subgroup a_group, folded a_r_congruent_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_l_coset_eq_rcong: assumes a: "a \ carrier G" shows "a <+ H = racong H `` {a}" by (rule subgroup.l_coset_eq_rcong [OF a_subgroup a_group, folded a_r_congruent_def a_l_coset_def, simplified monoid_record_simps]) (rule a) lemma (in abelian_subgroup) a_rcos_equation: shows "\ha \ a = h \ b; a \ carrier G; b \ carrier G; h \ H; ha \ H; hb \ H\ \ hb \ a \ (\h\H. {h \ b})" by (rule group.rcos_equation [OF a_group a_subgroup, folded a_r_congruent_def a_l_coset_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_rcos_disjoint: shows "\a \ a_rcosets H; b \ a_rcosets H; a\b\ \ a \ b = {}" by (rule group.rcos_disjoint [OF a_group a_subgroup, folded A_RCOSETS_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_rcos_self: shows "x \ carrier G \ x \ H +> x" by (rule group.rcos_self [OF a_group _ a_subgroup, folded a_r_coset_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_rcosets_part_G: shows "\(a_rcosets H) = carrier G" by (rule group.rcosets_part_G [OF a_group a_subgroup, folded A_RCOSETS_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_cosets_finite: "\c \ a_rcosets H; H \ carrier G; finite (carrier G)\ \ finite c" by (rule group.cosets_finite [OF a_group, folded A_RCOSETS_def, simplified monoid_record_simps]) lemma (in abelian_group) a_card_cosets_equal: "\c \ a_rcosets H; H \ carrier G; finite(carrier G)\ \ card c = card H" by (rule group.card_cosets_equal [OF a_group, folded A_RCOSETS_def, simplified monoid_record_simps]) lemma (in abelian_group) rcosets_subset_PowG: "additive_subgroup H G \ a_rcosets H \ Pow(carrier G)" by (rule group.rcosets_subset_PowG [OF a_group, folded A_RCOSETS_def, simplified monoid_record_simps], rule additive_subgroup.a_subgroup) theorem (in abelian_group) a_lagrange: "\finite(carrier G); additive_subgroup H G\ \ card(a_rcosets H) * card(H) = order(G)" by (rule group.lagrange [OF a_group, folded A_RCOSETS_def, simplified monoid_record_simps order_def, folded order_def]) (fast intro!: additive_subgroup.a_subgroup)+ subsubsection {* Factorization *} lemmas A_FactGroup_defs = A_FactGroup_def FactGroup_def lemma A_FactGroup_def': fixes G (structure) shows "G A_Mod H \ \carrier = a_rcosets\<^bsub>G\<^esub> H, mult = set_add G, one = H\" unfolding A_FactGroup_defs by (fold A_RCOSETS_def set_add_def) lemma (in abelian_subgroup) a_setmult_closed: "\K1 \ a_rcosets H; K2 \ a_rcosets H\ \ K1 <+> K2 \ a_rcosets H" by (rule normal.setmult_closed [OF a_normal, folded A_RCOSETS_def set_add_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_setinv_closed: "K \ a_rcosets H \ a_set_inv K \ a_rcosets H" by (rule normal.setinv_closed [OF a_normal, folded A_RCOSETS_def A_SET_INV_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_rcosets_assoc: "\M1 \ a_rcosets H; M2 \ a_rcosets H; M3 \ a_rcosets H\ \ M1 <+> M2 <+> M3 = M1 <+> (M2 <+> M3)" by (rule normal.rcosets_assoc [OF a_normal, folded A_RCOSETS_def set_add_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_subgroup_in_rcosets: "H \ a_rcosets H" by (rule subgroup.subgroup_in_rcosets [OF a_subgroup a_group, folded A_RCOSETS_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_rcosets_inv_mult_group_eq: "M \ a_rcosets H \ a_set_inv M <+> M = H" by (rule normal.rcosets_inv_mult_group_eq [OF a_normal, folded A_RCOSETS_def A_SET_INV_def set_add_def, simplified monoid_record_simps]) theorem (in abelian_subgroup) a_factorgroup_is_group: "group (G A_Mod H)" by (rule normal.factorgroup_is_group [OF a_normal, folded A_FactGroup_def, simplified monoid_record_simps]) text {* Since the Factorization is based on an \emph{abelian} subgroup, is results in a commutative group *} theorem (in abelian_subgroup) a_factorgroup_is_comm_group: "comm_group (G A_Mod H)" apply (intro comm_group.intro comm_monoid.intro) prefer 3 apply (rule a_factorgroup_is_group) apply (rule group.axioms[OF a_factorgroup_is_group]) apply (rule comm_monoid_axioms.intro) apply (unfold A_FactGroup_def FactGroup_def RCOSETS_def, fold set_add_def a_r_coset_def, clarsimp) apply (simp add: a_rcos_sum a_comm) done lemma add_A_FactGroup [simp]: "X \\<^bsub>(G A_Mod H)\<^esub> X' = X <+>\<^bsub>G\<^esub> X'" by (simp add: A_FactGroup_def set_add_def) lemma (in abelian_subgroup) a_inv_FactGroup: "X \ carrier (G A_Mod H) \ inv\<^bsub>G A_Mod H\<^esub> X = a_set_inv X" by (rule normal.inv_FactGroup [OF a_normal, folded A_FactGroup_def A_SET_INV_def, simplified monoid_record_simps]) text{*The coset map is a homomorphism from @{term G} to the quotient group @{term "G Mod H"}*} lemma (in abelian_subgroup) a_r_coset_hom_A_Mod: "(\a. H +> a) \ hom \carrier = carrier G, mult = add G, one = zero G\ (G A_Mod H)" by (rule normal.r_coset_hom_Mod [OF a_normal, folded A_FactGroup_def a_r_coset_def, simplified monoid_record_simps]) text {* The isomorphism theorems have been omitted from lifting, at least for now *} subsubsection{*The First Isomorphism Theorem*} text{*The quotient by the kernel of a homomorphism is isomorphic to the range of that homomorphism.*} lemmas a_kernel_defs = a_kernel_def kernel_def lemma a_kernel_def': "a_kernel R S h \ {x \ carrier R. h x = \\<^bsub>S\<^esub>}" by (rule a_kernel_def[unfolded kernel_def, simplified ring_record_simps]) subsubsection {* Homomorphisms *} lemma abelian_group_homI: assumes "abelian_group G" assumes "abelian_group H" assumes a_group_hom: "group_hom (| carrier = carrier G, mult = add G, one = zero G |) (| carrier = carrier H, mult = add H, one = zero H |) h" shows "abelian_group_hom G H h" proof - - interpret G!: abelian_group G by fact - interpret H!: abelian_group H by fact + interpret G: abelian_group G by fact + interpret H: abelian_group H by fact show ?thesis apply (intro abelian_group_hom.intro abelian_group_hom_axioms.intro) apply fact apply fact apply (rule a_group_hom) done qed lemma (in abelian_group_hom) is_abelian_group_hom: "abelian_group_hom G H h" .. lemma (in abelian_group_hom) hom_add [simp]: "[| x : carrier G; y : carrier G |] ==> h (x \\<^bsub>G\<^esub> y) = h x \\<^bsub>H\<^esub> h y" by (rule group_hom.hom_mult[OF a_group_hom, simplified ring_record_simps]) lemma (in abelian_group_hom) hom_closed [simp]: "x \ carrier G \ h x \ carrier H" by (rule group_hom.hom_closed[OF a_group_hom, simplified ring_record_simps]) lemma (in abelian_group_hom) zero_closed [simp]: "h \ \ carrier H" by (rule group_hom.one_closed[OF a_group_hom, simplified ring_record_simps]) lemma (in abelian_group_hom) hom_zero [simp]: "h \ = \\<^bsub>H\<^esub>" by (rule group_hom.hom_one[OF a_group_hom, simplified ring_record_simps]) lemma (in abelian_group_hom) a_inv_closed [simp]: "x \ carrier G ==> h (\x) \ carrier H" by (rule group_hom.inv_closed[OF a_group_hom, folded a_inv_def, simplified ring_record_simps]) lemma (in abelian_group_hom) hom_a_inv [simp]: "x \ carrier G ==> h (\x) = \\<^bsub>H\<^esub> (h x)" by (rule group_hom.hom_inv[OF a_group_hom, folded a_inv_def, simplified ring_record_simps]) lemma (in abelian_group_hom) additive_subgroup_a_kernel: "additive_subgroup (a_kernel G H h) G" apply (rule additive_subgroup.intro) apply (rule group_hom.subgroup_kernel[OF a_group_hom, folded a_kernel_def, simplified ring_record_simps]) done text{*The kernel of a homomorphism is an abelian subgroup*} lemma (in abelian_group_hom) abelian_subgroup_a_kernel: "abelian_subgroup (a_kernel G H h) G" apply (rule abelian_subgroupI) apply (rule group_hom.normal_kernel[OF a_group_hom, folded a_kernel_def, simplified ring_record_simps]) apply (simp add: G.a_comm) done lemma (in abelian_group_hom) A_FactGroup_nonempty: assumes X: "X \ carrier (G A_Mod a_kernel G H h)" shows "X \ {}" by (rule group_hom.FactGroup_nonempty[OF a_group_hom, folded a_kernel_def A_FactGroup_def, simplified ring_record_simps]) (rule X) lemma (in abelian_group_hom) FactGroup_contents_mem: assumes X: "X \ carrier (G A_Mod (a_kernel G H h))" shows "contents (h`X) \ carrier H" by (rule group_hom.FactGroup_contents_mem[OF a_group_hom, folded a_kernel_def A_FactGroup_def, simplified ring_record_simps]) (rule X) lemma (in abelian_group_hom) A_FactGroup_hom: "(\X. contents (h`X)) \ hom (G A_Mod (a_kernel G H h)) \carrier = carrier H, mult = add H, one = zero H\" by (rule group_hom.FactGroup_hom[OF a_group_hom, folded a_kernel_def A_FactGroup_def, simplified ring_record_simps]) lemma (in abelian_group_hom) A_FactGroup_inj_on: "inj_on (\X. contents (h ` X)) (carrier (G A_Mod a_kernel G H h))" by (rule group_hom.FactGroup_inj_on[OF a_group_hom, folded a_kernel_def A_FactGroup_def, simplified ring_record_simps]) text{*If the homomorphism @{term h} is onto @{term H}, then so is the homomorphism from the quotient group*} lemma (in abelian_group_hom) A_FactGroup_onto: assumes h: "h ` carrier G = carrier H" shows "(\X. contents (h ` X)) ` carrier (G A_Mod a_kernel G H h) = carrier H" by (rule group_hom.FactGroup_onto[OF a_group_hom, folded a_kernel_def A_FactGroup_def, simplified ring_record_simps]) (rule h) text{*If @{term h} is a homomorphism from @{term G} onto @{term H}, then the quotient group @{term "G Mod (kernel G H h)"} is isomorphic to @{term H}.*} theorem (in abelian_group_hom) A_FactGroup_iso: "h ` carrier G = carrier H \ (\X. contents (h`X)) \ (G A_Mod (a_kernel G H h)) \ (| carrier = carrier H, mult = add H, one = zero H |)" by (rule group_hom.FactGroup_iso[OF a_group_hom, folded a_kernel_def A_FactGroup_def, simplified ring_record_simps]) subsubsection {* Cosets *} text {* Not eveything from \texttt{CosetExt.thy} is lifted here. *} lemma (in additive_subgroup) a_Hcarr [simp]: assumes hH: "h \ H" shows "h \ carrier G" by (rule subgroup.mem_carrier [OF a_subgroup, simplified monoid_record_simps]) (rule hH) lemma (in abelian_subgroup) a_elemrcos_carrier: assumes acarr: "a \ carrier G" and a': "a' \ H +> a" shows "a' \ carrier G" by (rule subgroup.elemrcos_carrier [OF a_subgroup a_group, folded a_r_coset_def, simplified monoid_record_simps]) (rule acarr, rule a') lemma (in abelian_subgroup) a_rcos_const: assumes hH: "h \ H" shows "H +> h = H" by (rule subgroup.rcos_const [OF a_subgroup a_group, folded a_r_coset_def, simplified monoid_record_simps]) (rule hH) lemma (in abelian_subgroup) a_rcos_module_imp: assumes xcarr: "x \ carrier G" and x'cos: "x' \ H +> x" shows "(x' \ \x) \ H" by (rule subgroup.rcos_module_imp [OF a_subgroup a_group, folded a_r_coset_def a_inv_def, simplified monoid_record_simps]) (rule xcarr, rule x'cos) lemma (in abelian_subgroup) a_rcos_module_rev: assumes "x \ carrier G" "x' \ carrier G" and "(x' \ \x) \ H" shows "x' \ H +> x" using assms by (rule subgroup.rcos_module_rev [OF a_subgroup a_group, folded a_r_coset_def a_inv_def, simplified monoid_record_simps]) lemma (in abelian_subgroup) a_rcos_module: assumes "x \ carrier G" "x' \ carrier G" shows "(x' \ H +> x) = (x' \ \x \ H)" using assms by (rule subgroup.rcos_module [OF a_subgroup a_group, folded a_r_coset_def a_inv_def, simplified monoid_record_simps]) --"variant" lemma (in abelian_subgroup) a_rcos_module_minus: assumes "ring G" assumes carr: "x \ carrier G" "x' \ carrier G" shows "(x' \ H +> x) = (x' \ x \ H)" proof - - interpret G!: ring G by fact + interpret G: ring G by fact from carr have "(x' \ H +> x) = (x' \ \x \ H)" by (rule a_rcos_module) with carr show "(x' \ H +> x) = (x' \ x \ H)" by (simp add: minus_eq) qed lemma (in abelian_subgroup) a_repr_independence': assumes y: "y \ H +> x" and xcarr: "x \ carrier G" shows "H +> x = H +> y" apply (rule a_repr_independence) apply (rule y) apply (rule xcarr) apply (rule a_subgroup) done lemma (in abelian_subgroup) a_repr_independenceD: assumes ycarr: "y \ carrier G" and repr: "H +> x = H +> y" shows "y \ H +> x" by (rule group.repr_independenceD [OF a_group a_subgroup, folded a_r_coset_def, simplified monoid_record_simps]) (rule ycarr, rule repr) lemma (in abelian_subgroup) a_rcosets_carrier: "X \ a_rcosets H \ X \ carrier G" by (rule subgroup.rcosets_carrier [OF a_subgroup a_group, folded A_RCOSETS_def, simplified monoid_record_simps]) subsubsection {* Addition of Subgroups *} lemma (in abelian_monoid) set_add_closed: assumes Acarr: "A \ carrier G" and Bcarr: "B \ carrier G" shows "A <+> B \ carrier G" by (rule monoid.set_mult_closed [OF a_monoid, folded set_add_def, simplified monoid_record_simps]) (rule Acarr, rule Bcarr) lemma (in abelian_group) add_additive_subgroups: assumes subH: "additive_subgroup H G" and subK: "additive_subgroup K G" shows "additive_subgroup (H <+> K) G" apply (rule additive_subgroup.intro) apply (unfold set_add_def) apply (intro comm_group.mult_subgroups) apply (rule a_comm_group) apply (rule additive_subgroup.a_subgroup[OF subH]) apply (rule additive_subgroup.a_subgroup[OF subK]) done 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,808 +1,808 @@ (* Title: HOL/Algebra/Group.thy Author: Clemens Ballarin, started 4 February 2003 Based on work by Florian Kammueller, L C Paulson and Markus Wenzel. *) theory Group imports Lattice FuncSet begin section {* Monoids and Groups *} subsection {* Definitions *} text {* Definitions follow \cite{Jacobson:1985}. *} record 'a monoid = "'a partial_object" + mult :: "['a, 'a] \ 'a" (infixl "\\" 70) one :: 'a ("\\") constdefs (structure G) m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" ("inv\ _" [81] 80) "inv x == (THE y. y \ carrier G & x \ y = \ & y \ x = \)" Units :: "_ => 'a set" --{*The set of invertible elements*} "Units G == {y. y \ carrier G & (\x \ carrier G. x \ y = \ & y \ x = \)}" consts pow :: "[('a, 'm) monoid_scheme, 'a, 'b::number] => 'a" (infixr "'(^')\" 75) defs (overloaded) nat_pow_def: "pow G a n == nat_rec \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) n" int_pow_def: "pow G a z == let p = nat_rec \\<^bsub>G\<^esub> (%u b. b \\<^bsub>G\<^esub> a) in if neg z then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z)" 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) 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 [intro, simp]: 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 -- "Must avoid premature use of @{text hyp_subst_tac}." apply (rule_tac CollectI) apply (rule) apply (fast) apply (rule bexI [where x = "y' \ x'"]) apply (auto simp: m_assoc) done 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 (unfold Units_def m_inv_def, auto) apply (rule theI2, fast) apply (fast intro: inv_unique, fast) done 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, auto) apply (rule theI2, fast) apply (fast intro: inv_unique, fast) done lemma (in monoid) Units_r_inv [simp]: "x \ Units G ==> x \ inv x = \" apply (unfold Units_def m_inv_def, auto) apply (rule theI2, fast) apply (fast intro: inv_unique, fast) done 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 add: Units_l_inv) 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 text {* Power *} 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_pow: "x \ carrier G ==> (x (^) n) (^) m = x (^) (n * m::nat)" by (induct m) (simp, simp add: nat_pow_mult add_commute) (* 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: "group G" by (rule group_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 proof qed (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 = \" using Units_l_inv by simp subsection {* Cancellation Laws and Basic Properties *} lemma (in group) l_cancel [simp]: "[| x \ carrier G; y \ carrier G; z \ carrier G |] ==> (x \ y = x \ z) = (y = z)" using Units_l_inv by simp lemma (in group) r_inv [simp]: "x \ carrier G ==> x \ inv x = \" proof - assume x: "x \ carrier G" then have "inv x \ (x \ inv x) = inv x \ \" by (simp add: m_assoc [symmetric] l_inv) with x show ?thesis by (simp del: r_one) qed lemma (in group) r_cancel [simp]: "[| x \ carrier G; y \ carrier G; z \ carrier G |] ==> (y \ x = z \ x) = (y = z)" proof assume eq: "y \ x = z \ x" and G: "x \ carrier G" "y \ carrier G" "z \ carrier G" then have "y \ (x \ inv x) = z \ (x \ inv x)" by (simp add: m_assoc [symmetric] del: r_inv Units_r_inv) with G show "y = z" by simp next assume eq: "y = z" and G: "x \ carrier G" "y \ carrier G" "z \ carrier G" then show "y \ x = z \ x" by simp qed lemma (in group) inv_one [simp]: "inv \ = \" proof - have "inv \ = \ \ (inv \)" by (simp del: r_inv Units_r_inv) moreover have "... = \" by simp finally show ?thesis . qed 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 l_inv) (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" apply (simp add: m_inv_def) apply (rule the_equality) apply (simp add: inv_comm [of y x]) apply (rule r_cancel [THEN iffD1], auto) done text {* Power *} lemma (in group) int_pow_def2: "a (^) (z::int) = (if neg z then inv (a (^) (nat (-z))) else a (^) (nat z))" by (simp add: int_pow_def nat_pow_def Let_def) lemma (in group) int_pow_0 [simp]: "x (^) (0::int) = \" by (simp add: int_pow_def2) lemma (in group) int_pow_one [simp]: "\ (^) (z::int) = \" by (simp add: int_pow_def2) 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 subgroup_imp_subset: "subgroup H G \ H \ carrier G" by (rule subgroup.subset) lemma (in subgroup) subgroup_is_group [intro]: assumes "group G" shows "group (G\carrier := H\)" proof - interpret group G by fact show ?thesis apply (rule monoid.group_l_invI) apply (unfold_locales) [1] apply (auto intro: m_assoc l_inv mem_carrier) done qed text {* Since @{term H} is nonempty, it contains some element @{term x}. Since it is closed under inverse, it contains @{text "inv x"}. Since it is closed under product, it contains @{text "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 simp add: l_inv) 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 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" proof (rule classical) assume "finite (carrier G)" "~ 0 < card H" then have "finite H" by (blast intro: finite_subset [OF subset]) with prems have "subgroup {} G" by simp with subgroup_nonempty show ?thesis by contradiction qed (* lemma (in monoid) Units_subgroup: "subgroup (Units G) G" *) subsection {* Direct Products *} constdefs DirProd :: "_ \ _ \ ('a \ 'b) monoid" (infixr "\\" 80) "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 + 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 + 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 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" + 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 subsection {* Homomorphisms and Isomorphisms *} constdefs (structure G and H) hom :: "_ => _ => ('a => 'b) set" "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 (in group) hom_compose: "[|h \ hom G H; i \ hom H I|] ==> compose (carrier G) i h \ hom G I" apply (auto simp add: hom_def funcset_compose) apply (simp add: compose_def funcset_mem) done constdefs iso :: "_ => _ => ('a => 'b) set" (infixr "\" 60) "G \ H == {h. h \ hom G H & bij_betw h (carrier G) (carrier H)}" lemma iso_refl: "(%x. x) \ G \ G" by (simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) lemma (in group) iso_sym: "h \ G \ H \ Inv (carrier G) h \ H \ G" apply (simp add: iso_def bij_betw_Inv) apply (subgoal_tac "Inv (carrier G) h \ carrier H \ carrier G") prefer 2 apply (simp add: bij_betw_imp_funcset [OF bij_betw_Inv]) apply (simp add: hom_def bij_betw_def Inv_f_eq funcset_mem f_Inv_f) done lemma (in group) iso_trans: "[|h \ G \ H; i \ H \ I|] ==> (compose (carrier G) i h) \ G \ I" by (auto simp add: iso_def hom_compose bij_betw_compose) lemma DirProd_commute_iso: shows "(\(x,y). (y,x)) \ (G \\ H) \ (H \\ G)" by (auto simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) lemma DirProd_assoc_iso: shows "(\(x,y,z). (x,(y,z))) \ (G \\ H \\ I) \ (G \\ (H \\ I))" by (auto simp add: iso_def hom_def inj_on_def bij_betw_def Pi_def) 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: "h \ hom G H" 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 simp add: funcset_mem) qed lemma (in group_hom) one_closed [simp]: "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 (simp del: r_one) qed lemma (in group_hom) inv_closed [simp]: "x \ carrier G ==> h (inv x) \ carrier H" by simp lemma (in group_hom) hom_inv [simp]: "x \ carrier G ==> h (inv x) = inv\<^bsub>H\<^esub> (h x)" proof - assume x: "x \ carrier G" then have "h x \\<^bsub>H\<^esub> h (inv x) = \\<^bsub>H\<^esub>" by (simp add: hom_mult [symmetric] del: hom_mult) also from x have "... = h x \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h x)" by (simp add: hom_mult [symmetric] del: hom_mult) finally have "h x \\<^bsub>H\<^esub> h (inv x) = h x \\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h x)" . with x show ?thesis by (simp del: H.r_inv H.Units_r_inv) qed 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) r_one [simp]: "x \ carrier G ==> x \ \ = x" proof - assume G: "x \ carrier G" then have "x \ \ = \ \ x" by (simp add: m_comm) also from G have "... = x" by simp finally show ?thesis . qed*) lemma (in comm_monoid) nat_pow_distr: "[| x \ carrier G; y \ carrier G |] ==> (x \ y) (^) (n::nat) = x (^) n \ y (^) n" by (induct n) (simp, simp add: m_ac) 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" proof qed (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 (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) 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 = op =, le = op \ |)" proof qed 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) is_monoid [intro, simp]: "monoid G" by (auto intro: monoid.intro m_assoc) lemma (in group) subgroup_inv_equality: "[| subgroup H G; x \ H |] ==> m_inv (G (| carrier := H |)) x = inv x" apply (rule_tac inv_equality [THEN sym]) apply (rule group.l_inv [OF subgroup_imp_group, simplified], assumption+) apply (rule subsetD [OF subgroup.subset], assumption+) apply (rule subsetD [OF subgroup.subset], assumption) apply (rule_tac group.inv_closed [OF subgroup_imp_group, simplified], assumption+) done 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 theorem (in group) subgroups_complete_lattice: "complete_lattice (| carrier = {H. subgroup H G}, eq = op =, le = op \ |)" (is "complete_lattice ?L") proof (rule partial_order.complete_lattice_criterion1) show "partial_order ?L" by (rule subgroups_partial_order) next show "\G. greatest ?L G (carrier ?L)" proof show "greatest ?L (carrier G) (carrier ?L)" by (unfold greatest_def) (simp add: subgroup.subset subgroup_self) qed next fix A assume L: "A \ carrier ?L" and non_empty: "A ~= {}" then have Int_subgroup: "subgroup (\A) G" by (fastsimp intro: subgroups_Inter) show "\I. greatest ?L I (Lower ?L A)" proof show "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 fastsimp 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 (fastsimp 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 qed qed end diff --git a/src/HOL/Algebra/Ideal.thy b/src/HOL/Algebra/Ideal.thy --- a/src/HOL/Algebra/Ideal.thy +++ b/src/HOL/Algebra/Ideal.thy @@ -1,1035 +1,1035 @@ (* Title: HOL/Algebra/CIdeal.thy Author: Stephan Hohe, TU Muenchen *) theory Ideal imports Ring AbelCoset begin section {* Ideals *} subsection {* Definitions *} subsubsection {* General definition *} locale ideal = additive_subgroup I R + ring R for I and R (structure) + assumes I_l_closed: "\a \ I; x \ carrier R\ \ x \ a \ I" and I_r_closed: "\a \ I; x \ carrier R\ \ a \ x \ I" sublocale ideal \ abelian_subgroup I R apply (intro abelian_subgroupI3 abelian_group.intro) apply (rule ideal.axioms, rule ideal_axioms) apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms) apply (rule abelian_group.axioms, rule ring.axioms, rule ideal.axioms, rule ideal_axioms) done lemma (in ideal) is_ideal: "ideal I R" by (rule ideal_axioms) lemma idealI: fixes R (structure) assumes "ring R" assumes a_subgroup: "subgroup I \carrier = carrier R, mult = add R, one = zero R\" and I_l_closed: "\a x. \a \ I; x \ carrier R\ \ x \ a \ I" and I_r_closed: "\a x. \a \ I; x \ carrier R\ \ a \ x \ I" shows "ideal I R" proof - interpret ring R by fact show ?thesis apply (intro ideal.intro ideal_axioms.intro additive_subgroupI) apply (rule a_subgroup) apply (rule is_ring) apply (erule (1) I_l_closed) apply (erule (1) I_r_closed) done qed subsubsection (in ring) {* Ideals Generated by a Subset of @{term "carrier R"} *} constdefs (structure R) genideal :: "('a, 'b) ring_scheme \ 'a set \ 'a set" ("Idl\ _" [80] 79) "genideal R S \ Inter {I. ideal I R \ S \ I}" subsubsection {* Principal Ideals *} locale principalideal = ideal + assumes generate: "\i \ carrier R. I = Idl {i}" lemma (in principalideal) is_principalideal: shows "principalideal I R" by (rule principalideal_axioms) lemma principalidealI: fixes R (structure) assumes "ideal I R" assumes generate: "\i \ carrier R. I = Idl {i}" shows "principalideal I R" proof - interpret ideal I R by fact show ?thesis by (intro principalideal.intro principalideal_axioms.intro) (rule is_ideal, rule generate) qed subsubsection {* Maximal Ideals *} locale maximalideal = ideal + assumes I_notcarr: "carrier R \ I" and I_maximal: "\ideal J R; I \ J; J \ carrier R\ \ J = I \ J = carrier R" lemma (in maximalideal) is_maximalideal: shows "maximalideal I R" by (rule maximalideal_axioms) lemma maximalidealI: fixes R assumes "ideal I R" assumes I_notcarr: "carrier R \ I" and I_maximal: "\J. \ideal J R; I \ J; J \ carrier R\ \ J = I \ J = carrier R" shows "maximalideal I R" proof - interpret ideal I R by fact show ?thesis by (intro maximalideal.intro maximalideal_axioms.intro) (rule is_ideal, rule I_notcarr, rule I_maximal) qed subsubsection {* Prime Ideals *} locale primeideal = ideal + cring + assumes I_notcarr: "carrier R \ I" and I_prime: "\a \ carrier R; b \ carrier R; a \ b \ I\ \ a \ I \ b \ I" lemma (in primeideal) is_primeideal: shows "primeideal I R" by (rule primeideal_axioms) lemma primeidealI: fixes R (structure) assumes "ideal I R" assumes "cring R" assumes I_notcarr: "carrier R \ I" and I_prime: "\a b. \a \ carrier R; b \ carrier R; a \ b \ I\ \ a \ I \ b \ I" shows "primeideal I R" proof - interpret ideal I R by fact interpret cring R by fact show ?thesis by (intro primeideal.intro primeideal_axioms.intro) (rule is_ideal, rule is_cring, rule I_notcarr, rule I_prime) qed lemma primeidealI2: fixes R (structure) assumes "additive_subgroup I R" assumes "cring R" assumes I_l_closed: "\a x. \a \ I; x \ carrier R\ \ x \ a \ I" and I_r_closed: "\a x. \a \ I; x \ carrier R\ \ a \ x \ I" and I_notcarr: "carrier R \ I" and I_prime: "\a b. \a \ carrier R; b \ carrier R; a \ b \ I\ \ a \ I \ b \ I" shows "primeideal I R" proof - interpret additive_subgroup I R by fact interpret cring R by fact show ?thesis apply (intro_locales) apply (intro ideal_axioms.intro) apply (erule (1) I_l_closed) apply (erule (1) I_r_closed) apply (intro primeideal_axioms.intro) apply (rule I_notcarr) apply (erule (2) I_prime) done qed subsection {* Special Ideals *} lemma (in ring) zeroideal: shows "ideal {\} R" apply (intro idealI subgroup.intro) apply (rule is_ring) apply simp+ apply (fold a_inv_def, simp) apply simp+ done lemma (in ring) oneideal: shows "ideal (carrier R) R" apply (intro idealI subgroup.intro) apply (rule is_ring) apply simp+ apply (fold a_inv_def, simp) apply simp+ done lemma (in "domain") zeroprimeideal: shows "primeideal {\} R" apply (intro primeidealI) apply (rule zeroideal) apply (rule domain.axioms, rule domain_axioms) defer 1 apply (simp add: integral) proof (rule ccontr, simp) assume "carrier R = {\}" from this have "\ = \" by (rule one_zeroI) from this and one_not_zero show "False" by simp qed subsection {* General Ideal Properies *} lemma (in ideal) one_imp_carrier: assumes I_one_closed: "\ \ I" shows "I = carrier R" apply (rule) apply (rule) apply (rule a_Hcarr, simp) proof fix x assume xcarr: "x \ carrier R" from I_one_closed and this have "x \ \ \ I" by (intro I_l_closed) from this and xcarr show "x \ I" by simp qed lemma (in ideal) Icarr: assumes iI: "i \ I" shows "i \ carrier R" using iI by (rule a_Hcarr) subsection {* Intersection of Ideals *} text {* \paragraph{Intersection of two ideals} The intersection of any two ideals is again an ideal in @{term R} *} lemma (in ring) i_intersect: assumes "ideal I R" assumes "ideal J R" shows "ideal (I \ J) R" proof - interpret ideal I R by fact interpret ideal J R by fact show ?thesis apply (intro idealI subgroup.intro) apply (rule is_ring) apply (force simp add: a_subset) apply (simp add: a_inv_def[symmetric]) apply simp apply (simp add: a_inv_def[symmetric]) apply (clarsimp, rule) apply (fast intro: ideal.I_l_closed ideal.intro assms)+ apply (clarsimp, rule) apply (fast intro: ideal.I_r_closed ideal.intro assms)+ done qed text {* The intersection of any Number of Ideals is again an Ideal in @{term R} *} lemma (in ring) i_Intersect: assumes Sideals: "\I. I \ S \ ideal I R" and notempty: "S \ {}" shows "ideal (Inter S) R" apply (unfold_locales) apply (simp_all add: Inter_def INTER_def) apply (rule, simp) defer 1 apply rule defer 1 apply rule defer 1 apply (fold a_inv_def, rule) defer 1 apply rule defer 1 apply rule defer 1 proof - fix x assume "\I\S. x \ I" hence xI: "\I. I \ S \ x \ I" by simp from notempty have "\I0. I0 \ S" by blast from this obtain I0 where I0S: "I0 \ S" by auto interpret ideal I0 R by (rule Sideals[OF I0S]) from xI[OF I0S] have "x \ I0" . from this and a_subset show "x \ carrier R" by fast next fix x y assume "\I\S. x \ I" hence xI: "\I. I \ S \ x \ I" by simp assume "\I\S. y \ I" hence yI: "\I. I \ S \ y \ I" by simp fix J assume JS: "J \ S" interpret ideal J R by (rule Sideals[OF JS]) from xI[OF JS] and yI[OF JS] show "x \ y \ J" by (rule a_closed) next fix J assume JS: "J \ S" interpret ideal J R by (rule Sideals[OF JS]) show "\ \ J" by simp next fix x assume "\I\S. x \ I" hence xI: "\I. I \ S \ x \ I" by simp fix J assume JS: "J \ S" interpret ideal J R by (rule Sideals[OF JS]) from xI[OF JS] show "\ x \ J" by (rule a_inv_closed) next fix x y assume "\I\S. x \ I" hence xI: "\I. I \ S \ x \ I" by simp assume ycarr: "y \ carrier R" fix J assume JS: "J \ S" interpret ideal J R by (rule Sideals[OF JS]) from xI[OF JS] and ycarr show "y \ x \ J" by (rule I_l_closed) next fix x y assume "\I\S. x \ I" hence xI: "\I. I \ S \ x \ I" by simp assume ycarr: "y \ carrier R" fix J assume JS: "J \ S" interpret ideal J R by (rule Sideals[OF JS]) from xI[OF JS] and ycarr show "x \ y \ J" by (rule I_r_closed) qed subsection {* Addition of Ideals *} lemma (in ring) add_ideals: assumes idealI: "ideal I R" and idealJ: "ideal J R" shows "ideal (I <+> J) R" apply (rule ideal.intro) apply (rule add_additive_subgroups) apply (intro ideal.axioms[OF idealI]) apply (intro ideal.axioms[OF idealJ]) apply (rule is_ring) apply (rule ideal_axioms.intro) apply (simp add: set_add_defs, clarsimp) defer 1 apply (simp add: set_add_defs, clarsimp) defer 1 proof - fix x i j assume xcarr: "x \ carrier R" and iI: "i \ I" and jJ: "j \ J" from xcarr ideal.Icarr[OF idealI iI] ideal.Icarr[OF idealJ jJ] have c: "(i \ j) \ x = (i \ x) \ (j \ x)" by algebra from xcarr and iI have a: "i \ x \ I" by (simp add: ideal.I_r_closed[OF idealI]) from xcarr and jJ have b: "j \ x \ J" by (simp add: ideal.I_r_closed[OF idealJ]) from a b c show "\ha\I. \ka\J. (i \ j) \ x = ha \ ka" by fast next fix x i j assume xcarr: "x \ carrier R" and iI: "i \ I" and jJ: "j \ J" from xcarr ideal.Icarr[OF idealI iI] ideal.Icarr[OF idealJ jJ] have c: "x \ (i \ j) = (x \ i) \ (x \ j)" by algebra from xcarr and iI have a: "x \ i \ I" by (simp add: ideal.I_l_closed[OF idealI]) from xcarr and jJ have b: "x \ j \ J" by (simp add: ideal.I_l_closed[OF idealJ]) from a b c show "\ha\I. \ka\J. x \ (i \ j) = ha \ ka" by fast qed subsection (in ring) {* Ideals generated by a subset of @{term "carrier R"} *} text {* @{term genideal} generates an ideal *} lemma (in ring) genideal_ideal: assumes Scarr: "S \ carrier R" shows "ideal (Idl S) R" unfolding genideal_def proof (rule i_Intersect, fast, simp) from oneideal and Scarr show "\I. ideal I R \ S \ I" by fast qed lemma (in ring) genideal_self: assumes "S \ carrier R" shows "S \ Idl S" unfolding genideal_def by fast lemma (in ring) genideal_self': assumes carr: "i \ carrier R" shows "i \ Idl {i}" proof - from carr have "{i} \ Idl {i}" by (fast intro!: genideal_self) thus "i \ Idl {i}" by fast qed text {* @{term genideal} generates the minimal ideal *} lemma (in ring) genideal_minimal: assumes a: "ideal I R" and b: "S \ I" shows "Idl S \ I" unfolding genideal_def by (rule, elim InterD, simp add: a b) text {* Generated ideals and subsets *} lemma (in ring) Idl_subset_ideal: assumes Iideal: "ideal I R" and Hcarr: "H \ carrier R" shows "(Idl H \ I) = (H \ I)" proof assume a: "Idl H \ I" from Hcarr have "H \ Idl H" by (rule genideal_self) from this and a show "H \ I" by simp next fix x assume HI: "H \ I" from Iideal and HI have "I \ {I. ideal I R \ H \ I}" by fast from this show "Idl H \ I" unfolding genideal_def by fast qed lemma (in ring) subset_Idl_subset: assumes Icarr: "I \ carrier R" and HI: "H \ I" shows "Idl H \ Idl I" proof - from HI and genideal_self[OF Icarr] have HIdlI: "H \ Idl I" by fast from Icarr have Iideal: "ideal (Idl I) R" by (rule genideal_ideal) from HI and Icarr have "H \ carrier R" by fast from Iideal and this have "(H \ Idl I) = (Idl H \ Idl I)" by (rule Idl_subset_ideal[symmetric]) from HIdlI and this show "Idl H \ Idl I" by simp qed lemma (in ring) Idl_subset_ideal': assumes acarr: "a \ carrier R" and bcarr: "b \ carrier R" shows "(Idl {a} \ Idl {b}) = (a \ Idl {b})" apply (subst Idl_subset_ideal[OF genideal_ideal[of "{b}"], of "{a}"]) apply (fast intro: bcarr, fast intro: acarr) apply fast done lemma (in ring) genideal_zero: "Idl {\} = {\}" apply rule apply (rule genideal_minimal[OF zeroideal], simp) apply (simp add: genideal_self') done lemma (in ring) genideal_one: "Idl {\} = carrier R" proof - interpret ideal "Idl {\}" "R" by (rule genideal_ideal, fast intro: one_closed) show "Idl {\} = carrier R" apply (rule, rule a_subset) apply (simp add: one_imp_carrier genideal_self') done qed text {* Generation of Principal Ideals in Commutative Rings *} constdefs (structure R) cgenideal :: "('a, 'b) monoid_scheme \ 'a \ 'a set" ("PIdl\ _" [80] 79) "cgenideal R a \ { x \ a | x. x \ carrier R }" text {* genhideal (?) really generates an ideal *} lemma (in cring) cgenideal_ideal: assumes acarr: "a \ carrier R" shows "ideal (PIdl a) R" apply (unfold cgenideal_def) apply (rule idealI[OF is_ring]) apply (rule subgroup.intro) apply (simp_all add: monoid_record_simps) apply (blast intro: acarr m_closed) apply clarsimp defer 1 defer 1 apply (fold a_inv_def, clarsimp) defer 1 apply clarsimp defer 1 apply clarsimp defer 1 proof - fix x y assume xcarr: "x \ carrier R" and ycarr: "y \ carrier R" note carr = acarr xcarr ycarr from carr have "x \ a \ y \ a = (x \ y) \ a" by (simp add: l_distr) from this and carr show "\z. x \ a \ y \ a = z \ a \ z \ carrier R" by fast next from l_null[OF acarr, symmetric] and zero_closed show "\x. \ = x \ a \ x \ carrier R" by fast next fix x assume xcarr: "x \ carrier R" note carr = acarr xcarr from carr have "\ (x \ a) = (\ x) \ a" by (simp add: l_minus) from this and carr show "\z. \ (x \ a) = z \ a \ z \ carrier R" by fast next fix x y assume xcarr: "x \ carrier R" and ycarr: "y \ carrier R" note carr = acarr xcarr ycarr from carr have "y \ a \ x = (y \ x) \ a" by (simp add: m_assoc, simp add: m_comm) from this and carr show "\z. y \ a \ x = z \ a \ z \ carrier R" by fast next fix x y assume xcarr: "x \ carrier R" and ycarr: "y \ carrier R" note carr = acarr xcarr ycarr from carr have "x \ (y \ a) = (x \ y) \ a" by (simp add: m_assoc) from this and carr show "\z. x \ (y \ a) = z \ a \ z \ carrier R" by fast qed lemma (in ring) cgenideal_self: assumes icarr: "i \ carrier R" shows "i \ PIdl i" unfolding cgenideal_def proof simp from icarr have "i = \ \ i" by simp from this and icarr show "\x. i = x \ i \ x \ carrier R" by fast qed text {* @{const "cgenideal"} is minimal *} lemma (in ring) cgenideal_minimal: assumes "ideal J R" assumes aJ: "a \ J" shows "PIdl a \ J" proof - interpret ideal J R by fact show ?thesis unfolding cgenideal_def apply rule apply clarify using aJ apply (erule I_l_closed) done qed lemma (in cring) cgenideal_eq_genideal: assumes icarr: "i \ carrier R" shows "PIdl i = Idl {i}" apply rule apply (intro cgenideal_minimal) apply (rule genideal_ideal, fast intro: icarr) apply (rule genideal_self', fast intro: icarr) apply (intro genideal_minimal) apply (rule cgenideal_ideal [OF icarr]) apply (simp, rule cgenideal_self [OF icarr]) done lemma (in cring) cgenideal_eq_rcos: "PIdl i = carrier R #> i" unfolding cgenideal_def r_coset_def by fast lemma (in cring) cgenideal_is_principalideal: assumes icarr: "i \ carrier R" shows "principalideal (PIdl i) R" apply (rule principalidealI) apply (rule cgenideal_ideal [OF icarr]) proof - from icarr have "PIdl i = Idl {i}" by (rule cgenideal_eq_genideal) from icarr and this show "\i'\carrier R. PIdl i = Idl {i'}" by fast qed subsection {* Union of Ideals *} lemma (in ring) union_genideal: assumes idealI: "ideal I R" and idealJ: "ideal J R" shows "Idl (I \ J) = I <+> J" apply rule apply (rule ring.genideal_minimal) apply (rule is_ring) apply (rule add_ideals[OF idealI idealJ]) apply (rule) apply (simp add: set_add_defs) apply (elim disjE) defer 1 defer 1 apply (rule) apply (simp add: set_add_defs genideal_def) apply clarsimp defer 1 proof - fix x assume xI: "x \ I" have ZJ: "\ \ J" by (intro additive_subgroup.zero_closed, rule ideal.axioms[OF idealJ]) from ideal.Icarr[OF idealI xI] have "x = x \ \" by algebra from xI and ZJ and this show "\h\I. \k\J. x = h \ k" by fast next fix x assume xJ: "x \ J" have ZI: "\ \ I" by (intro additive_subgroup.zero_closed, rule ideal.axioms[OF idealI]) from ideal.Icarr[OF idealJ xJ] have "x = \ \ x" by algebra from ZI and xJ and this show "\h\I. \k\J. x = h \ k" by fast next fix i j K assume iI: "i \ I" and jJ: "j \ J" and idealK: "ideal K R" and IK: "I \ K" and JK: "J \ K" from iI and IK have iK: "i \ K" by fast from jJ and JK have jK: "j \ K" by fast from iK and jK show "i \ j \ K" by (intro additive_subgroup.a_closed) (rule ideal.axioms[OF idealK]) qed subsection {* Properties of Principal Ideals *} text {* @{text "\"} generates the zero ideal *} lemma (in ring) zero_genideal: shows "Idl {\} = {\}" apply rule apply (simp add: genideal_minimal zeroideal) apply (fast intro!: genideal_self) done text {* @{text "\"} generates the unit ideal *} lemma (in ring) one_genideal: shows "Idl {\} = carrier R" proof - have "\ \ Idl {\}" by (simp add: genideal_self') thus "Idl {\} = carrier R" by (intro ideal.one_imp_carrier, fast intro: genideal_ideal) qed text {* The zero ideal is a principal ideal *} corollary (in ring) zeropideal: shows "principalideal {\} R" apply (rule principalidealI) apply (rule zeroideal) apply (blast intro!: zero_closed zero_genideal[symmetric]) done text {* The unit ideal is a principal ideal *} corollary (in ring) onepideal: shows "principalideal (carrier R) R" apply (rule principalidealI) apply (rule oneideal) apply (blast intro!: one_closed one_genideal[symmetric]) done text {* Every principal ideal is a right coset of the carrier *} lemma (in principalideal) rcos_generate: assumes "cring R" shows "\x\I. I = carrier R #> x" proof - interpret cring R by fact from generate obtain i where icarr: "i \ carrier R" and I1: "I = Idl {i}" by fast+ from icarr and genideal_self[of "{i}"] have "i \ Idl {i}" by fast hence iI: "i \ I" by (simp add: I1) from I1 icarr have I2: "I = PIdl i" by (simp add: cgenideal_eq_genideal) have "PIdl i = carrier R #> i" unfolding cgenideal_def r_coset_def by fast from I2 and this have "I = carrier R #> i" by simp from iI and this show "\x\I. I = carrier R #> x" by fast qed subsection {* Prime Ideals *} lemma (in ideal) primeidealCD: assumes "cring R" assumes notprime: "\ primeideal I R" shows "carrier R = I \ (\a b. a \ carrier R \ b \ carrier R \ a \ b \ I \ a \ I \ b \ I)" proof (rule ccontr, clarsimp) interpret cring R by fact assume InR: "carrier R \ I" and "\a. a \ carrier R \ (\b. a \ b \ I \ b \ carrier R \ a \ I \ b \ I)" hence I_prime: "\ a b. \a \ carrier R; b \ carrier R; a \ b \ I\ \ a \ I \ b \ I" by simp have "primeideal I R" apply (rule primeideal.intro [OF is_ideal is_cring]) apply (rule primeideal_axioms.intro) apply (rule InR) apply (erule (2) I_prime) done from this and notprime show "False" by simp qed lemma (in ideal) primeidealCE: assumes "cring R" assumes notprime: "\ primeideal I R" obtains "carrier R = I" | "\a b. a \ carrier R \ b \ carrier R \ a \ b \ I \ a \ I \ b \ I" proof - - interpret R!: cring R by fact + interpret R: cring R by fact assume "carrier R = I ==> thesis" and "\a b. a \ carrier R \ b \ carrier R \ a \ b \ I \ a \ I \ b \ I \ thesis" then show thesis using primeidealCD [OF R.is_cring notprime] by blast qed text {* If @{text "{\}"} is a prime ideal of a commutative ring, the ring is a domain *} lemma (in cring) zeroprimeideal_domainI: assumes pi: "primeideal {\} R" shows "domain R" apply (rule domain.intro, rule is_cring) apply (rule domain_axioms.intro) proof (rule ccontr, simp) interpret primeideal "{\}" "R" by (rule pi) assume "\ = \" hence "carrier R = {\}" by (rule one_zeroD) from this[symmetric] and I_notcarr show "False" by simp next interpret primeideal "{\}" "R" by (rule pi) fix a b assume ab: "a \ b = \" and carr: "a \ carrier R" "b \ carrier R" from ab have abI: "a \ b \ {\}" by fast from carr and this have "a \ {\} \ b \ {\}" by (rule I_prime) thus "a = \ \ b = \" by simp qed corollary (in cring) domain_eq_zeroprimeideal: shows "domain R = primeideal {\} R" apply rule apply (erule domain.zeroprimeideal) apply (erule zeroprimeideal_domainI) done subsection {* Maximal Ideals *} lemma (in ideal) helper_I_closed: assumes carr: "a \ carrier R" "x \ carrier R" "y \ carrier R" and axI: "a \ x \ I" shows "a \ (x \ y) \ I" proof - from axI and carr have "(a \ x) \ y \ I" by (simp add: I_r_closed) also from carr have "(a \ x) \ y = a \ (x \ y)" by (simp add: m_assoc) finally show "a \ (x \ y) \ I" . qed lemma (in ideal) helper_max_prime: assumes "cring R" assumes acarr: "a \ carrier R" shows "ideal {x\carrier R. a \ x \ I} R" proof - interpret cring R by fact show ?thesis apply (rule idealI) apply (rule cring.axioms[OF is_cring]) apply (rule subgroup.intro) apply (simp, fast) apply clarsimp apply (simp add: r_distr acarr) apply (simp add: acarr) apply (simp add: a_inv_def[symmetric], clarify) defer 1 apply clarsimp defer 1 apply (fast intro!: helper_I_closed acarr) proof - fix x assume xcarr: "x \ carrier R" and ax: "a \ x \ I" from ax and acarr xcarr have "\(a \ x) \ I" by simp also from acarr xcarr have "\(a \ x) = a \ (\x)" by algebra finally show "a \ (\x) \ I" . from acarr have "a \ \ = \" by simp next fix x y assume xcarr: "x \ carrier R" and ycarr: "y \ carrier R" and ayI: "a \ y \ I" from ayI and acarr xcarr ycarr have "a \ (y \ x) \ I" by (simp add: helper_I_closed) moreover from xcarr ycarr have "y \ x = x \ y" by (simp add: m_comm) ultimately show "a \ (x \ y) \ I" by simp qed qed text {* In a cring every maximal ideal is prime *} lemma (in cring) maximalideal_is_prime: assumes "maximalideal I R" shows "primeideal I R" proof - interpret maximalideal I R by fact show ?thesis apply (rule ccontr) apply (rule primeidealCE) apply (rule is_cring) apply assumption apply (simp add: I_notcarr) proof - assume "\a b. a \ carrier R \ b \ carrier R \ a \ b \ I \ a \ I \ b \ I" from this obtain a b where acarr: "a \ carrier R" and bcarr: "b \ carrier R" and abI: "a \ b \ I" and anI: "a \ I" and bnI: "b \ I" by fast def J \ "{x\carrier R. a \ x \ I}" from is_cring and acarr have idealJ: "ideal J R" unfolding J_def by (rule helper_max_prime) have IsubJ: "I \ J" proof fix x assume xI: "x \ I" from this and acarr have "a \ x \ I" by (intro I_l_closed) from xI[THEN a_Hcarr] this show "x \ J" unfolding J_def by fast qed from abI and acarr bcarr have "b \ J" unfolding J_def by fast from bnI and this have JnI: "J \ I" by fast from acarr have "a = a \ \" by algebra from this and anI have "a \ \ \ I" by simp from one_closed and this have "\ \ J" unfolding J_def by fast hence Jncarr: "J \ carrier R" by fast interpret ideal J R by (rule idealJ) have "J = I \ J = carrier R" apply (intro I_maximal) apply (rule idealJ) apply (rule IsubJ) apply (rule a_subset) done from this and JnI and Jncarr show "False" by simp qed qed subsection {* Derived Theorems *} --"A non-zero cring that has only the two trivial ideals is a field" lemma (in cring) trivialideals_fieldI: assumes carrnzero: "carrier R \ {\}" and haveideals: "{I. ideal I R} = {{\}, carrier R}" shows "field R" apply (rule cring_fieldI) apply (rule, rule, rule) apply (erule Units_closed) defer 1 apply rule defer 1 proof (rule ccontr, simp) assume zUnit: "\ \ Units R" hence a: "\ \ inv \ = \" by (rule Units_r_inv) from zUnit have "\ \ inv \ = \" by (intro l_null, rule Units_inv_closed) from a[symmetric] and this have "\ = \" by simp hence "carrier R = {\}" by (rule one_zeroD) from this and carrnzero show "False" by simp next fix x assume xcarr': "x \ carrier R - {\}" hence xcarr: "x \ carrier R" by fast from xcarr' have xnZ: "x \ \" by fast from xcarr have xIdl: "ideal (PIdl x) R" by (intro cgenideal_ideal, fast) from xcarr have "x \ PIdl x" by (intro cgenideal_self, fast) from this and xnZ have "PIdl x \ {\}" by fast from haveideals and this have "PIdl x = carrier R" by (blast intro!: xIdl) hence "\ \ PIdl x" by simp hence "\y. \ = y \ x \ y \ carrier R" unfolding cgenideal_def by blast from this obtain y where ycarr: " y \ carrier R" and ylinv: "\ = y \ x" by fast+ from ylinv and xcarr ycarr have yrinv: "\ = x \ y" by (simp add: m_comm) from ycarr and ylinv[symmetric] and yrinv[symmetric] have "\y \ carrier R. y \ x = \ \ x \ y = \" by fast from this and xcarr show "x \ Units R" unfolding Units_def by fast qed lemma (in field) all_ideals: shows "{I. ideal I R} = {{\}, carrier R}" apply (rule, rule) proof - fix I assume a: "I \ {I. ideal I R}" with this interpret ideal I R by simp show "I \ {{\}, carrier R}" proof (cases "\a. a \ I - {\}") assume "\a. a \ I - {\}" from this obtain a where aI: "a \ I" and anZ: "a \ \" by fast+ from aI[THEN a_Hcarr] anZ have aUnit: "a \ Units R" by (simp add: field_Units) hence a: "a \ inv a = \" by (rule Units_r_inv) from aI and aUnit have "a \ inv a \ I" by (simp add: I_r_closed del: Units_r_inv) hence oneI: "\ \ I" by (simp add: a[symmetric]) have "carrier R \ I" proof fix x assume xcarr: "x \ carrier R" from oneI and this have "\ \ x \ I" by (rule I_r_closed) from this and xcarr show "x \ I" by simp qed from this and a_subset have "I = carrier R" by fast thus "I \ {{\}, carrier R}" by fast next assume "\ (\a. a \ I - {\})" hence IZ: "\a. a \ I \ a = \" by simp have a: "I \ {\}" proof fix x assume "x \ I" hence "x = \" by (rule IZ) thus "x \ {\}" by fast qed have "\ \ I" by simp hence "{\} \ I" by fast from this and a have "I = {\}" by fast thus "I \ {{\}, carrier R}" by fast qed qed (simp add: zeroideal oneideal) --"Jacobson Theorem 2.2" lemma (in cring) trivialideals_eq_field: assumes carrnzero: "carrier R \ {\}" shows "({I. ideal I R} = {{\}, carrier R}) = field R" by (fast intro!: trivialideals_fieldI[OF carrnzero] field.all_ideals) text {* Like zeroprimeideal for domains *} lemma (in field) zeromaximalideal: "maximalideal {\} R" apply (rule maximalidealI) apply (rule zeroideal) proof- from one_not_zero have "\ \ {\}" by simp from this and one_closed show "carrier R \ {\}" by fast next fix J assume Jideal: "ideal J R" hence "J \ {I. ideal I R}" by fast from this and all_ideals show "J = {\} \ J = carrier R" by simp qed lemma (in cring) zeromaximalideal_fieldI: assumes zeromax: "maximalideal {\} R" shows "field R" apply (rule trivialideals_fieldI, rule maximalideal.I_notcarr[OF zeromax]) apply rule apply clarsimp defer 1 apply (simp add: zeroideal oneideal) proof - fix J assume Jn0: "J \ {\}" and idealJ: "ideal J R" interpret ideal J R by (rule idealJ) have "{\} \ J" by (rule ccontr, simp) from zeromax and idealJ and this and a_subset have "J = {\} \ J = carrier R" by (rule maximalideal.I_maximal) from this and Jn0 show "J = carrier R" by simp qed lemma (in cring) zeromaximalideal_eq_field: "maximalideal {\} R = field R" apply rule apply (erule zeromaximalideal_fieldI) apply (erule field.zeromaximalideal) done end diff --git a/src/HOL/Algebra/IntRing.thy b/src/HOL/Algebra/IntRing.thy --- a/src/HOL/Algebra/IntRing.thy +++ b/src/HOL/Algebra/IntRing.thy @@ -1,492 +1,492 @@ (* Title: HOL/Algebra/IntRing.thy Author: Stephan Hohe, TU Muenchen *) theory IntRing imports QuotRing Lattice Int Primes begin section {* The Ring of Integers *} subsection {* Some properties of @{typ int} *} lemma dvds_imp_abseq: "\l dvd k; k dvd l\ \ abs l = abs (k::int)" apply (subst abs_split, rule conjI) apply (clarsimp, subst abs_split, rule conjI) apply (clarsimp) apply (cases "k=0", simp) apply (cases "l=0", simp) apply (simp add: zdvd_anti_sym) apply clarsimp apply (cases "k=0", simp) apply (simp add: zdvd_anti_sym) apply (clarsimp, subst abs_split, rule conjI) apply (clarsimp) apply (cases "l=0", simp) apply (simp add: zdvd_anti_sym) apply (clarsimp) apply (subgoal_tac "-l = -k", simp) apply (intro zdvd_anti_sym, simp+) done lemma abseq_imp_dvd: assumes a_lk: "abs l = abs (k::int)" shows "l dvd k" proof - from a_lk have "nat (abs l) = nat (abs k)" by simp hence "nat (abs l) dvd nat (abs k)" by simp hence "int (nat (abs l)) dvd k" by (subst int_dvd_iff) hence "abs l dvd k" by simp thus "l dvd k" apply (unfold dvd_def, cases "l<0") defer 1 apply clarsimp proof (clarsimp) fix k assume l0: "l < 0" have "- (l * k) = l * (-k)" by simp thus "\ka. - (l * k) = l * ka" by fast qed qed lemma dvds_eq_abseq: "(l dvd k \ k dvd l) = (abs l = abs (k::int))" apply rule apply (simp add: dvds_imp_abseq) apply (rule conjI) apply (simp add: abseq_imp_dvd)+ done subsection {* @{text "\"}: The Set of Integers as Algebraic Structure *} constdefs int_ring :: "int ring" ("\") "int_ring \ \carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\" lemma int_Zcarr [intro!, simp]: "k \ carrier \" by (simp add: int_ring_def) lemma int_is_cring: "cring \" unfolding int_ring_def apply (rule cringI) apply (rule abelian_groupI, simp_all) defer 1 apply (rule comm_monoidI, simp_all) apply (rule zadd_zmult_distrib) apply (fast intro: zadd_zminus_inverse2) done (* lemma int_is_domain: "domain \" apply (intro domain.intro domain_axioms.intro) apply (rule int_is_cring) apply (unfold int_ring_def, simp+) done *) subsection {* Interpretations *} text {* Since definitions of derived operations are global, their interpretation needs to be done as early as possible --- that is, with as few assumptions as possible. *} -interpretation int!: monoid \ +interpretation int: monoid \ where "carrier \ = UNIV" and "mult \ x y = x * y" and "one \ = 1" and "pow \ x n = x^n" proof - -- "Specification" show "monoid \" proof qed (auto simp: int_ring_def) - then interpret int!: monoid \ . + then interpret int: monoid \ . -- "Carrier" show "carrier \ = UNIV" by (simp add: int_ring_def) -- "Operations" { fix x y show "mult \ x y = x * y" by (simp add: int_ring_def) } note mult = this show one: "one \ = 1" by (simp add: int_ring_def) show "pow \ x n = x^n" by (induct n) (simp, simp add: int_ring_def)+ qed -interpretation int!: comm_monoid \ +interpretation int: comm_monoid \ where "finprod \ f A = (if finite A then setprod f A else undefined)" proof - -- "Specification" show "comm_monoid \" proof qed (auto simp: int_ring_def) - then interpret int!: comm_monoid \ . + then interpret int: comm_monoid \ . -- "Operations" { fix x y have "mult \ x y = x * y" by (simp add: int_ring_def) } note mult = this have one: "one \ = 1" by (simp add: int_ring_def) show "finprod \ f A = (if finite A then setprod f A else undefined)" proof (cases "finite A") case True then show ?thesis proof induct case empty show ?case by (simp add: one) next case insert then show ?case by (simp add: Pi_def mult) qed next case False then show ?thesis by (simp add: finprod_def) qed qed -interpretation int!: abelian_monoid \ +interpretation int: abelian_monoid \ where "zero \ = 0" and "add \ x y = x + y" and "finsum \ f A = (if finite A then setsum f A else undefined)" proof - -- "Specification" show "abelian_monoid \" proof qed (auto simp: int_ring_def) - then interpret int!: abelian_monoid \ . + then interpret int: abelian_monoid \ . -- "Operations" { fix x y show "add \ x y = x + y" by (simp add: int_ring_def) } note add = this show zero: "zero \ = 0" by (simp add: int_ring_def) show "finsum \ f A = (if finite A then setsum f A else undefined)" proof (cases "finite A") case True then show ?thesis proof induct case empty show ?case by (simp add: zero) next case insert then show ?case by (simp add: Pi_def add) qed next case False then show ?thesis by (simp add: finsum_def finprod_def) qed qed -interpretation int!: abelian_group \ +interpretation int: abelian_group \ where "a_inv \ x = - x" and "a_minus \ x y = x - y" proof - -- "Specification" show "abelian_group \" proof (rule abelian_groupI) show "!!x. x \ carrier \ ==> EX y : carrier \. y \\<^bsub>\\<^esub> x = \\<^bsub>\\<^esub>" by (simp add: int_ring_def) arith qed (auto simp: int_ring_def) - then interpret int!: abelian_group \ . + then interpret int: abelian_group \ . -- "Operations" { fix x y have "add \ x y = x + y" by (simp add: int_ring_def) } note add = this have zero: "zero \ = 0" by (simp add: int_ring_def) { fix x have "add \ (-x) x = zero \" by (simp add: add zero) then show "a_inv \ x = - x" by (simp add: int.minus_equality) } note a_inv = this show "a_minus \ x y = x - y" by (simp add: int.minus_eq add a_inv) qed -interpretation int!: "domain" \ +interpretation int: "domain" \ proof qed (auto simp: int_ring_def left_distrib right_distrib) text {* Removal of occurrences of @{term UNIV} in interpretation result --- experimental. *} lemma UNIV: "x \ UNIV = True" "A \ UNIV = True" "(ALL x : UNIV. P x) = (ALL x. P x)" "(EX x : UNIV. P x) = (EX x. P x)" "(True --> Q) = Q" "(True ==> PROP R) == PROP R" by simp_all -interpretation int! (* FIXME [unfolded UNIV] *) : +interpretation int (* FIXME [unfolded UNIV] *) : partial_order "(| carrier = UNIV::int set, eq = op =, le = op \ |)" where "carrier (| carrier = UNIV::int set, eq = op =, le = op \ |) = UNIV" and "le (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = (x \ y)" and "lless (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = (x < y)" proof - show "partial_order (| carrier = UNIV::int set, eq = op =, le = op \ |)" proof qed simp_all show "carrier (| carrier = UNIV::int set, eq = op =, le = op \ |) = UNIV" by simp show "le (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = (x \ y)" by simp show "lless (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = (x < y)" by (simp add: lless_def) auto qed -interpretation int! (* FIXME [unfolded UNIV] *) : +interpretation int (* FIXME [unfolded UNIV] *) : lattice "(| carrier = UNIV::int set, eq = op =, le = op \ |)" where "join (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = max x y" and "meet (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = min x y" proof - let ?Z = "(| carrier = UNIV::int set, eq = op =, le = op \ |)" show "lattice ?Z" apply unfold_locales apply (simp add: least_def Upper_def) apply arith apply (simp add: greatest_def Lower_def) apply arith done - then interpret int!: lattice "?Z" . + then interpret int: lattice "?Z" . show "join ?Z x y = max x y" apply (rule int.joinI) apply (simp_all add: least_def Upper_def) apply arith done show "meet ?Z x y = min x y" apply (rule int.meetI) apply (simp_all add: greatest_def Lower_def) apply arith done qed -interpretation int! (* [unfolded UNIV] *) : +interpretation int (* [unfolded UNIV] *) : total_order "(| carrier = UNIV::int set, eq = op =, le = op \ |)" proof qed clarsimp subsection {* Generated Ideals of @{text "\"} *} lemma int_Idl: "Idl\<^bsub>\\<^esub> {a} = {x * a | x. True}" apply (subst int.cgenideal_eq_genideal[symmetric]) apply (simp add: int_ring_def) apply (simp add: cgenideal_def int_ring_def) done lemma multiples_principalideal: "principalideal {x * a | x. True } \" apply (subst int_Idl[symmetric], rule principalidealI) apply (rule int.genideal_ideal, simp) apply fast done lemma prime_primeideal: assumes prime: "prime (nat p)" shows "primeideal (Idl\<^bsub>\\<^esub> {p}) \" apply (rule primeidealI) apply (rule int.genideal_ideal, simp) apply (rule int_is_cring) apply (simp add: int.cgenideal_eq_genideal[symmetric] cgenideal_def) apply (simp add: int_ring_def) apply clarsimp defer 1 apply (simp add: int.cgenideal_eq_genideal[symmetric] cgenideal_def) apply (simp add: int_ring_def) apply (elim exE) proof - fix a b x from prime have ppos: "0 <= p" by (simp add: prime_def) have unnat: "!!x. nat p dvd nat (abs x) ==> p dvd x" proof - fix x assume "nat p dvd nat (abs x)" hence "int (nat p) dvd x" by (simp add: int_dvd_iff[symmetric]) thus "p dvd x" by (simp add: ppos) qed assume "a * b = x * p" hence "p dvd a * b" by simp hence "nat p dvd nat (abs (a * b))" using ppos by (simp add: nat_dvd_iff) hence "nat p dvd (nat (abs a) * nat (abs b))" by (simp add: nat_abs_mult_distrib) hence "nat p dvd nat (abs a) | nat p dvd nat (abs b)" by (rule prime_dvd_mult[OF prime]) hence "p dvd a | p dvd b" by (fast intro: unnat) thus "(EX x. a = x * p) | (EX x. b = x * p)" proof assume "p dvd a" hence "EX x. a = p * x" by (simp add: dvd_def) from this obtain x where "a = p * x" by fast hence "a = x * p" by simp hence "EX x. a = x * p" by simp thus "(EX x. a = x * p) | (EX x. b = x * p)" .. next assume "p dvd b" hence "EX x. b = p * x" by (simp add: dvd_def) from this obtain x where "b = p * x" by fast hence "b = x * p" by simp hence "EX x. b = x * p" by simp thus "(EX x. a = x * p) | (EX x. b = x * p)" .. qed next assume "UNIV = {uu. EX x. uu = x * p}" from this obtain x where "1 = x * p" by best from this [symmetric] have "p * x = 1" by (subst zmult_commute) hence "\p * x\ = 1" by simp hence "\p\ = 1" by (rule abs_zmult_eq_1) from this and prime show "False" by (simp add: prime_def) qed subsection {* Ideals and Divisibility *} lemma int_Idl_subset_ideal: "Idl\<^bsub>\\<^esub> {k} \ Idl\<^bsub>\\<^esub> {l} = (k \ Idl\<^bsub>\\<^esub> {l})" by (rule int.Idl_subset_ideal', simp+) lemma Idl_subset_eq_dvd: "(Idl\<^bsub>\\<^esub> {k} \ Idl\<^bsub>\\<^esub> {l}) = (l dvd k)" apply (subst int_Idl_subset_ideal, subst int_Idl, simp) apply (rule, clarify) apply (simp add: dvd_def) apply (simp add: dvd_def mult_ac) done lemma dvds_eq_Idl: "(l dvd k \ k dvd l) = (Idl\<^bsub>\\<^esub> {k} = Idl\<^bsub>\\<^esub> {l})" proof - have a: "l dvd k = (Idl\<^bsub>\\<^esub> {k} \ Idl\<^bsub>\\<^esub> {l})" by (rule Idl_subset_eq_dvd[symmetric]) have b: "k dvd l = (Idl\<^bsub>\\<^esub> {l} \ Idl\<^bsub>\\<^esub> {k})" by (rule Idl_subset_eq_dvd[symmetric]) have "(l dvd k \ k dvd l) = ((Idl\<^bsub>\\<^esub> {k} \ Idl\<^bsub>\\<^esub> {l}) \ (Idl\<^bsub>\\<^esub> {l} \ Idl\<^bsub>\\<^esub> {k}))" by (subst a, subst b, simp) also have "((Idl\<^bsub>\\<^esub> {k} \ Idl\<^bsub>\\<^esub> {l}) \ (Idl\<^bsub>\\<^esub> {l} \ Idl\<^bsub>\\<^esub> {k})) = (Idl\<^bsub>\\<^esub> {k} = Idl\<^bsub>\\<^esub> {l})" by (rule, fast+) finally show ?thesis . qed lemma Idl_eq_abs: "(Idl\<^bsub>\\<^esub> {k} = Idl\<^bsub>\\<^esub> {l}) = (abs l = abs k)" apply (subst dvds_eq_abseq[symmetric]) apply (rule dvds_eq_Idl[symmetric]) done subsection {* Ideals and the Modulus *} constdefs ZMod :: "int => int => int set" "ZMod k r == (Idl\<^bsub>\\<^esub> {k}) +>\<^bsub>\\<^esub> r" lemmas ZMod_defs = ZMod_def genideal_def lemma rcos_zfact: assumes kIl: "k \ ZMod l r" shows "EX x. k = x * l + r" proof - from kIl[unfolded ZMod_def] have "\xl\Idl\<^bsub>\\<^esub> {l}. k = xl + r" by (simp add: a_r_coset_defs int_ring_def) from this obtain xl where xl: "xl \ Idl\<^bsub>\\<^esub> {l}" and k: "k = xl + r" by auto from xl obtain x where "xl = x * l" by (simp add: int_Idl, fast) from k and this have "k = x * l + r" by simp thus "\x. k = x * l + r" .. qed lemma ZMod_imp_zmod: assumes zmods: "ZMod m a = ZMod m b" shows "a mod m = b mod m" proof - interpret ideal "Idl\<^bsub>\\<^esub> {m}" \ by (rule int.genideal_ideal, fast) from zmods have "b \ ZMod m a" unfolding ZMod_def by (simp add: a_repr_independenceD) from this have "EX x. b = x * m + a" by (rule rcos_zfact) from this obtain x where "b = x * m + a" by fast hence "b mod m = (x * m + a) mod m" by simp also have "\ = ((x * m) mod m) + (a mod m)" by (simp add: mod_add_eq) also have "\ = a mod m" by simp finally have "b mod m = a mod m" . thus "a mod m = b mod m" .. qed lemma ZMod_mod: shows "ZMod m a = ZMod m (a mod m)" proof - interpret ideal "Idl\<^bsub>\\<^esub> {m}" \ by (rule int.genideal_ideal, fast) show ?thesis unfolding ZMod_def apply (rule a_repr_independence'[symmetric]) apply (simp add: int_Idl a_r_coset_defs) apply (simp add: int_ring_def) proof - have "a = m * (a div m) + (a mod m)" by (simp add: zmod_zdiv_equality) hence "a = (a div m) * m + (a mod m)" by simp thus "\h. (\x. h = x * m) \ a = h + a mod m" by fast qed simp qed lemma zmod_imp_ZMod: assumes modeq: "a mod m = b mod m" shows "ZMod m a = ZMod m b" proof - have "ZMod m a = ZMod m (a mod m)" by (rule ZMod_mod) also have "\ = ZMod m (b mod m)" by (simp add: modeq[symmetric]) also have "\ = ZMod m b" by (rule ZMod_mod[symmetric]) finally show ?thesis . qed corollary ZMod_eq_mod: shows "(ZMod m a = ZMod m b) = (a mod m = b mod m)" by (rule, erule ZMod_imp_zmod, erule zmod_imp_ZMod) subsection {* Factorization *} constdefs ZFact :: "int \ int set ring" "ZFact k == \ Quot (Idl\<^bsub>\\<^esub> {k})" lemmas ZFact_defs = ZFact_def FactRing_def lemma ZFact_is_cring: shows "cring (ZFact k)" apply (unfold ZFact_def) apply (rule ideal.quotient_is_cring) apply (intro ring.genideal_ideal) apply (simp add: cring.axioms[OF int_is_cring] ring.intro) apply simp apply (rule int_is_cring) done lemma ZFact_zero: "carrier (ZFact 0) = (\a. {{a}})" apply (insert int.genideal_zero) apply (simp add: ZFact_defs A_RCOSETS_defs r_coset_def int_ring_def ring_record_simps) done lemma ZFact_one: "carrier (ZFact 1) = {UNIV}" apply (simp only: ZFact_defs A_RCOSETS_defs r_coset_def int_ring_def ring_record_simps) apply (subst int.genideal_one[unfolded int_ring_def, simplified ring_record_simps]) apply (rule, rule, clarsimp) apply (rule, rule, clarsimp) apply (rule, clarsimp, arith) apply (rule, clarsimp) apply (rule exI[of _ "0"], clarsimp) done lemma ZFact_prime_is_domain: assumes pprime: "prime (nat p)" shows "domain (ZFact p)" apply (unfold ZFact_def) apply (rule primeideal.quotient_is_domain) apply (rule prime_primeideal[OF pprime]) done end diff --git a/src/HOL/Algebra/RingHom.thy b/src/HOL/Algebra/RingHom.thy --- a/src/HOL/Algebra/RingHom.thy +++ b/src/HOL/Algebra/RingHom.thy @@ -1,206 +1,206 @@ (* Title: HOL/Algebra/RingHom.thy Author: Stephan Hohe, TU Muenchen *) theory RingHom imports Ideal begin section {* Homomorphisms of Non-Commutative Rings *} text {* Lifting existing lemmas in a @{text ring_hom_ring} locale *} locale ring_hom_ring = R: ring R + S: ring S for R (structure) and S (structure) + fixes h assumes homh: "h \ ring_hom R S" notes hom_mult [simp] = ring_hom_mult [OF homh] and hom_one [simp] = ring_hom_one [OF homh] sublocale ring_hom_cring \ ring: ring_hom_ring proof qed (rule homh) sublocale ring_hom_ring \ abelian_group: abelian_group_hom R S apply (rule abelian_group_homI) apply (rule R.is_abelian_group) apply (rule S.is_abelian_group) apply (intro group_hom.intro group_hom_axioms.intro) apply (rule R.a_group) apply (rule S.a_group) apply (insert homh, unfold hom_def ring_hom_def) apply simp done lemma (in ring_hom_ring) is_ring_hom_ring: "ring_hom_ring R S h" by (rule ring_hom_ring_axioms) lemma ring_hom_ringI: fixes R (structure) and S (structure) assumes "ring R" "ring S" assumes (* morphism: "h \ carrier R \ carrier S" *) hom_closed: "!!x. x \ carrier R ==> h x \ carrier S" and compatible_mult: "!!x y. [| x : carrier R; y : carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_add: "!!x y. [| x : carrier R; y : carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_one: "h \ = \\<^bsub>S\<^esub>" shows "ring_hom_ring R S h" proof - interpret ring R by fact interpret ring S by fact show ?thesis apply unfold_locales apply (unfold ring_hom_def, safe) apply (simp add: hom_closed Pi_def) apply (erule (1) compatible_mult) apply (erule (1) compatible_add) apply (rule compatible_one) done qed lemma ring_hom_ringI2: assumes "ring R" "ring S" assumes h: "h \ ring_hom R S" shows "ring_hom_ring R S h" proof - - interpret R!: ring R by fact - interpret S!: ring S by fact + interpret R: ring R by fact + interpret S: ring S by fact show ?thesis apply (intro ring_hom_ring.intro ring_hom_ring_axioms.intro) apply (rule R.is_ring) apply (rule S.is_ring) apply (rule h) done qed lemma ring_hom_ringI3: fixes R (structure) and S (structure) assumes "abelian_group_hom R S h" "ring R" "ring S" assumes compatible_mult: "!!x y. [| x : carrier R; y : carrier R |] ==> h (x \ y) = h x \\<^bsub>S\<^esub> h y" and compatible_one: "h \ = \\<^bsub>S\<^esub>" shows "ring_hom_ring R S h" proof - interpret abelian_group_hom R S h by fact - interpret R!: ring R by fact - interpret S!: ring S by fact + interpret R: ring R by fact + interpret S: ring S by fact show ?thesis apply (intro ring_hom_ring.intro ring_hom_ring_axioms.intro, rule R.is_ring, rule S.is_ring) apply (insert group_hom.homh[OF a_group_hom]) apply (unfold hom_def ring_hom_def, simp) apply safe apply (erule (1) compatible_mult) apply (rule compatible_one) done qed lemma ring_hom_cringI: assumes "ring_hom_ring R S h" "cring R" "cring S" shows "ring_hom_cring R S h" proof - interpret ring_hom_ring R S h by fact - interpret R!: cring R by fact - interpret S!: cring S by fact + interpret R: cring R by fact + interpret S: cring S by fact show ?thesis by (intro ring_hom_cring.intro ring_hom_cring_axioms.intro) (rule R.is_cring, rule S.is_cring, rule homh) qed subsection {* The Kernel of a Ring Homomorphism *} --"the kernel of a ring homomorphism is an ideal" lemma (in ring_hom_ring) kernel_is_ideal: shows "ideal (a_kernel R S h) R" apply (rule idealI) apply (rule R.is_ring) apply (rule additive_subgroup.a_subgroup[OF additive_subgroup_a_kernel]) apply (unfold a_kernel_def', simp+) done text {* Elements of the kernel are mapped to zero *} lemma (in abelian_group_hom) kernel_zero [simp]: "i \ a_kernel R S h \ h i = \\<^bsub>S\<^esub>" by (simp add: a_kernel_defs) subsection {* Cosets *} text {* Cosets of the kernel correspond to the elements of the image of the homomorphism *} lemma (in ring_hom_ring) rcos_imp_homeq: assumes acarr: "a \ carrier R" and xrcos: "x \ a_kernel R S h +> a" shows "h x = h a" proof - interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal) from xrcos have "\i \ a_kernel R S h. x = i \ a" by (simp add: a_r_coset_defs) from this obtain i where iker: "i \ a_kernel R S h" and x: "x = i \ a" by fast+ note carr = acarr iker[THEN a_Hcarr] from x have "h x = h (i \ a)" by simp also from carr have "\ = h i \\<^bsub>S\<^esub> h a" by simp also from iker have "\ = \\<^bsub>S\<^esub> \\<^bsub>S\<^esub> h a" by simp also from carr have "\ = h a" by simp finally show "h x = h a" . qed lemma (in ring_hom_ring) homeq_imp_rcos: assumes acarr: "a \ carrier R" and xcarr: "x \ carrier R" and hx: "h x = h a" shows "x \ a_kernel R S h +> a" proof - interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal) note carr = acarr xcarr note hcarr = acarr[THEN hom_closed] xcarr[THEN hom_closed] from hx and hcarr have a: "h x \\<^bsub>S\<^esub> \\<^bsub>S\<^esub>h a = \\<^bsub>S\<^esub>" by algebra from carr have "h x \\<^bsub>S\<^esub> \\<^bsub>S\<^esub>h a = h (x \ \a)" by simp from a and this have b: "h (x \ \a) = \\<^bsub>S\<^esub>" by simp from carr have "x \ \a \ carrier R" by simp from this and b have "x \ \a \ a_kernel R S h" unfolding a_kernel_def' by fast from this and carr show "x \ a_kernel R S h +> a" by (simp add: a_rcos_module_rev) qed corollary (in ring_hom_ring) rcos_eq_homeq: assumes acarr: "a \ carrier R" shows "(a_kernel R S h) +> a = {x \ carrier R. h x = h a}" apply rule defer 1 apply clarsimp defer 1 proof interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal) fix x assume xrcos: "x \ a_kernel R S h +> a" from acarr and this have xcarr: "x \ carrier R" by (rule a_elemrcos_carrier) from xrcos have "h x = h a" by (rule rcos_imp_homeq[OF acarr]) from xcarr and this show "x \ {x \ carrier R. h x = h a}" by fast next interpret ideal "a_kernel R S h" "R" by (rule kernel_is_ideal) fix x assume xcarr: "x \ carrier R" and hx: "h x = h a" from acarr xcarr hx show "x \ a_kernel R S h +> a" by (rule homeq_imp_rcos) qed 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,1907 +1,1907 @@ (* 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 @{text "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}, 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_def: "up R == {f. f \ UNIV -> carrier R & (EX n. bound \\<^bsub>R\<^esub> n f)}" definition UP :: "('a, 'm) ring_scheme => ('a, nat => 'a) up_ring" where UP_def: "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; EX 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 ==> EX 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 "EX 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 fastsimp 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 auto then show "EX 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" using mem_upD [of p R] mem_upD [of q R] up_add_closed up_a_inv_closed a_minus_def [of _ R] 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 "EX 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 "EX 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 cring) cring: "cring R" .. lemma (in UP_cring) UP_algebra: "algebra R P" by (auto intro!: algebraI R.cring 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: R.r_null 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: ivl_disj_int_singleton 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 ivl_disj_int_singleton 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) simp_all 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 "EX m. deg R p - 1 < m & coeff P p m ~= \" by (unfold bound_def) fast then have "EX 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 fastsimp 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 "EX 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_anti_sym deg_aboveI deg_belowI) text {* Degree and polynomial operations *} lemma deg_add [simp]: assumes R: "p \ carrier P" "q \ carrier P" shows "deg R (p \\<^bsub>P\<^esub> q) <= max (deg R p) (deg R q)" proof (cases "deg R p <= deg R q") case True show ?thesis by (rule deg_aboveI) (simp_all add: True R deg_aboveD) next case False show ?thesis by (rule deg_aboveI) (simp_all add: False R deg_aboveD) qed 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 (fastsimp intro: le_anti_sym deg_aboveI deg_belowI) lemma deg_const [simp]: assumes R: "a \ carrier R" shows "deg R (monom P a 0) = 0" proof (rule le_anti_sym) 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_anti_sym) 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_anti_sym) 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_anti_sym) 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_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, @{text 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_anti_sym) 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_anti_sym) 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: ivl_disj_int_singleton 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: ivl_disj_int_singleton 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: ivl_disj_int_singleton 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 fastsimp 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 (* FIXME print_locale ring_hom_cring fails *) 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 @{text 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]: "[| finite A; f \ A -> carrier R |] ==> h (finsum R f A) = finsum S (h o f) A" proof (induct set: finite) case empty then show ?case by simp next case insert then show ?case by (simp add: Pi_def) qed 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 le_maxI1 le_maxI2) 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 @{text UP_cring} with the additional assumption that @{text 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 apply (unfold Eval_def) apply intro_locales apply (rule ring_hom_cring.axioms) apply (rule ring_hom_cring.intro) apply unfold_locales apply (rule eval_ring_hom) apply rule done 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" proof (rule ring_hom_cring.intro) show "ring_hom_cring_axioms P S Phi" by (rule ring_hom_cring_axioms.intro) (rule Phi) qed unfold_locales theorem UP_universal_property: assumes S: "s \ carrier S" shows "EX! Phi. Phi \ ring_hom P S \ extensional (carrier P) & Phi (monom P \ 1) = s & (ALL 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" unfolding a_minus_def [OF p q] unfolding coeff_add [OF p a_inv_closed [OF q]] unfolding coeff_a_inv [OF q] using coeff_closed [OF p, of n] using coeff_closed [OF q, of n] by algebra 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)" proof - 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))" and ?lg = "lcoeff g" show ?thesis (*JE: we distinguish some particular cases where the solution is almost direct.*) proof (cases "deg R f < deg R g") case True (*JE: if the degree of f is smaller than the one of g the solution is straightforward.*) (* CB: avoid exI3 *) have "?pred \\<^bsub>P\<^esub> f 0" using True by force then show ?thesis by fast next case False then have deg_g_le_deg_f: "deg R g \ deg R f" by simp { (*JE: we now apply the induction hypothesis with some additional facts required*) from f_in_P deg_g_le_deg_f show ?thesis proof (induct n \ "deg R f" arbitrary: "f" rule: nat_less_induct) fix n f assume hypo: "\mx. x \ carrier P \ deg R g \ deg R x \ m = deg R x \ (\q r (k::nat). q \ carrier P \ r \ carrier P \ lcoeff g (^) k \\<^bsub>P\<^esub> x = g \\<^bsub>P\<^esub> q \\<^bsub>P\<^esub> r & (r = \\<^bsub>P\<^esub> | deg R r < deg R g))" and prem: "n = deg R f" and f_in_P [simp]: "f \ carrier P" and deg_g_le_deg_f: "deg R g \ deg R f" let ?k = "1::nat" and ?r = "(g \\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> (lcoeff g \\<^bsub>P\<^esub> f)" and ?q = "monom P (lcoeff f) (deg R f - deg R g)" show "\ 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)" proof - (*JE: we first extablish the existence of a triple satisfying the previous equation. Then we will have to prove the second part of the predicate.*) have exist: "lcoeff g (^) ?k \\<^bsub>P\<^esub> f = g \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> ?r" using minus_add using sym [OF a_assoc [of "g \\<^bsub>P\<^esub> ?q" "\\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q)" "lcoeff g \\<^bsub>P\<^esub> f"]] using r_neg by auto show ?thesis proof (cases "deg R (\\<^bsub>P\<^esub> ?r) < deg R g") (*JE: if the degree of the remainder satisfies the statement property we are done*) case True { show ?thesis proof (rule exI3 [of _ ?q "\\<^bsub>P\<^esub> ?r" ?k], intro conjI) show "lcoeff g (^) ?k \\<^bsub>P\<^esub> f = g \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> ?r" using exist by simp show "\\<^bsub>P\<^esub> ?r = \\<^bsub>P\<^esub> \ deg R (\\<^bsub>P\<^esub> ?r) < deg R g" using True by simp qed (simp_all) } next case False note n_deg_r_l_deg_g = False { (*JE: otherwise, we verify the conditions of the induction hypothesis.*) show ?thesis proof (cases "deg R f = 0") (*JE: the solutions are different if the degree of f is zero or not*) case True { have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp have "lcoeff g (^) (1::nat) \\<^bsub>P\<^esub> f = g \\<^bsub>P\<^esub> f \\<^bsub>P\<^esub> \\<^bsub>P\<^esub>" unfolding deg_g apply simp unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]] using deg_zero_impl_monom [OF g_in_P deg_g] by simp then show ?thesis using f_in_P by blast } next case False note deg_f_nzero = False { (*JE: now it only remains the case where the induction hypothesis can be used.*) (*JE: we first prove that the degree of the remainder is smaller than the one of f*) have deg_remainder_l_f: "deg R (\\<^bsub>P\<^esub> ?r) < n" proof - have "deg R (\\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp also have "\ < deg R f" proof (rule deg_lcoeff_cancel) show "deg R (\\<^bsub>P\<^esub> (lcoeff g \\<^bsub>P\<^esub> f)) \ deg R f" using deg_smult_ring [of "lcoeff g" f] using prem using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp show "deg R (g \\<^bsub>P\<^esub> ?q) \ deg R f" using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f by simp show "coeff P (g \\<^bsub>P\<^esub> ?q) (deg R f) = \ coeff P (\\<^bsub>P\<^esub> (lcoeff g \\<^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 lcoeff f else \))" "(\i. if deg R g = i then coeff P g i \ lcoeff f else \)"] using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\i. coeff P g i \ lcoeff f)"] unfolding Pi_def using deg_g_le_deg_f by force qed (simp_all add: deg_f_nzero) finally show "deg R (\\<^bsub>P\<^esub> ?r) < n" unfolding prem . qed moreover have "\\<^bsub>P\<^esub> ?r \ carrier P" by simp moreover obtain m where deg_rem_eq_m: "deg R (\\<^bsub>P\<^esub> ?r) = m" by auto moreover have "deg R g \ deg R (\\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*) ultimately obtain q' r' k' where rem_desc: "lcoeff g (^) (k'::nat) \\<^bsub>P\<^esub> (\\<^bsub>P\<^esub> ?r) = 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 hypo by blast (*JE: we now prove that the new quotient, remainder and exponent can be used to get the quotient, remainder and exponent of the long division theorem*) show ?thesis proof (rule exI3 [of _ "((lcoeff g (^) k') \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI) show "(lcoeff g (^) (Suc k')) \\<^bsub>P\<^esub> f = g \\<^bsub>P\<^esub> ((lcoeff g (^) k') \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> q') \\<^bsub>P\<^esub> r'" proof - have "(lcoeff g (^) (Suc k')) \\<^bsub>P\<^esub> f = (lcoeff g (^) k') \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q \\<^bsub>P\<^esub> \\<^bsub>P\<^esub> ?r)" using smult_assoc1 exist by simp also have "\ = (lcoeff g (^) k') \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q) \\<^bsub>P\<^esub> ((lcoeff g (^) k') \\<^bsub>P\<^esub> ( \\<^bsub>P\<^esub> ?r))" using UP_smult_r_distr by simp also have "\ = (lcoeff g (^) k') \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q) \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> q' \\<^bsub>P\<^esub> r')" using rem_desc by simp also have "\ = (lcoeff g (^) 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 "lcoeff g (^) k' \\<^bsub>P\<^esub> (g \\<^bsub>P\<^esub> ?q)" "g \\<^bsub>P\<^esub> q'" "r'"]] using q'_in_carrier r'_in_carrier by simp also have "\ = (lcoeff g (^) 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 "\ = (((lcoeff g (^) 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 by auto also have "\ = ((lcoeff g (^) 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 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 proof qed 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" proof qed (simp add: 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)" unfolding P.minus_eq [OF mon1_closed mon0_closed] unfolding hom_add [OF mon1_closed min_mon0_closed] unfolding hom_a_inv [OF mon0_closed] using R.minus_eq [symmetric] mon1_closed mon0_closed by auto 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 proof qed 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_def: "INTEG == (| carrier = UNIV, mult = op *, one = 1, zero = 0, add = op + |)" lemma INTEG_cring: "cring INTEG" by (unfold INTEG_def) (auto intro!: cringI abelian_groupI comm_monoidI zadd_zminus_inverse2 zadd_zmult_distrib) 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" +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/Complex.thy b/src/HOL/Complex.thy --- a/src/HOL/Complex.thy +++ b/src/HOL/Complex.thy @@ -1,721 +1,721 @@ (* Title: Complex.thy Author: Jacques D. Fleuriot Copyright: 2001 University of Edinburgh Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 *) header {* Complex Numbers: Rectangular and Polar Representations *} theory Complex imports Transcendental begin datatype complex = Complex real real primrec Re :: "complex \ real" where Re: "Re (Complex x y) = x" primrec Im :: "complex \ real" where Im: "Im (Complex x y) = y" lemma complex_surj [simp]: "Complex (Re z) (Im z) = z" by (induct z) simp lemma complex_equality [intro?]: "\Re x = Re y; Im x = Im y\ \ x = y" by (induct x, induct y) simp lemma expand_complex_eq: "x = y \ Re x = Re y \ Im x = Im y" by (induct x, induct y) simp lemmas complex_Re_Im_cancel_iff = expand_complex_eq subsection {* Addition and Subtraction *} instantiation complex :: ab_group_add begin definition complex_zero_def: "0 = Complex 0 0" definition complex_add_def: "x + y = Complex (Re x + Re y) (Im x + Im y)" definition complex_minus_def: "- x = Complex (- Re x) (- Im x)" definition complex_diff_def: "x - (y\complex) = x + - y" lemma Complex_eq_0 [simp]: "Complex a b = 0 \ a = 0 \ b = 0" by (simp add: complex_zero_def) lemma complex_Re_zero [simp]: "Re 0 = 0" by (simp add: complex_zero_def) lemma complex_Im_zero [simp]: "Im 0 = 0" by (simp add: complex_zero_def) lemma complex_add [simp]: "Complex a b + Complex c d = Complex (a + c) (b + d)" by (simp add: complex_add_def) lemma complex_Re_add [simp]: "Re (x + y) = Re x + Re y" by (simp add: complex_add_def) lemma complex_Im_add [simp]: "Im (x + y) = Im x + Im y" by (simp add: complex_add_def) lemma complex_minus [simp]: "- (Complex a b) = Complex (- a) (- b)" by (simp add: complex_minus_def) lemma complex_Re_minus [simp]: "Re (- x) = - Re x" by (simp add: complex_minus_def) lemma complex_Im_minus [simp]: "Im (- x) = - Im x" by (simp add: complex_minus_def) lemma complex_diff [simp]: "Complex a b - Complex c d = Complex (a - c) (b - d)" by (simp add: complex_diff_def) lemma complex_Re_diff [simp]: "Re (x - y) = Re x - Re y" by (simp add: complex_diff_def) lemma complex_Im_diff [simp]: "Im (x - y) = Im x - Im y" by (simp add: complex_diff_def) instance by intro_classes (simp_all add: complex_add_def complex_diff_def) end subsection {* Multiplication and Division *} instantiation complex :: "{field, division_by_zero}" begin definition complex_one_def: "1 = Complex 1 0" definition complex_mult_def: "x * y = Complex (Re x * Re y - Im x * Im y) (Re x * Im y + Im x * Re y)" definition complex_inverse_def: "inverse x = Complex (Re x / ((Re x)\ + (Im x)\)) (- Im x / ((Re x)\ + (Im x)\))" definition complex_divide_def: "x / (y\complex) = x * inverse y" lemma Complex_eq_1 [simp]: "(Complex a b = 1) = (a = 1 \ b = 0)" by (simp add: complex_one_def) lemma complex_Re_one [simp]: "Re 1 = 1" by (simp add: complex_one_def) lemma complex_Im_one [simp]: "Im 1 = 0" by (simp add: complex_one_def) lemma complex_mult [simp]: "Complex a b * Complex c d = Complex (a * c - b * d) (a * d + b * c)" by (simp add: complex_mult_def) lemma complex_Re_mult [simp]: "Re (x * y) = Re x * Re y - Im x * Im y" by (simp add: complex_mult_def) lemma complex_Im_mult [simp]: "Im (x * y) = Re x * Im y + Im x * Re y" by (simp add: complex_mult_def) lemma complex_inverse [simp]: "inverse (Complex a b) = Complex (a / (a\ + b\)) (- b / (a\ + b\))" by (simp add: complex_inverse_def) lemma complex_Re_inverse: "Re (inverse x) = Re x / ((Re x)\ + (Im x)\)" by (simp add: complex_inverse_def) lemma complex_Im_inverse: "Im (inverse x) = - Im x / ((Re x)\ + (Im x)\)" by (simp add: complex_inverse_def) instance by intro_classes (simp_all add: complex_mult_def right_distrib left_distrib right_diff_distrib left_diff_distrib complex_inverse_def complex_divide_def power2_eq_square add_divide_distrib [symmetric] expand_complex_eq) end subsection {* Exponentiation *} instantiation complex :: recpower begin primrec power_complex where "z ^ 0 = (1\complex)" | "z ^ Suc n = (z\complex) * z ^ n" instance proof qed simp_all declare power_complex.simps [simp del] end subsection {* Numerals and Arithmetic *} instantiation complex :: number_ring begin definition number_of_complex where complex_number_of_def: "number_of w = (of_int w \ complex)" instance by intro_classes (simp only: complex_number_of_def) end lemma complex_Re_of_nat [simp]: "Re (of_nat n) = of_nat n" by (induct n) simp_all lemma complex_Im_of_nat [simp]: "Im (of_nat n) = 0" by (induct n) simp_all lemma complex_Re_of_int [simp]: "Re (of_int z) = of_int z" by (cases z rule: int_diff_cases) simp lemma complex_Im_of_int [simp]: "Im (of_int z) = 0" by (cases z rule: int_diff_cases) simp lemma complex_Re_number_of [simp]: "Re (number_of v) = number_of v" unfolding number_of_eq by (rule complex_Re_of_int) lemma complex_Im_number_of [simp]: "Im (number_of v) = 0" unfolding number_of_eq by (rule complex_Im_of_int) lemma Complex_eq_number_of [simp]: "(Complex a b = number_of w) = (a = number_of w \ b = 0)" by (simp add: expand_complex_eq) subsection {* Scalar Multiplication *} instantiation complex :: real_field begin definition complex_scaleR_def: "scaleR r x = Complex (r * Re x) (r * Im x)" lemma complex_scaleR [simp]: "scaleR r (Complex a b) = Complex (r * a) (r * b)" unfolding complex_scaleR_def by simp lemma complex_Re_scaleR [simp]: "Re (scaleR r x) = r * Re x" unfolding complex_scaleR_def by simp lemma complex_Im_scaleR [simp]: "Im (scaleR r x) = r * Im x" unfolding complex_scaleR_def by simp instance proof fix a b :: real and x y :: complex show "scaleR a (x + y) = scaleR a x + scaleR a y" by (simp add: expand_complex_eq right_distrib) show "scaleR (a + b) x = scaleR a x + scaleR b x" by (simp add: expand_complex_eq left_distrib) show "scaleR a (scaleR b x) = scaleR (a * b) x" by (simp add: expand_complex_eq mult_assoc) show "scaleR 1 x = x" by (simp add: expand_complex_eq) show "scaleR a x * y = scaleR a (x * y)" by (simp add: expand_complex_eq algebra_simps) show "x * scaleR a y = scaleR a (x * y)" by (simp add: expand_complex_eq algebra_simps) qed end subsection{* Properties of Embedding from Reals *} abbreviation complex_of_real :: "real \ complex" where "complex_of_real \ of_real" lemma complex_of_real_def: "complex_of_real r = Complex r 0" by (simp add: of_real_def complex_scaleR_def) lemma Re_complex_of_real [simp]: "Re (complex_of_real z) = z" by (simp add: complex_of_real_def) lemma Im_complex_of_real [simp]: "Im (complex_of_real z) = 0" by (simp add: complex_of_real_def) lemma Complex_add_complex_of_real [simp]: "Complex x y + complex_of_real r = Complex (x+r) y" by (simp add: complex_of_real_def) lemma complex_of_real_add_Complex [simp]: "complex_of_real r + Complex x y = Complex (r+x) y" by (simp add: complex_of_real_def) lemma Complex_mult_complex_of_real: "Complex x y * complex_of_real r = Complex (x*r) (y*r)" by (simp add: complex_of_real_def) lemma complex_of_real_mult_Complex: "complex_of_real r * Complex x y = Complex (r*x) (r*y)" by (simp add: complex_of_real_def) subsection {* Vector Norm *} instantiation complex :: real_normed_field begin definition complex_norm_def: "norm z = sqrt ((Re z)\ + (Im z)\)" abbreviation cmod :: "complex \ real" where "cmod \ norm" definition complex_sgn_def: "sgn x = x /\<^sub>R cmod x" lemmas cmod_def = complex_norm_def lemma complex_norm [simp]: "cmod (Complex x y) = sqrt (x\ + y\)" by (simp add: complex_norm_def) instance proof fix r :: real and x y :: complex show "0 \ norm x" by (induct x) simp show "(norm x = 0) = (x = 0)" by (induct x) simp show "norm (x + y) \ norm x + norm y" by (induct x, induct y) (simp add: real_sqrt_sum_squares_triangle_ineq) show "norm (scaleR r x) = \r\ * norm x" by (induct x) (simp add: power_mult_distrib right_distrib [symmetric] real_sqrt_mult) show "norm (x * y) = norm x * norm y" by (induct x, induct y) (simp add: real_sqrt_mult [symmetric] power2_eq_square algebra_simps) show "sgn x = x /\<^sub>R cmod x" by(simp add: complex_sgn_def) qed end lemma cmod_unit_one [simp]: "cmod (Complex (cos a) (sin a)) = 1" by simp lemma cmod_complex_polar [simp]: "cmod (complex_of_real r * Complex (cos a) (sin a)) = abs r" by (simp add: norm_mult) lemma complex_Re_le_cmod: "Re x \ cmod x" unfolding complex_norm_def by (rule real_sqrt_sum_squares_ge1) lemma complex_mod_minus_le_complex_mod [simp]: "- cmod x \ cmod x" by (rule order_trans [OF _ norm_ge_zero], simp) lemma complex_mod_triangle_ineq2 [simp]: "cmod(b + a) - cmod b \ cmod a" by (rule ord_le_eq_trans [OF norm_triangle_ineq2], simp) lemmas real_sum_squared_expand = power2_sum [where 'a=real] lemma abs_Re_le_cmod: "\Re x\ \ cmod x" by (cases x) simp lemma abs_Im_le_cmod: "\Im x\ \ cmod x" by (cases x) simp subsection {* Completeness of the Complexes *} -interpretation Re!: bounded_linear "Re" +interpretation Re: bounded_linear "Re" apply (unfold_locales, simp, simp) apply (rule_tac x=1 in exI) apply (simp add: complex_norm_def) done -interpretation Im!: bounded_linear "Im" +interpretation Im: bounded_linear "Im" apply (unfold_locales, simp, simp) apply (rule_tac x=1 in exI) apply (simp add: complex_norm_def) done lemma LIMSEQ_Complex: "\X ----> a; Y ----> b\ \ (\n. Complex (X n) (Y n)) ----> Complex a b" apply (rule LIMSEQ_I) apply (subgoal_tac "0 < r / sqrt 2") apply (drule_tac r="r / sqrt 2" in LIMSEQ_D, safe) apply (drule_tac r="r / sqrt 2" in LIMSEQ_D, safe) apply (rename_tac M N, rule_tac x="max M N" in exI, safe) apply (simp add: real_sqrt_sum_squares_less) apply (simp add: divide_pos_pos) done instance complex :: banach proof fix X :: "nat \ complex" assume X: "Cauchy X" from Re.Cauchy [OF X] have 1: "(\n. Re (X n)) ----> lim (\n. Re (X n))" by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) from Im.Cauchy [OF X] have 2: "(\n. Im (X n)) ----> lim (\n. Im (X n))" by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) have "X ----> Complex (lim (\n. Re (X n))) (lim (\n. Im (X n)))" using LIMSEQ_Complex [OF 1 2] by simp thus "convergent X" by (rule convergentI) qed subsection {* The Complex Number @{term "\"} *} definition "ii" :: complex ("\") where i_def: "ii \ Complex 0 1" lemma complex_Re_i [simp]: "Re ii = 0" by (simp add: i_def) lemma complex_Im_i [simp]: "Im ii = 1" by (simp add: i_def) lemma Complex_eq_i [simp]: "(Complex x y = ii) = (x = 0 \ y = 1)" by (simp add: i_def) lemma complex_i_not_zero [simp]: "ii \ 0" by (simp add: expand_complex_eq) lemma complex_i_not_one [simp]: "ii \ 1" by (simp add: expand_complex_eq) lemma complex_i_not_number_of [simp]: "ii \ number_of w" by (simp add: expand_complex_eq) lemma i_mult_Complex [simp]: "ii * Complex a b = Complex (- b) a" by (simp add: expand_complex_eq) lemma Complex_mult_i [simp]: "Complex a b * ii = Complex (- b) a" by (simp add: expand_complex_eq) lemma i_complex_of_real [simp]: "ii * complex_of_real r = Complex 0 r" by (simp add: i_def complex_of_real_def) lemma complex_of_real_i [simp]: "complex_of_real r * ii = Complex 0 r" by (simp add: i_def complex_of_real_def) lemma i_squared [simp]: "ii * ii = -1" by (simp add: i_def) lemma power2_i [simp]: "ii\ = -1" by (simp add: power2_eq_square) lemma inverse_i [simp]: "inverse ii = - ii" by (rule inverse_unique, simp) subsection {* Complex Conjugation *} definition cnj :: "complex \ complex" where "cnj z = Complex (Re z) (- Im z)" lemma complex_cnj [simp]: "cnj (Complex a b) = Complex a (- b)" by (simp add: cnj_def) lemma complex_Re_cnj [simp]: "Re (cnj x) = Re x" by (simp add: cnj_def) lemma complex_Im_cnj [simp]: "Im (cnj x) = - Im x" by (simp add: cnj_def) lemma complex_cnj_cancel_iff [simp]: "(cnj x = cnj y) = (x = y)" by (simp add: expand_complex_eq) lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z" by (simp add: cnj_def) lemma complex_cnj_zero [simp]: "cnj 0 = 0" by (simp add: expand_complex_eq) lemma complex_cnj_zero_iff [iff]: "(cnj z = 0) = (z = 0)" by (simp add: expand_complex_eq) lemma complex_cnj_add: "cnj (x + y) = cnj x + cnj y" by (simp add: expand_complex_eq) lemma complex_cnj_diff: "cnj (x - y) = cnj x - cnj y" by (simp add: expand_complex_eq) lemma complex_cnj_minus: "cnj (- x) = - cnj x" by (simp add: expand_complex_eq) lemma complex_cnj_one [simp]: "cnj 1 = 1" by (simp add: expand_complex_eq) lemma complex_cnj_mult: "cnj (x * y) = cnj x * cnj y" by (simp add: expand_complex_eq) lemma complex_cnj_inverse: "cnj (inverse x) = inverse (cnj x)" by (simp add: complex_inverse_def) lemma complex_cnj_divide: "cnj (x / y) = cnj x / cnj y" by (simp add: complex_divide_def complex_cnj_mult complex_cnj_inverse) lemma complex_cnj_power: "cnj (x ^ n) = cnj x ^ n" by (induct n, simp_all add: complex_cnj_mult) lemma complex_cnj_of_nat [simp]: "cnj (of_nat n) = of_nat n" by (simp add: expand_complex_eq) lemma complex_cnj_of_int [simp]: "cnj (of_int z) = of_int z" by (simp add: expand_complex_eq) lemma complex_cnj_number_of [simp]: "cnj (number_of w) = number_of w" by (simp add: expand_complex_eq) lemma complex_cnj_scaleR: "cnj (scaleR r x) = scaleR r (cnj x)" by (simp add: expand_complex_eq) lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" by (simp add: complex_norm_def) lemma complex_cnj_complex_of_real [simp]: "cnj (of_real x) = of_real x" by (simp add: expand_complex_eq) lemma complex_cnj_i [simp]: "cnj ii = - ii" by (simp add: expand_complex_eq) lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re z)" by (simp add: expand_complex_eq) lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im z) * ii" by (simp add: expand_complex_eq) lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\ + (Im z)\)" by (simp add: expand_complex_eq power2_eq_square) lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\" by (simp add: norm_mult power2_eq_square) -interpretation cnj!: bounded_linear "cnj" +interpretation cnj: bounded_linear "cnj" apply (unfold_locales) apply (rule complex_cnj_add) apply (rule complex_cnj_scaleR) apply (rule_tac x=1 in exI, simp) done subsection{*The Functions @{term sgn} and @{term arg}*} text {*------------ Argand -------------*} definition arg :: "complex => real" where "arg z = (SOME a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a \ pi)" lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" by (simp add: complex_sgn_def divide_inverse scaleR_conv_of_real mult_commute) lemma i_mult_eq: "ii * ii = complex_of_real (-1)" by (simp add: i_def complex_of_real_def) lemma i_mult_eq2 [simp]: "ii * ii = -(1::complex)" by (simp add: i_def complex_one_def) lemma complex_eq_cancel_iff2 [simp]: "(Complex x y = complex_of_real xa) = (x = xa & y = 0)" by (simp add: complex_of_real_def) lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) lemma complex_inverse_complex_split: "inverse(complex_of_real x + ii * complex_of_real y) = complex_of_real(x/(x ^ 2 + y ^ 2)) - ii * complex_of_real(y/(x ^ 2 + y ^ 2))" by (simp add: complex_of_real_def i_def diff_minus divide_inverse) (*----------------------------------------------------------------------------*) (* Many of the theorems below need to be moved elsewhere e.g. Transc. Also *) (* many of the theorems are not used - so should they be kept? *) (*----------------------------------------------------------------------------*) lemma cos_arg_i_mult_zero_pos: "0 < y ==> cos (arg(Complex 0 y)) = 0" apply (simp add: arg_def abs_if) apply (rule_tac a = "pi/2" in someI2, auto) apply (rule order_less_trans [of _ 0], auto) done lemma cos_arg_i_mult_zero_neg: "y < 0 ==> cos (arg(Complex 0 y)) = 0" apply (simp add: arg_def abs_if) apply (rule_tac a = "- pi/2" in someI2, auto) apply (rule order_trans [of _ 0], auto) done lemma cos_arg_i_mult_zero [simp]: "y \ 0 ==> cos (arg(Complex 0 y)) = 0" by (auto simp add: linorder_neq_iff cos_arg_i_mult_zero_pos cos_arg_i_mult_zero_neg) subsection{*Finally! Polar Form for Complex Numbers*} definition (* abbreviation for (cos a + i sin a) *) cis :: "real => complex" where "cis a = Complex (cos a) (sin a)" definition (* abbreviation for r*(cos a + i sin a) *) rcis :: "[real, real] => complex" where "rcis r a = complex_of_real r * cis a" definition (* e ^ (x + iy) *) expi :: "complex => complex" where "expi z = complex_of_real(exp (Re z)) * cis (Im z)" lemma complex_split_polar: "\r a. z = complex_of_real r * (Complex (cos a) (sin a))" apply (induct z) apply (auto simp add: polar_Ex complex_of_real_mult_Complex) done lemma rcis_Ex: "\r a. z = rcis r a" apply (induct z) apply (simp add: rcis_def cis_def polar_Ex complex_of_real_mult_Complex) done lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a" by (simp add: rcis_def cis_def) lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" by (simp add: rcis_def cis_def) lemma sin_cos_squared_add2_mult: "(r * cos a)\ + (r * sin a)\ = r\" proof - have "(r * cos a)\ + (r * sin a)\ = r\ * ((cos a)\ + (sin a)\)" by (simp only: power_mult_distrib right_distrib) thus ?thesis by simp qed lemma complex_mod_rcis [simp]: "cmod(rcis r a) = abs r" by (simp add: rcis_def cis_def sin_cos_squared_add2_mult) lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" by (simp add: cmod_def power2_eq_square) lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0" by simp (*---------------------------------------------------------------------------*) (* (r1 * cis a) * (r2 * cis b) = r1 * r2 * cis (a + b) *) (*---------------------------------------------------------------------------*) lemma cis_rcis_eq: "cis a = rcis 1 a" by (simp add: rcis_def) lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1*r2) (a + b)" by (simp add: rcis_def cis_def cos_add sin_add right_distrib right_diff_distrib complex_of_real_def) lemma cis_mult: "cis a * cis b = cis (a + b)" by (simp add: cis_rcis_eq rcis_mult) lemma cis_zero [simp]: "cis 0 = 1" by (simp add: cis_def complex_one_def) lemma rcis_zero_mod [simp]: "rcis 0 a = 0" by (simp add: rcis_def) lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r" by (simp add: rcis_def) lemma complex_of_real_minus_one: "complex_of_real (-(1::real)) = -(1::complex)" by (simp add: complex_of_real_def complex_one_def) lemma complex_i_mult_minus [simp]: "ii * (ii * x) = - x" by (simp add: mult_assoc [symmetric]) lemma cis_real_of_nat_Suc_mult: "cis (real (Suc n) * a) = cis a * cis (real n * a)" by (simp add: cis_def real_of_nat_Suc left_distrib cos_add sin_add right_distrib) lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" apply (induct_tac "n") apply (auto simp add: cis_real_of_nat_Suc_mult) done lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)" by (simp add: rcis_def power_mult_distrib DeMoivre) lemma cis_inverse [simp]: "inverse(cis a) = cis (-a)" by (simp add: cis_def complex_inverse_complex_split diff_minus) lemma rcis_inverse: "inverse(rcis r a) = rcis (1/r) (-a)" by (simp add: divide_inverse rcis_def) lemma cis_divide: "cis a / cis b = cis (a - b)" by (simp add: complex_divide_def cis_mult real_diff_def) lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)" apply (simp add: complex_divide_def) apply (case_tac "r2=0", simp) apply (simp add: rcis_inverse rcis_mult real_diff_def) done lemma Re_cis [simp]: "Re(cis a) = cos a" by (simp add: cis_def) lemma Im_cis [simp]: "Im(cis a) = sin a" by (simp add: cis_def) lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)" by (auto simp add: DeMoivre) lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im(cis a ^ n)" by (auto simp add: DeMoivre) lemma expi_add: "expi(a + b) = expi(a) * expi(b)" by (simp add: expi_def exp_add cis_mult [symmetric] mult_ac) lemma expi_zero [simp]: "expi (0::complex) = 1" by (simp add: expi_def) lemma complex_expi_Ex: "\a r. z = complex_of_real r * expi a" apply (insert rcis_Ex [of z]) apply (auto simp add: expi_def rcis_def mult_assoc [symmetric]) apply (rule_tac x = "ii * complex_of_real a" in exI, auto) done lemma expi_two_pi_i [simp]: "expi((2::complex) * complex_of_real pi * ii) = 1" by (simp add: expi_def cis_def) end diff --git a/src/HOL/Decision_Procs/Dense_Linear_Order.thy b/src/HOL/Decision_Procs/Dense_Linear_Order.thy --- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy +++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy @@ -1,932 +1,932 @@ (* Title : HOL/Decision_Procs/Dense_Linear_Order.thy Author : Amine Chaieb, TU Muenchen *) header {* Dense linear order without endpoints and a quantifier elimination procedure in Ferrante and Rackoff style *} theory Dense_Linear_Order imports Plain Groebner_Basis Main uses "~~/src/HOL/Tools/Qelim/langford_data.ML" "~~/src/HOL/Tools/Qelim/ferrante_rackoff_data.ML" ("~~/src/HOL/Tools/Qelim/langford.ML") ("~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML") begin setup {* Langford_Data.setup #> Ferrante_Rackoff_Data.setup *} context linorder begin lemma less_not_permute[noatp]: "\ (x < y \ y < x)" by (simp add: not_less linear) lemma gather_simps[noatp]: shows "(\x. (\y \ L. y < x) \ (\y \ U. x < y) \ x < u \ P x) \ (\x. (\y \ L. y < x) \ (\y \ (insert u U). x < y) \ P x)" and "(\x. (\y \ L. y < x) \ (\y \ U. x < y) \ l < x \ P x) \ (\x. (\y \ (insert l L). y < x) \ (\y \ U. x < y) \ P x)" "(\x. (\y \ L. y < x) \ (\y \ U. x < y) \ x < u) \ (\x. (\y \ L. y < x) \ (\y \ (insert u U). x < y))" and "(\x. (\y \ L. y < x) \ (\y \ U. x < y) \ l < x) \ (\x. (\y \ (insert l L). y < x) \ (\y \ U. x < y))" by auto lemma gather_start[noatp]: "(\x. P x) \ (\x. (\y \ {}. y < x) \ (\y\ {}. x < y) \ P x)" by simp text{* Theorems for @{text "\z. \x. x < z \ (P x \ P\<^bsub>-\\<^esub>)"}*} lemma minf_lt[noatp]: "\z . \x. x < z \ (x < t \ True)" by auto lemma minf_gt[noatp]: "\z . \x. x < z \ (t < x \ False)" by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le) lemma minf_le[noatp]: "\z. \x. x < z \ (x \ t \ True)" by (auto simp add: less_le) lemma minf_ge[noatp]: "\z. \x. x < z \ (t \ x \ False)" by (auto simp add: less_le not_less not_le) lemma minf_eq[noatp]: "\z. \x. x < z \ (x = t \ False)" by auto lemma minf_neq[noatp]: "\z. \x. x < z \ (x \ t \ True)" by auto lemma minf_P[noatp]: "\z. \x. x < z \ (P \ P)" by blast text{* Theorems for @{text "\z. \x. x < z \ (P x \ P\<^bsub>+\\<^esub>)"}*} lemma pinf_gt[noatp]: "\z . \x. z < x \ (t < x \ True)" by auto lemma pinf_lt[noatp]: "\z . \x. z < x \ (x < t \ False)" by (simp add: not_less) (rule exI[where x="t"], auto simp add: less_le) lemma pinf_ge[noatp]: "\z. \x. z < x \ (t \ x \ True)" by (auto simp add: less_le) lemma pinf_le[noatp]: "\z. \x. z < x \ (x \ t \ False)" by (auto simp add: less_le not_less not_le) lemma pinf_eq[noatp]: "\z. \x. z < x \ (x = t \ False)" by auto lemma pinf_neq[noatp]: "\z. \x. z < x \ (x \ t \ True)" by auto lemma pinf_P[noatp]: "\z. \x. z < x \ (P \ P)" by blast lemma nmi_lt[noatp]: "t \ U \ \x. \True \ x < t \ (\ u\ U. u \ x)" by auto lemma nmi_gt[noatp]: "t \ U \ \x. \False \ t < x \ (\ u\ U. u \ x)" by (auto simp add: le_less) lemma nmi_le[noatp]: "t \ U \ \x. \True \ x\ t \ (\ u\ U. u \ x)" by auto lemma nmi_ge[noatp]: "t \ U \ \x. \False \ t\ x \ (\ u\ U. u \ x)" by auto lemma nmi_eq[noatp]: "t \ U \ \x. \False \ x = t \ (\ u\ U. u \ x)" by auto lemma nmi_neq[noatp]: "t \ U \\x. \True \ x \ t \ (\ u\ U. u \ x)" by auto lemma nmi_P[noatp]: "\ x. ~P \ P \ (\ u\ U. u \ x)" by auto lemma nmi_conj[noatp]: "\\x. \P1' \ P1 x \ (\ u\ U. u \ x) ; \x. \P2' \ P2 x \ (\ u\ U. u \ x)\ \ \x. \(P1' \ P2') \ (P1 x \ P2 x) \ (\ u\ U. u \ x)" by auto lemma nmi_disj[noatp]: "\\x. \P1' \ P1 x \ (\ u\ U. u \ x) ; \x. \P2' \ P2 x \ (\ u\ U. u \ x)\ \ \x. \(P1' \ P2') \ (P1 x \ P2 x) \ (\ u\ U. u \ x)" by auto lemma npi_lt[noatp]: "t \ U \ \x. \False \ x < t \ (\ u\ U. x \ u)" by (auto simp add: le_less) lemma npi_gt[noatp]: "t \ U \ \x. \True \ t < x \ (\ u\ U. x \ u)" by auto lemma npi_le[noatp]: "t \ U \ \x. \False \ x \ t \ (\ u\ U. x \ u)" by auto lemma npi_ge[noatp]: "t \ U \ \x. \True \ t \ x \ (\ u\ U. x \ u)" by auto lemma npi_eq[noatp]: "t \ U \ \x. \False \ x = t \ (\ u\ U. x \ u)" by auto lemma npi_neq[noatp]: "t \ U \ \x. \True \ x \ t \ (\ u\ U. x \ u )" by auto lemma npi_P[noatp]: "\ x. ~P \ P \ (\ u\ U. x \ u)" by auto lemma npi_conj[noatp]: "\\x. \P1' \ P1 x \ (\ u\ U. x \ u) ; \x. \P2' \ P2 x \ (\ u\ U. x \ u)\ \ \x. \(P1' \ P2') \ (P1 x \ P2 x) \ (\ u\ U. x \ u)" by auto lemma npi_disj[noatp]: "\\x. \P1' \ P1 x \ (\ u\ U. x \ u) ; \x. \P2' \ P2 x \ (\ u\ U. x \ u)\ \ \x. \(P1' \ P2') \ (P1 x \ P2 x) \ (\ u\ U. x \ u)" by auto lemma lin_dense_lt[noatp]: "t \ U \ \x l u. (\ t. l < t \ t < u \ t \ U) \ l< x \ x < u \ x < t \ (\ y. l < y \ y < u \ y < t)" proof(clarsimp) fix x l u y assume tU: "t \ U" and noU: "\t. l < t \ t < u \ t \ U" and lx: "l < x" and xu: "xy" by auto {assume H: "t < y" from less_trans[OF lx px] less_trans[OF H yu] have "l < t \ t < u" by simp with tU noU have "False" by auto} hence "\ t < y" by auto hence "y \ t" by (simp add: not_less) thus "y < t" using tny by (simp add: less_le) qed lemma lin_dense_gt[noatp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l < x \ x < u \ t < x \ (\ y. l < y \ y < u \ t < y)" proof(clarsimp) fix x l u y assume tU: "t \ U" and noU: "\t. l < t \ t < u \ t \ U" and lx: "l < x" and xu: "xy" by auto {assume H: "y< t" from less_trans[OF ly H] less_trans[OF px xu] have "l < t \ t < u" by simp with tU noU have "False" by auto} hence "\ y y" by (auto simp add: not_less) thus "t < y" using tny by (simp add:less_le) qed lemma lin_dense_le[noatp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x \ t \ (\ y. l < y \ y < u \ y\ t)" proof(clarsimp) fix x l u y assume tU: "t \ U" and noU: "\t. l < t \ t < u \ t \ U" and lx: "l < x" and xu: "x t" and ly: "ly" by auto {assume H: "t < y" from less_le_trans[OF lx px] less_trans[OF H yu] have "l < t \ t < u" by simp with tU noU have "False" by auto} hence "\ t < y" by auto thus "y \ t" by (simp add: not_less) qed lemma lin_dense_ge[noatp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ t \ x \ (\ y. l < y \ y < u \ t \ y)" proof(clarsimp) fix x l u y assume tU: "t \ U" and noU: "\t. l < t \ t < u \ t \ U" and lx: "l < x" and xu: "x x" and ly: "ly" by auto {assume H: "y< t" from less_trans[OF ly H] le_less_trans[OF px xu] have "l < t \ t < u" by simp with tU noU have "False" by auto} hence "\ y y" by (simp add: not_less) qed lemma lin_dense_eq[noatp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x = t \ (\ y. l < y \ y < u \ y= t)" by auto lemma lin_dense_neq[noatp]: "t \ U \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ x \ t \ (\ y. l < y \ y < u \ y\ t)" by auto lemma lin_dense_P[noatp]: "\x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P \ (\ y. l < y \ y < u \ P)" by auto lemma lin_dense_conj[noatp]: "\\x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P1 x \ (\ y. l < y \ y < u \ P1 y) ; \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P2 x \ (\ y. l < y \ y < u \ P2 y)\ \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ (P1 x \ P2 x) \ (\ y. l < y \ y < u \ (P1 y \ P2 y))" by blast lemma lin_dense_disj[noatp]: "\\x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P1 x \ (\ y. l < y \ y < u \ P1 y) ; \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ P2 x \ (\ y. l < y \ y < u \ P2 y)\ \ \x l u. (\ t. l < t \ t< u \ t \ U) \ l< x \ x < u \ (P1 x \ P2 x) \ (\ y. l < y \ y < u \ (P1 y \ P2 y))" by blast lemma npmibnd[noatp]: "\\x. \ MP \ P x \ (\ u\ U. u \ x); \x. \PP \ P x \ (\ u\ U. x \ u)\ \ \x. \ MP \ \PP \ P x \ (\ u\ U. \ u' \ U. u \ x \ x \ u')" by auto lemma finite_set_intervals[noatp]: assumes px: "P x" and lx: "l \ x" and xu: "x \ u" and linS: "l\ S" and uinS: "u \ S" and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" shows "\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a \ x \ x \ b \ P x" proof- let ?Mx = "{y. y\ S \ y \ x}" let ?xM = "{y. y\ S \ x \ y}" let ?a = "Max ?Mx" let ?b = "Min ?xM" have MxS: "?Mx \ S" by blast hence fMx: "finite ?Mx" using fS finite_subset by auto from lx linS have linMx: "l \ ?Mx" by blast hence Mxne: "?Mx \ {}" by blast have xMS: "?xM \ S" by blast hence fxM: "finite ?xM" using fS finite_subset by auto from xu uinS have linxM: "u \ ?xM" by blast hence xMne: "?xM \ {}" by blast have ax:"?a \ x" using Mxne fMx by auto have xb:"x \ ?b" using xMne fxM by auto have "?a \ ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \ S" using MxS by blast have "?b \ ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \ S" using xMS by blast have noy:"\ y. ?a < y \ y < ?b \ y \ S" proof(clarsimp) fix y assume ay: "?a < y" and yb: "y < ?b" and yS: "y \ S" from yS have "y\ ?Mx \ y\ ?xM" by (auto simp add: linear) moreover {assume "y \ ?Mx" hence "y \ ?a" using Mxne fMx by auto with ay have "False" by (simp add: not_le[symmetric])} moreover {assume "y \ ?xM" hence "?b \ y" using xMne fxM by auto with yb have "False" by (simp add: not_le[symmetric])} ultimately show "False" by blast qed from ainS binS noy ax xb px show ?thesis by blast qed lemma finite_set_intervals2[noatp]: assumes px: "P x" and lx: "l \ x" and xu: "x \ u" and linS: "l\ S" and uinS: "u \ S" and fS:"finite S" and lS: "\ x\ S. l \ x" and Su: "\ x\ S. x \ u" shows "(\ s\ S. P s) \ (\ a \ S. \ b \ S. (\ y. a < y \ y < b \ y \ S) \ a < x \ x < b \ P x)" proof- from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] obtain a and b where as: "a\ S" and bs: "b\ S" and noS:"\y. a < y \ y < b \ y \ S" and axb: "a \ x \ x \ b \ P x" by auto from axb have "x= a \ x= b \ (a < x \ x < b)" by (auto simp add: le_less) thus ?thesis using px as bs noS by blast qed end section {* The classical QE after Langford for dense linear orders *} context dense_linear_order begin lemma interval_empty_iff: "{y. x < y \ y < z} = {} \ \ x < z" by (auto dest: dense) lemma dlo_qe_bnds[noatp]: assumes ne: "L \ {}" and neU: "U \ {}" and fL: "finite L" and fU: "finite U" shows "(\x. (\y \ L. y < x) \ (\y \ U. x < y)) \ (\ l \ L. \u \ U. l < u)" proof (simp only: atomize_eq, rule iffI) assume H: "\x. (\y\L. y < x) \ (\y\U. x < y)" then obtain x where xL: "\y\L. y < x" and xU: "\y\U. x < y" by blast {fix l u assume l: "l \ L" and u: "u \ U" have "l < x" using xL l by blast also have "x < u" using xU u by blast finally (less_trans) have "l < u" .} thus "\l\L. \u\U. l < u" by blast next assume H: "\l\L. \u\U. l < u" let ?ML = "Max L" let ?MU = "Min U" from fL ne have th1: "?ML \ L" and th1': "\l\L. l \ ?ML" by auto from fU neU have th2: "?MU \ U" and th2': "\u\U. ?MU \ u" by auto from th1 th2 H have "?ML < ?MU" by auto with dense obtain w where th3: "?ML < w" and th4: "w < ?MU" by blast from th3 th1' have "\l \ L. l < w" by auto moreover from th4 th2' have "\u \ U. w < u" by auto ultimately show "\x. (\y\L. y < x) \ (\y\U. x < y)" by auto qed lemma dlo_qe_noub[noatp]: assumes ne: "L \ {}" and fL: "finite L" shows "(\x. (\y \ L. y < x) \ (\y \ {}. x < y)) \ True" proof(simp add: atomize_eq) from gt_ex[of "Max L"] obtain M where M: "Max L < M" by blast from ne fL have "\x \ L. x \ Max L" by simp with M have "\x\L. x < M" by (auto intro: le_less_trans) thus "\x. \y\L. y < x" by blast qed lemma dlo_qe_nolb[noatp]: assumes ne: "U \ {}" and fU: "finite U" shows "(\x. (\y \ {}. y < x) \ (\y \ U. x < y)) \ True" proof(simp add: atomize_eq) from lt_ex[of "Min U"] obtain M where M: "M < Min U" by blast from ne fU have "\x \ U. Min U \ x" by simp with M have "\x\U. M < x" by (auto intro: less_le_trans) thus "\x. \y\U. x < y" by blast qed lemma exists_neq[noatp]: "\(x::'a). x \ t" "\(x::'a). t \ x" using gt_ex[of t] by auto lemmas dlo_simps[noatp] = order_refl less_irrefl not_less not_le exists_neq le_less neq_iff linear less_not_permute lemma axiom[noatp]: "dense_linear_order (op \) (op <)" by (rule dense_linear_order_axioms) lemma atoms[noatp]: shows "TERM (less :: 'a \ _)" and "TERM (less_eq :: 'a \ _)" and "TERM (op = :: 'a \ _)" . declare axiom[langford qe: dlo_qe_bnds dlo_qe_nolb dlo_qe_noub gather: gather_start gather_simps atoms: atoms] declare dlo_simps[langfordsimp] end (* FIXME: Move to HOL -- together with the conj_aci_rule in langford.ML *) lemma dnf[noatp]: "(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))" by blast+ lemmas weak_dnf_simps[noatp] = simp_thms dnf lemma nnf_simps[noatp]: "(\(P \ Q)) = (\P \ \Q)" "(\(P \ Q)) = (\P \ \Q)" "(P \ Q) = (\P \ Q)" "(P = Q) = ((P \ Q) \ (\P \ \ Q))" "(\ \(P)) = P" by blast+ lemma ex_distrib[noatp]: "(\x. P x \ Q x) \ ((\x. P x) \ (\x. Q x))" by blast lemmas dnf_simps[noatp] = weak_dnf_simps nnf_simps ex_distrib use "~~/src/HOL/Tools/Qelim/langford.ML" method_setup dlo = {* Scan.succeed (SIMPLE_METHOD' o LangfordQE.dlo_tac) *} "Langford's algorithm for quantifier elimination in dense linear orders" section {* Contructive dense linear orders yield QE for linear arithmetic over ordered Fields *} text {* Linear order without upper bounds *} locale linorder_stupid_syntax = linorder begin notation less_eq ("op \") and less_eq ("(_/ \ _)" [51, 51] 50) and less ("op \") and less ("(_/ \ _)" [51, 51] 50) end locale linorder_no_ub = linorder_stupid_syntax + assumes gt_ex: "\y. less x y" begin lemma ge_ex[noatp]: "\y. x \ y" using gt_ex by auto text {* Theorems for @{text "\z. \x. z \ x \ (P x \ P\<^bsub>+\\<^esub>)"} *} lemma pinf_conj[noatp]: assumes ex1: "\z1. \x. z1 \ x \ (P1 x \ P1')" and ex2: "\z2. \x. z2 \ x \ (P2 x \ P2')" shows "\z. \x. z \ x \ ((P1 x \ P2 x) \ (P1' \ P2'))" proof- from ex1 ex2 obtain z1 and z2 where z1: "\x. z1 \ x \ (P1 x \ P1')" and z2: "\x. z2 \ x \ (P2 x \ P2')" by blast from gt_ex obtain z where z:"ord.max less_eq z1 z2 \ z" by blast from z have zz1: "z1 \ z" and zz2: "z2 \ z" by simp_all {fix x assume H: "z \ x" from less_trans[OF zz1 H] less_trans[OF zz2 H] have "(P1 x \ P2 x) \ (P1' \ P2')" using z1 zz1 z2 zz2 by auto } thus ?thesis by blast qed lemma pinf_disj[noatp]: assumes ex1: "\z1. \x. z1 \ x \ (P1 x \ P1')" and ex2: "\z2. \x. z2 \ x \ (P2 x \ P2')" shows "\z. \x. z \ x \ ((P1 x \ P2 x) \ (P1' \ P2'))" proof- from ex1 ex2 obtain z1 and z2 where z1: "\x. z1 \ x \ (P1 x \ P1')" and z2: "\x. z2 \ x \ (P2 x \ P2')" by blast from gt_ex obtain z where z:"ord.max less_eq z1 z2 \ z" by blast from z have zz1: "z1 \ z" and zz2: "z2 \ z" by simp_all {fix x assume H: "z \ x" from less_trans[OF zz1 H] less_trans[OF zz2 H] have "(P1 x \ P2 x) \ (P1' \ P2')" using z1 zz1 z2 zz2 by auto } thus ?thesis by blast qed lemma pinf_ex[noatp]: assumes ex:"\z. \x. z \ x \ (P x \ P1)" and p1: P1 shows "\ x. P x" proof- from ex obtain z where z: "\x. z \ x \ (P x \ P1)" by blast from gt_ex obtain x where x: "z \ x" by blast from z x p1 show ?thesis by blast qed end text {* Linear order without upper bounds *} locale linorder_no_lb = linorder_stupid_syntax + assumes lt_ex: "\y. less y x" begin lemma le_ex[noatp]: "\y. y \ x" using lt_ex by auto text {* Theorems for @{text "\z. \x. x \ z \ (P x \ P\<^bsub>-\\<^esub>)"} *} lemma minf_conj[noatp]: assumes ex1: "\z1. \x. x \ z1 \ (P1 x \ P1')" and ex2: "\z2. \x. x \ z2 \ (P2 x \ P2')" shows "\z. \x. x \ z \ ((P1 x \ P2 x) \ (P1' \ P2'))" proof- from ex1 ex2 obtain z1 and z2 where z1: "\x. x \ z1 \ (P1 x \ P1')"and z2: "\x. x \ z2 \ (P2 x \ P2')" by blast from lt_ex obtain z where z:"z \ ord.min less_eq z1 z2" by blast from z have zz1: "z \ z1" and zz2: "z \ z2" by simp_all {fix x assume H: "x \ z" from less_trans[OF H zz1] less_trans[OF H zz2] have "(P1 x \ P2 x) \ (P1' \ P2')" using z1 zz1 z2 zz2 by auto } thus ?thesis by blast qed lemma minf_disj[noatp]: assumes ex1: "\z1. \x. x \ z1 \ (P1 x \ P1')" and ex2: "\z2. \x. x \ z2 \ (P2 x \ P2')" shows "\z. \x. x \ z \ ((P1 x \ P2 x) \ (P1' \ P2'))" proof- from ex1 ex2 obtain z1 and z2 where z1: "\x. x \ z1 \ (P1 x \ P1')"and z2: "\x. x \ z2 \ (P2 x \ P2')" by blast from lt_ex obtain z where z:"z \ ord.min less_eq z1 z2" by blast from z have zz1: "z \ z1" and zz2: "z \ z2" by simp_all {fix x assume H: "x \ z" from less_trans[OF H zz1] less_trans[OF H zz2] have "(P1 x \ P2 x) \ (P1' \ P2')" using z1 zz1 z2 zz2 by auto } thus ?thesis by blast qed lemma minf_ex[noatp]: assumes ex:"\z. \x. x \ z \ (P x \ P1)" and p1: P1 shows "\ x. P x" proof- from ex obtain z where z: "\x. x \ z \ (P x \ P1)" by blast from lt_ex obtain x where x: "x \ z" by blast from z x p1 show ?thesis by blast qed end locale constr_dense_linear_order = linorder_no_lb + linorder_no_ub + fixes between assumes between_less: "less x y \ less x (between x y) \ less (between x y) y" and between_same: "between x x = x" sublocale constr_dense_linear_order < dense_linear_order apply unfold_locales using gt_ex lt_ex between_less by (auto, rule_tac x="between x y" in exI, simp) context constr_dense_linear_order begin lemma rinf_U[noatp]: assumes fU: "finite U" and lin_dense: "\x l u. (\ t. l \ t \ t\ u \ t \ U) \ l\ x \ x \ u \ P x \ (\ y. l \ y \ y \ u \ P y )" and nmpiU: "\x. \ MP \ \PP \ P x \ (\ u\ U. \ u' \ U. u \ x \ x \ u')" and nmi: "\ MP" and npi: "\ PP" and ex: "\ x. P x" shows "\ u\ U. \ u' \ U. P (between u u')" proof- from ex obtain x where px: "P x" by blast from px nmi npi nmpiU have "\ u\ U. \ u' \ U. u \ x \ x \ u'" by auto then obtain u and u' where uU:"u\ U" and uU': "u' \ U" and ux:"u \ x" and xu':"x \ u'" by auto from uU have Une: "U \ {}" by auto term "linorder.Min less_eq" let ?l = "linorder.Min less_eq U" let ?u = "linorder.Max less_eq U" have linM: "?l \ U" using fU Une by simp have uinM: "?u \ U" using fU Une by simp have lM: "\ t\ U. ?l \ t" using Une fU by auto have Mu: "\ t\ U. t \ ?u" using Une fU by auto have th:"?l \ u" using uU Une lM by auto from order_trans[OF th ux] have lx: "?l \ x" . have th: "u' \ ?u" using uU' Une Mu by simp from order_trans[OF xu' th] have xu: "x \ ?u" . from finite_set_intervals2[where P="P",OF px lx xu linM uinM fU lM Mu] have "(\ s\ U. P s) \ (\ t1\ U. \ t2 \ U. (\ y. t1 \ y \ y \ t2 \ y \ U) \ t1 \ x \ x \ t2 \ P x)" . moreover { fix u assume um: "u\U" and pu: "P u" have "between u u = u" by (simp add: between_same) with um pu have "P (between u u)" by simp with um have ?thesis by blast} moreover{ assume "\ t1\ U. \ t2 \ U. (\ y. t1 \ y \ y \ t2 \ y \ U) \ t1 \ x \ x \ t2 \ P x" then obtain t1 and t2 where t1M: "t1 \ U" and t2M: "t2\ U" and noM: "\ y. t1 \ y \ y \ t2 \ y \ U" and t1x: "t1 \ x" and xt2: "x \ t2" and px: "P x" by blast from less_trans[OF t1x xt2] have t1t2: "t1 \ t2" . let ?u = "between t1 t2" from between_less t1t2 have t1lu: "t1 \ ?u" and ut2: "?u \ t2" by auto from lin_dense noM t1x xt2 px t1lu ut2 have "P ?u" by blast with t1M t2M have ?thesis by blast} ultimately show ?thesis by blast qed theorem fr_eq[noatp]: assumes fU: "finite U" and lin_dense: "\x l u. (\ t. l \ t \ t\ u \ t \ U) \ l\ x \ x \ u \ P x \ (\ y. l \ y \ y \ u \ P y )" and nmibnd: "\x. \ MP \ P x \ (\ u\ U. u \ x)" and npibnd: "\x. \PP \ P x \ (\ u\ U. x \ u)" and mi: "\z. \x. x \ z \ (P x = MP)" and pi: "\z. \x. z \ x \ (P x = PP)" shows "(\ x. P x) \ (MP \ PP \ (\ u \ U. \ u'\ U. P (between u u')))" (is "_ \ (_ \ _ \ ?F)" is "?E \ ?D") proof- { assume px: "\ x. P x" have "MP \ PP \ (\ MP \ \ PP)" by blast moreover {assume "MP \ PP" hence "?D" by blast} moreover {assume nmi: "\ MP" and npi: "\ PP" from npmibnd[OF nmibnd npibnd] have nmpiU: "\x. \ MP \ \PP \ P x \ (\ u\ U. \ u' \ U. u \ x \ x \ u')" . from rinf_U[OF fU lin_dense nmpiU nmi npi px] have "?D" by blast} ultimately have "?D" by blast} moreover { assume "?D" moreover {assume m:"MP" from minf_ex[OF mi m] have "?E" .} moreover {assume p: "PP" from pinf_ex[OF pi p] have "?E" . } moreover {assume f:"?F" hence "?E" by blast} ultimately have "?E" by blast} ultimately have "?E = ?D" by blast thus "?E \ ?D" by simp qed lemmas minf_thms[noatp] = minf_conj minf_disj minf_eq minf_neq minf_lt minf_le minf_gt minf_ge minf_P lemmas pinf_thms[noatp] = pinf_conj pinf_disj pinf_eq pinf_neq pinf_lt pinf_le pinf_gt pinf_ge pinf_P lemmas nmi_thms[noatp] = nmi_conj nmi_disj nmi_eq nmi_neq nmi_lt nmi_le nmi_gt nmi_ge nmi_P lemmas npi_thms[noatp] = npi_conj npi_disj npi_eq npi_neq npi_lt npi_le npi_gt npi_ge npi_P lemmas lin_dense_thms[noatp] = lin_dense_conj lin_dense_disj lin_dense_eq lin_dense_neq lin_dense_lt lin_dense_le lin_dense_gt lin_dense_ge lin_dense_P lemma ferrack_axiom[noatp]: "constr_dense_linear_order less_eq less between" by (rule constr_dense_linear_order_axioms) lemma atoms[noatp]: shows "TERM (less :: 'a \ _)" and "TERM (less_eq :: 'a \ _)" and "TERM (op = :: 'a \ _)" . declare ferrack_axiom [ferrack minf: minf_thms pinf: pinf_thms nmi: nmi_thms npi: npi_thms lindense: lin_dense_thms qe: fr_eq atoms: atoms] declaration {* let fun simps phi = map (Morphism.thm phi) [@{thm "not_less"}, @{thm "not_le"}] fun generic_whatis phi = let val [lt, le] = map (Morphism.term phi) [@{term "op \"}, @{term "op \"}] fun h x t = case term_of t of Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq else Ferrante_Rackoff_Data.Nox | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq else Ferrante_Rackoff_Data.Nox | b$y$z => if Term.could_unify (b, lt) then if term_of x aconv y then Ferrante_Rackoff_Data.Lt else if term_of x aconv z then Ferrante_Rackoff_Data.Gt else Ferrante_Rackoff_Data.Nox else if Term.could_unify (b, le) then if term_of x aconv y then Ferrante_Rackoff_Data.Le else if term_of x aconv z then Ferrante_Rackoff_Data.Ge else Ferrante_Rackoff_Data.Nox else Ferrante_Rackoff_Data.Nox | _ => Ferrante_Rackoff_Data.Nox in h end fun ss phi = HOL_ss addsimps (simps phi) in Ferrante_Rackoff_Data.funs @{thm "ferrack_axiom"} {isolate_conv = K (K (K Thm.reflexive)), whatis = generic_whatis, simpset = ss} end *} end use "~~/src/HOL/Tools/Qelim/ferrante_rackoff.ML" method_setup ferrack = {* Scan.succeed (SIMPLE_METHOD' o FerranteRackoff.dlo_tac) *} "Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders" subsection {* Ferrante and Rackoff algorithm over ordered fields *} lemma neg_prod_lt:"(c\'a\ordered_field) < 0 \ ((c*x < 0) == (x > 0))" proof- assume H: "c < 0" have "c*x < 0 = (0/c < x)" by (simp only: neg_divide_less_eq[OF H] algebra_simps) also have "\ = (0 < x)" by simp finally show "(c*x < 0) == (x > 0)" by simp qed lemma pos_prod_lt:"(c\'a\ordered_field) > 0 \ ((c*x < 0) == (x < 0))" proof- assume H: "c > 0" hence "c*x < 0 = (0/c > x)" by (simp only: pos_less_divide_eq[OF H] algebra_simps) also have "\ = (0 > x)" by simp finally show "(c*x < 0) == (x < 0)" by simp qed lemma neg_prod_sum_lt: "(c\'a\ordered_field) < 0 \ ((c*x + t< 0) == (x > (- 1/c)*t))" proof- assume H: "c < 0" have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp) also have "\ = (-t/c < x)" by (simp only: neg_divide_less_eq[OF H] algebra_simps) also have "\ = ((- 1/c)*t < x)" by simp finally show "(c*x + t < 0) == (x > (- 1/c)*t)" by simp qed lemma pos_prod_sum_lt:"(c\'a\ordered_field) > 0 \ ((c*x + t < 0) == (x < (- 1/c)*t))" proof- assume H: "c > 0" have "c*x + t< 0 = (c*x < -t)" by (subst less_iff_diff_less_0 [of "c*x" "-t"], simp) also have "\ = (-t/c > x)" by (simp only: pos_less_divide_eq[OF H] algebra_simps) also have "\ = ((- 1/c)*t > x)" by simp finally show "(c*x + t < 0) == (x < (- 1/c)*t)" by simp qed lemma sum_lt:"((x::'a::pordered_ab_group_add) + t < 0) == (x < - t)" using less_diff_eq[where a= x and b=t and c=0] by simp lemma neg_prod_le:"(c\'a\ordered_field) < 0 \ ((c*x <= 0) == (x >= 0))" proof- assume H: "c < 0" have "c*x <= 0 = (0/c <= x)" by (simp only: neg_divide_le_eq[OF H] algebra_simps) also have "\ = (0 <= x)" by simp finally show "(c*x <= 0) == (x >= 0)" by simp qed lemma pos_prod_le:"(c\'a\ordered_field) > 0 \ ((c*x <= 0) == (x <= 0))" proof- assume H: "c > 0" hence "c*x <= 0 = (0/c >= x)" by (simp only: pos_le_divide_eq[OF H] algebra_simps) also have "\ = (0 >= x)" by simp finally show "(c*x <= 0) == (x <= 0)" by simp qed lemma neg_prod_sum_le: "(c\'a\ordered_field) < 0 \ ((c*x + t <= 0) == (x >= (- 1/c)*t))" proof- assume H: "c < 0" have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp) also have "\ = (-t/c <= x)" by (simp only: neg_divide_le_eq[OF H] algebra_simps) also have "\ = ((- 1/c)*t <= x)" by simp finally show "(c*x + t <= 0) == (x >= (- 1/c)*t)" by simp qed lemma pos_prod_sum_le:"(c\'a\ordered_field) > 0 \ ((c*x + t <= 0) == (x <= (- 1/c)*t))" proof- assume H: "c > 0" have "c*x + t <= 0 = (c*x <= -t)" by (subst le_iff_diff_le_0 [of "c*x" "-t"], simp) also have "\ = (-t/c >= x)" by (simp only: pos_le_divide_eq[OF H] algebra_simps) also have "\ = ((- 1/c)*t >= x)" by simp finally show "(c*x + t <= 0) == (x <= (- 1/c)*t)" by simp qed lemma sum_le:"((x::'a::pordered_ab_group_add) + t <= 0) == (x <= - t)" using le_diff_eq[where a= x and b=t and c=0] by simp lemma nz_prod_eq:"(c\'a\ordered_field) \ 0 \ ((c*x = 0) == (x = 0))" by simp lemma nz_prod_sum_eq: "(c\'a\ordered_field) \ 0 \ ((c*x + t = 0) == (x = (- 1/c)*t))" proof- assume H: "c \ 0" have "c*x + t = 0 = (c*x = -t)" by (subst eq_iff_diff_eq_0 [of "c*x" "-t"], simp) also have "\ = (x = -t/c)" by (simp only: nonzero_eq_divide_eq[OF H] algebra_simps) finally show "(c*x + t = 0) == (x = (- 1/c)*t)" by simp qed lemma sum_eq:"((x::'a::pordered_ab_group_add) + t = 0) == (x = - t)" using eq_diff_eq[where a= x and b=t and c=0] by simp -interpretation class_ordered_field_dense_linear_order!: constr_dense_linear_order +interpretation class_ordered_field_dense_linear_order: constr_dense_linear_order "op <=" "op <" "\ x y. 1/2 * ((x::'a::{ordered_field,recpower,number_ring}) + y)" proof (unfold_locales, dlo, dlo, auto) fix x y::'a assume lt: "x < y" from less_half_sum[OF lt] show "x < (x + y) /2" by simp next fix x y::'a assume lt: "x < y" from gt_half_sum[OF lt] show "(x + y) /2 < y" by simp qed declaration{* let fun earlier [] x y = false | earlier (h::t) x y = if h aconvc y then false else if h aconvc x then true else earlier t x y; fun dest_frac ct = case term_of ct of Const (@{const_name "HOL.divide"},_) $ a $ b=> Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b)) | t => Rat.rat_of_int (snd (HOLogic.dest_number t)) fun mk_frac phi cT x = let val (a, b) = Rat.quotient_of_rat x in if b = 1 then Numeral.mk_cnumber cT a else Thm.capply (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"}) (Numeral.mk_cnumber cT a)) (Numeral.mk_cnumber cT b) end fun whatis x ct = case term_of ct of Const(@{const_name "HOL.plus"}, _)$(Const(@{const_name "HOL.times"},_)$_$y)$_ => if y aconv term_of x then ("c*x+t",[(funpow 2 Thm.dest_arg1) ct, Thm.dest_arg ct]) else ("Nox",[]) | Const(@{const_name "HOL.plus"}, _)$y$_ => if y aconv term_of x then ("x+t",[Thm.dest_arg ct]) else ("Nox",[]) | Const(@{const_name "HOL.times"}, _)$_$y => if y aconv term_of x then ("c*x",[Thm.dest_arg1 ct]) else ("Nox",[]) | t => if t aconv term_of x then ("x",[]) else ("Nox",[]); fun xnormalize_conv ctxt [] ct = reflexive ct | xnormalize_conv ctxt (vs as (x::_)) ct = case term_of ct of Const(@{const_name HOL.less},_)$_$Const(@{const_name "HOL.zero"},_) => (case whatis x (Thm.dest_arg1 ct) of ("c*x+t",[c,t]) => let val cr = dest_frac c val clt = Thm.dest_fun2 ct val cz = Thm.dest_arg ct val neg = cr let val T = ctyp_of_term x val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_lt"} val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th in rth end | ("c*x",[c]) => let val cr = dest_frac c val clt = Thm.dest_fun2 ct val cz = Thm.dest_arg ct val neg = cr reflexive ct) | Const(@{const_name HOL.less_eq},_)$_$Const(@{const_name "HOL.zero"},_) => (case whatis x (Thm.dest_arg1 ct) of ("c*x+t",[c,t]) => let val T = ctyp_of_term x val cr = dest_frac c val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"} val cz = Thm.dest_arg ct val neg = cr let val T = ctyp_of_term x val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_le"} val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th in rth end | ("c*x",[c]) => let val T = ctyp_of_term x val cr = dest_frac c val clt = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "op <"} val cz = Thm.dest_arg ct val neg = cr reflexive ct) | Const("op =",_)$_$Const(@{const_name "HOL.zero"},_) => (case whatis x (Thm.dest_arg1 ct) of ("c*x+t",[c,t]) => let val T = ctyp_of_term x val cr = dest_frac c val ceq = Thm.dest_fun2 ct val cz = Thm.dest_arg ct val cthp = Simplifier.rewrite (local_simpset_of ctxt) (Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz))) val cth = equal_elim (symmetric cthp) TrueI val th = implies_elim (instantiate' [SOME T] (map SOME [c,x,t]) @{thm nz_prod_sum_eq}) cth val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th in rth end | ("x+t",[t]) => let val T = ctyp_of_term x val th = instantiate' [SOME T] [SOME x, SOME t] @{thm "sum_eq"} val rth = Conv.fconv_rule (Conv.arg_conv (Conv.binop_conv (Normalizer.semiring_normalize_ord_conv ctxt (earlier vs)))) th in rth end | ("c*x",[c]) => let val T = ctyp_of_term x val cr = dest_frac c val ceq = Thm.dest_fun2 ct val cz = Thm.dest_arg ct val cthp = Simplifier.rewrite (local_simpset_of ctxt) (Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply ceq c) cz))) val cth = equal_elim (symmetric cthp) TrueI val rth = implies_elim (instantiate' [SOME T] (map SOME [c,x]) @{thm nz_prod_eq}) cth in rth end | _ => reflexive ct); local val less_iff_diff_less_0 = mk_meta_eq @{thm "less_iff_diff_less_0"} val le_iff_diff_le_0 = mk_meta_eq @{thm "le_iff_diff_le_0"} val eq_iff_diff_eq_0 = mk_meta_eq @{thm "eq_iff_diff_eq_0"} in fun field_isolate_conv phi ctxt vs ct = case term_of ct of Const(@{const_name HOL.less},_)$a$b => let val (ca,cb) = Thm.dest_binop ct val T = ctyp_of_term ca val th = instantiate' [SOME T] [SOME ca, SOME cb] less_iff_diff_less_0 val nth = Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) in rth end | Const(@{const_name HOL.less_eq},_)$a$b => let val (ca,cb) = Thm.dest_binop ct val T = ctyp_of_term ca val th = instantiate' [SOME T] [SOME ca, SOME cb] le_iff_diff_le_0 val nth = Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) in rth end | Const("op =",_)$a$b => let val (ca,cb) = Thm.dest_binop ct val T = ctyp_of_term ca val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0 val nth = Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th val rth = transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth)) in rth end | @{term "Not"} $(Const("op =",_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct | _ => reflexive ct end; fun classfield_whatis phi = let fun h x t = case term_of t of Const("op =", _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq else Ferrante_Rackoff_Data.Nox | @{term "Not"}$(Const("op =", _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq else Ferrante_Rackoff_Data.Nox | Const(@{const_name HOL.less},_)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Lt else if term_of x aconv z then Ferrante_Rackoff_Data.Gt else Ferrante_Rackoff_Data.Nox | Const (@{const_name HOL.less_eq},_)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Le else if term_of x aconv z then Ferrante_Rackoff_Data.Ge else Ferrante_Rackoff_Data.Nox | _ => Ferrante_Rackoff_Data.Nox in h end; fun class_field_ss phi = HOL_basic_ss addsimps ([@{thm "linorder_not_less"}, @{thm "linorder_not_le"}]) addsplits [@{thm "abs_split"},@{thm "split_max"}, @{thm "split_min"}] in Ferrante_Rackoff_Data.funs @{thm "class_ordered_field_dense_linear_order.ferrack_axiom"} {isolate_conv = field_isolate_conv, whatis = classfield_whatis, simpset = class_field_ss} end *} lemma upper_bound_finite_set: assumes fS: "finite S" shows "\(a::'a::linorder). \x \ S. f x \ a" proof(induct rule: finite_induct[OF fS]) case 1 thus ?case by simp next case (2 x F) from "2.hyps" obtain a where a:"\x \F. f x \ a" by blast let ?a = "max a (f x)" have m: "a \ ?a" "f x \ ?a" by simp_all {fix y assume y: "y \ insert x F" {assume "y = x" hence "f y \ ?a" using m by simp} moreover {assume yF: "y\ F" from a[rule_format, OF yF] m have "f y \ ?a" by (simp add: max_def)} ultimately have "f y \ ?a" using y by blast} then show ?case by blast qed lemma lower_bound_finite_set: assumes fS: "finite S" shows "\(a::'a::linorder). \x \ S. f x \ a" proof(induct rule: finite_induct[OF fS]) case 1 thus ?case by simp next case (2 x F) from "2.hyps" obtain a where a:"\x \F. f x \ a" by blast let ?a = "min a (f x)" have m: "a \ ?a" "f x \ ?a" by simp_all {fix y assume y: "y \ insert x F" {assume "y = x" hence "f y \ ?a" using m by simp} moreover {assume yF: "y\ F" from a[rule_format, OF yF] m have "f y \ ?a" by (simp add: min_def)} ultimately have "f y \ ?a" using y by blast} then show ?case by blast qed lemma bound_finite_set: assumes f: "finite S" shows "\a. \x \S. (f x:: 'a::linorder) \ a" proof- let ?F = "f ` S" from f have fF: "finite ?F" by simp let ?a = "Max ?F" {assume "S = {}" hence ?thesis by blast} moreover {assume Se: "S \ {}" hence Fe: "?F \ {}" by simp {fix x assume x: "x \ S" hence th0: "f x \ ?F" by simp hence "f x \ ?a" using Max_ge[OF fF th0] ..} hence ?thesis by blast} ultimately show ?thesis by blast qed end diff --git a/src/HOL/Divides.thy b/src/HOL/Divides.thy --- a/src/HOL/Divides.thy +++ b/src/HOL/Divides.thy @@ -1,1156 +1,1156 @@ (* Title: HOL/Divides.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1999 University of Cambridge *) header {* The division operators div and mod *} theory Divides imports Nat Power Product_Type uses "~~/src/Provers/Arith/cancel_div_mod.ML" begin subsection {* Syntactic division operations *} class div = dvd + fixes div :: "'a \ 'a \ 'a" (infixl "div" 70) and mod :: "'a \ 'a \ 'a" (infixl "mod" 70) subsection {* Abstract division in commutative semirings. *} class semiring_div = comm_semiring_1_cancel + div + assumes mod_div_equality: "a div b * b + a mod b = a" and div_by_0 [simp]: "a div 0 = 0" and div_0 [simp]: "0 div a = 0" and div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" begin text {* @{const div} and @{const mod} *} lemma mod_div_equality2: "b * (a div b) + a mod b = a" unfolding mult_commute [of b] by (rule mod_div_equality) lemma mod_div_equality': "a mod b + a div b * b = a" using mod_div_equality [of a b] by (simp only: add_ac) lemma div_mod_equality: "((a div b) * b + a mod b) + c = a + c" by (simp add: mod_div_equality) lemma div_mod_equality2: "(b * (a div b) + a mod b) + c = a + c" by (simp add: mod_div_equality2) lemma mod_by_0 [simp]: "a mod 0 = a" using mod_div_equality [of a zero] by simp lemma mod_0 [simp]: "0 mod a = 0" using mod_div_equality [of zero a] div_0 by simp lemma div_mult_self2 [simp]: assumes "b \ 0" shows "(a + b * c) div b = c + a div b" using assms div_mult_self1 [of b a c] by (simp add: mult_commute) lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" proof (cases "b = 0") case True then show ?thesis by simp next case False have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" by (simp add: mod_div_equality) also from False div_mult_self1 [of b a c] have "\ = (c + a div b) * b + (a + c * b) mod b" by (simp add: algebra_simps) finally have "a = a div b * b + (a + c * b) mod b" by (simp add: add_commute [of a] add_assoc left_distrib) then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" by (simp add: mod_div_equality) then show ?thesis by simp qed lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" by (simp add: mult_commute [of b]) lemma div_mult_self1_is_id [simp]: "b \ 0 \ b * a div b = a" using div_mult_self2 [of b 0 a] by simp lemma div_mult_self2_is_id [simp]: "b \ 0 \ a * b div b = a" using div_mult_self1 [of b 0 a] by simp lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" using mod_mult_self2 [of 0 b a] by simp lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" using mod_mult_self1 [of 0 a b] by simp lemma div_by_1 [simp]: "a div 1 = a" using div_mult_self2_is_id [of 1 a] zero_neq_one by simp lemma mod_by_1 [simp]: "a mod 1 = 0" proof - from mod_div_equality [of a one] div_by_1 have "a + a mod 1 = a" by simp then have "a + a mod 1 = a + 0" by simp then show ?thesis by (rule add_left_imp_eq) qed lemma mod_self [simp]: "a mod a = 0" using mod_mult_self2_is_0 [of 1] by simp lemma div_self [simp]: "a \ 0 \ a div a = 1" using div_mult_self2_is_id [of _ 1] by simp lemma div_add_self1 [simp]: assumes "b \ 0" shows "(b + a) div b = a div b + 1" using assms div_mult_self1 [of b a 1] by (simp add: add_commute) lemma div_add_self2 [simp]: assumes "b \ 0" shows "(a + b) div b = a div b + 1" using assms div_add_self1 [of b a] by (simp add: add_commute) lemma mod_add_self1 [simp]: "(b + a) mod b = a mod b" using mod_mult_self1 [of a 1 b] by (simp add: add_commute) lemma mod_add_self2 [simp]: "(a + b) mod b = a mod b" using mod_mult_self1 [of a 1 b] by simp lemma mod_div_decomp: fixes a b obtains q r where "q = a div b" and "r = a mod b" and "a = q * b + r" proof - from mod_div_equality have "a = a div b * b + a mod b" by simp moreover have "a div b = a div b" .. moreover have "a mod b = a mod b" .. note that ultimately show thesis by blast qed lemma dvd_eq_mod_eq_0 [code unfold]: "a dvd b \ b mod a = 0" proof assume "b mod a = 0" with mod_div_equality [of b a] have "b div a * a = b" by simp then have "b = a * (b div a)" unfolding mult_commute .. then have "\c. b = a * c" .. then show "a dvd b" unfolding dvd_def . next assume "a dvd b" then have "\c. b = a * c" unfolding dvd_def . then obtain c where "b = a * c" .. then have "b mod a = a * c mod a" by simp then have "b mod a = c * a mod a" by (simp add: mult_commute) then show "b mod a = 0" by simp qed lemma mod_div_trivial [simp]: "a mod b div b = 0" proof (cases "b = 0") assume "b = 0" thus ?thesis by simp next assume "b \ 0" hence "a div b + a mod b div b = (a mod b + a div b * b) div b" by (rule div_mult_self1 [symmetric]) also have "\ = a div b" by (simp only: mod_div_equality') also have "\ = a div b + 0" by simp finally show ?thesis by (rule add_left_imp_eq) qed lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" proof - have "a mod b mod b = (a mod b + a div b * b) mod b" by (simp only: mod_mult_self1) also have "\ = a mod b" by (simp only: mod_div_equality') finally show ?thesis . qed lemma dvd_imp_mod_0: "a dvd b \ b mod a = 0" by (rule dvd_eq_mod_eq_0[THEN iffD1]) lemma dvd_div_mult_self: "a dvd b \ (b div a) * a = b" by (subst (2) mod_div_equality [of b a, symmetric]) (simp add:dvd_imp_mod_0) lemma dvd_div_mult: "a dvd b \ (b div a) * c = b * c div a" apply (cases "a = 0") apply simp apply (auto simp: dvd_def mult_assoc) done lemma div_dvd_div[simp]: "a dvd b \ a dvd c \ (b div a dvd c div a) = (b dvd c)" apply (cases "a = 0") apply simp apply (unfold dvd_def) apply auto apply(blast intro:mult_assoc[symmetric]) apply(fastsimp simp add: mult_assoc) done lemma dvd_mod_imp_dvd: "[| k dvd m mod n; k dvd n |] ==> k dvd m" apply (subgoal_tac "k dvd (m div n) *n + m mod n") apply (simp add: mod_div_equality) apply (simp only: dvd_add dvd_mult) done text {* Addition respects modular equivalence. *} lemma mod_add_left_eq: "(a + b) mod c = (a mod c + b) mod c" proof - have "(a + b) mod c = (a div c * c + a mod c + b) mod c" by (simp only: mod_div_equality) also have "\ = (a mod c + b + a div c * c) mod c" by (simp only: add_ac) also have "\ = (a mod c + b) mod c" by (rule mod_mult_self1) finally show ?thesis . qed lemma mod_add_right_eq: "(a + b) mod c = (a + b mod c) mod c" proof - have "(a + b) mod c = (a + (b div c * c + b mod c)) mod c" by (simp only: mod_div_equality) also have "\ = (a + b mod c + b div c * c) mod c" by (simp only: add_ac) also have "\ = (a + b mod c) mod c" by (rule mod_mult_self1) finally show ?thesis . qed lemma mod_add_eq: "(a + b) mod c = (a mod c + b mod c) mod c" by (rule trans [OF mod_add_left_eq mod_add_right_eq]) lemma mod_add_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a + b) mod c = (a' + b') mod c" proof - have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" unfolding assms .. thus ?thesis by (simp only: mod_add_eq [symmetric]) qed text {* Multiplication respects modular equivalence. *} lemma mod_mult_left_eq: "(a * b) mod c = ((a mod c) * b) mod c" proof - have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" by (simp only: mod_div_equality) also have "\ = (a mod c * b + a div c * b * c) mod c" by (simp only: algebra_simps) also have "\ = (a mod c * b) mod c" by (rule mod_mult_self1) finally show ?thesis . qed lemma mod_mult_right_eq: "(a * b) mod c = (a * (b mod c)) mod c" proof - have "(a * b) mod c = (a * (b div c * c + b mod c)) mod c" by (simp only: mod_div_equality) also have "\ = (a * (b mod c) + a * (b div c) * c) mod c" by (simp only: algebra_simps) also have "\ = (a * (b mod c)) mod c" by (rule mod_mult_self1) finally show ?thesis . qed lemma mod_mult_eq: "(a * b) mod c = ((a mod c) * (b mod c)) mod c" by (rule trans [OF mod_mult_left_eq mod_mult_right_eq]) lemma mod_mult_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a * b) mod c = (a' * b') mod c" proof - have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" unfolding assms .. thus ?thesis by (simp only: mod_mult_eq [symmetric]) qed lemma mod_mod_cancel: assumes "c dvd b" shows "a mod b mod c = a mod c" proof - from `c dvd b` obtain k where "b = c * k" by (rule dvdE) have "a mod b mod c = a mod (c * k) mod c" by (simp only: `b = c * k`) also have "\ = (a mod (c * k) + a div (c * k) * k * c) mod c" by (simp only: mod_mult_self1) also have "\ = (a div (c * k) * (c * k) + a mod (c * k)) mod c" by (simp only: add_ac mult_ac) also have "\ = a mod c" by (simp only: mod_div_equality) finally show ?thesis . qed end lemma div_mult_div_if_dvd: "(y::'a::{semiring_div,no_zero_divisors}) dvd x \ z dvd w \ (x div y) * (w div z) = (x * w) div (y * z)" unfolding dvd_def apply clarify apply (case_tac "y = 0") apply simp apply (case_tac "z = 0") apply simp apply (simp add: algebra_simps) apply (subst mult_assoc [symmetric]) apply (simp add: no_zero_divisors) done lemma div_power: "(y::'a::{semiring_div,no_zero_divisors,recpower}) dvd x \ (x div y)^n = x^n div y^n" apply (induct n) apply simp apply(simp add: div_mult_div_if_dvd dvd_power_same) done class ring_div = semiring_div + comm_ring_1 begin text {* Negation respects modular equivalence. *} lemma mod_minus_eq: "(- a) mod b = (- (a mod b)) mod b" proof - have "(- a) mod b = (- (a div b * b + a mod b)) mod b" by (simp only: mod_div_equality) also have "\ = (- (a mod b) + - (a div b) * b) mod b" by (simp only: minus_add_distrib minus_mult_left add_ac) also have "\ = (- (a mod b)) mod b" by (rule mod_mult_self1) finally show ?thesis . qed lemma mod_minus_cong: assumes "a mod b = a' mod b" shows "(- a) mod b = (- a') mod b" proof - have "(- (a mod b)) mod b = (- (a' mod b)) mod b" unfolding assms .. thus ?thesis by (simp only: mod_minus_eq [symmetric]) qed text {* Subtraction respects modular equivalence. *} lemma mod_diff_left_eq: "(a - b) mod c = (a mod c - b) mod c" unfolding diff_minus by (intro mod_add_cong mod_minus_cong) simp_all lemma mod_diff_right_eq: "(a - b) mod c = (a - b mod c) mod c" unfolding diff_minus by (intro mod_add_cong mod_minus_cong) simp_all lemma mod_diff_eq: "(a - b) mod c = (a mod c - b mod c) mod c" unfolding diff_minus by (intro mod_add_cong mod_minus_cong) simp_all lemma mod_diff_cong: assumes "a mod c = a' mod c" assumes "b mod c = b' mod c" shows "(a - b) mod c = (a' - b') mod c" unfolding diff_minus using assms by (intro mod_add_cong mod_minus_cong) lemma dvd_neg_div: "y dvd x \ -x div y = - (x div y)" apply (case_tac "y = 0") apply simp apply (auto simp add: dvd_def) apply (subgoal_tac "-(y * k) = y * - k") apply (erule ssubst) apply (erule div_mult_self1_is_id) apply simp done lemma dvd_div_neg: "y dvd x \ x div -y = - (x div y)" apply (case_tac "y = 0") apply simp apply (auto simp add: dvd_def) apply (subgoal_tac "y * k = -y * -k") apply (erule ssubst) apply (rule div_mult_self1_is_id) apply simp apply simp done end subsection {* Division on @{typ nat} *} text {* We define @{const div} and @{const mod} on @{typ nat} by means of a characteristic relation with two input arguments @{term "m\nat"}, @{term "n\nat"} and two output arguments @{term "q\nat"}(uotient) and @{term "r\nat"}(emainder). *} definition divmod_rel :: "nat \ nat \ nat \ nat \ bool" where "divmod_rel m n q r \ m = q * n + r \ (if n > 0 then 0 \ r \ r < n else q = 0)" text {* @{const divmod_rel} is total: *} lemma divmod_rel_ex: obtains q r where "divmod_rel m n q r" proof (cases "n = 0") case True with that show thesis by (auto simp add: divmod_rel_def) next case False have "\q r. m = q * n + r \ r < n" proof (induct m) case 0 with `n \ 0` have "(0\nat) = 0 * n + 0 \ 0 < n" by simp then show ?case by blast next case (Suc m) then obtain q' r' where m: "m = q' * n + r'" and n: "r' < n" by auto then show ?case proof (cases "Suc r' < n") case True from m n have "Suc m = q' * n + Suc r'" by simp with True show ?thesis by blast next case False then have "n \ Suc r'" by auto moreover from n have "Suc r' \ n" by auto ultimately have "n = Suc r'" by auto with m have "Suc m = Suc q' * n + 0" by simp with `n \ 0` show ?thesis by blast qed qed with that show thesis using `n \ 0` by (auto simp add: divmod_rel_def) qed text {* @{const divmod_rel} is injective: *} lemma divmod_rel_unique_div: assumes "divmod_rel m n q r" and "divmod_rel m n q' r'" shows "q = q'" proof (cases "n = 0") case True with assms show ?thesis by (simp add: divmod_rel_def) next case False have aux: "\q r q' r'. q' * n + r' = q * n + r \ r < n \ q' \ (q\nat)" apply (rule leI) apply (subst less_iff_Suc_add) apply (auto simp add: add_mult_distrib) done from `n \ 0` assms show ?thesis by (auto simp add: divmod_rel_def intro: order_antisym dest: aux sym) qed lemma divmod_rel_unique_mod: assumes "divmod_rel m n q r" and "divmod_rel m n q' r'" shows "r = r'" proof - from assms have "q = q'" by (rule divmod_rel_unique_div) with assms show ?thesis by (simp add: divmod_rel_def) qed text {* We instantiate divisibility on the natural numbers by means of @{const divmod_rel}: *} instantiation nat :: semiring_div begin definition divmod :: "nat \ nat \ nat \ nat" where [code del]: "divmod m n = (THE (q, r). divmod_rel m n q r)" definition div_nat where "m div n = fst (divmod m n)" definition mod_nat where "m mod n = snd (divmod m n)" lemma divmod_div_mod: "divmod m n = (m div n, m mod n)" unfolding div_nat_def mod_nat_def by simp lemma divmod_eq: assumes "divmod_rel m n q r" shows "divmod m n = (q, r)" using assms by (auto simp add: divmod_def dest: divmod_rel_unique_div divmod_rel_unique_mod) lemma div_eq: assumes "divmod_rel m n q r" shows "m div n = q" using assms by (auto dest: divmod_eq simp add: div_nat_def) lemma mod_eq: assumes "divmod_rel m n q r" shows "m mod n = r" using assms by (auto dest: divmod_eq simp add: mod_nat_def) lemma divmod_rel: "divmod_rel m n (m div n) (m mod n)" proof - from divmod_rel_ex obtain q r where rel: "divmod_rel m n q r" . moreover with div_eq mod_eq have "m div n = q" and "m mod n = r" by simp_all ultimately show ?thesis by simp qed lemma divmod_zero: "divmod m 0 = (0, m)" proof - from divmod_rel [of m 0] show ?thesis unfolding divmod_div_mod divmod_rel_def by simp qed lemma divmod_base: assumes "m < n" shows "divmod m n = (0, m)" proof - from divmod_rel [of m n] show ?thesis unfolding divmod_div_mod divmod_rel_def using assms by (cases "m div n = 0") (auto simp add: gr0_conv_Suc [of "m div n"]) qed lemma divmod_step: assumes "0 < n" and "n \ m" shows "divmod m n = (Suc ((m - n) div n), (m - n) mod n)" proof - from divmod_rel have divmod_m_n: "divmod_rel m n (m div n) (m mod n)" . with assms have m_div_n: "m div n \ 1" by (cases "m div n") (auto simp add: divmod_rel_def) from assms divmod_m_n have "divmod_rel (m - n) n (m div n - Suc 0) (m mod n)" by (cases "m div n") (auto simp add: divmod_rel_def) with divmod_eq have "divmod (m - n) n = (m div n - Suc 0, m mod n)" by simp moreover from divmod_div_mod have "divmod (m - n) n = ((m - n) div n, (m - n) mod n)" . ultimately have "m div n = Suc ((m - n) div n)" and "m mod n = (m - n) mod n" using m_div_n by simp_all then show ?thesis using divmod_div_mod by simp qed text {* The ''recursion'' equations for @{const div} and @{const mod} *} lemma div_less [simp]: fixes m n :: nat assumes "m < n" shows "m div n = 0" using assms divmod_base divmod_div_mod by simp lemma le_div_geq: fixes m n :: nat assumes "0 < n" and "n \ m" shows "m div n = Suc ((m - n) div n)" using assms divmod_step divmod_div_mod by simp lemma mod_less [simp]: fixes m n :: nat assumes "m < n" shows "m mod n = m" using assms divmod_base divmod_div_mod by simp lemma le_mod_geq: fixes m n :: nat assumes "n \ m" shows "m mod n = (m - n) mod n" using assms divmod_step divmod_div_mod by (cases "n = 0") simp_all instance proof fix m n :: nat show "m div n * n + m mod n = m" using divmod_rel [of m n] by (simp add: divmod_rel_def) next fix n :: nat show "n div 0 = 0" using divmod_zero divmod_div_mod [of n 0] by simp next fix n :: nat show "0 div n = 0" using divmod_rel [of 0 n] by (cases n) (simp_all add: divmod_rel_def) next fix m n q :: nat assume "n \ 0" then show "(q + m * n) div n = m + q div n" by (induct m) (simp_all add: le_div_geq) qed end text {* Simproc for cancelling @{const div} and @{const mod} *} (*lemmas mod_div_equality_nat = semiring_div_class.times_div_mod_plus_zero_one.mod_div_equality [of "m\nat" n, standard] lemmas mod_div_equality2_nat = mod_div_equality2 [of "n\nat" m, standard*) ML {* structure CancelDivModData = struct val div_name = @{const_name div}; val mod_name = @{const_name mod}; val mk_binop = HOLogic.mk_binop; val mk_sum = Nat_Arith.mk_sum; val dest_sum = Nat_Arith.dest_sum; (*logic*) val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}] val trans = trans val prove_eq_sums = let val simps = @{thm add_0} :: @{thm add_0_right} :: @{thms add_ac} in Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac simps) end; end; structure CancelDivMod = CancelDivModFun(CancelDivModData); val cancel_div_mod_proc = Simplifier.simproc (the_context ()) "cancel_div_mod" ["(m::nat) + n"] (K CancelDivMod.proc); Addsimprocs[cancel_div_mod_proc]; *} text {* code generator setup *} lemma divmod_if [code]: "divmod m n = (if n = 0 \ m < n then (0, m) else let (q, r) = divmod (m - n) n in (Suc q, r))" by (simp add: divmod_zero divmod_base divmod_step) (simp add: divmod_div_mod) code_modulename SML Divides Nat code_modulename OCaml Divides Nat code_modulename Haskell Divides Nat subsubsection {* Quotient *} lemma div_geq: "0 < n \ \ m < n \ m div n = Suc ((m - n) div n)" by (simp add: le_div_geq linorder_not_less) lemma div_if: "0 < n \ m div n = (if m < n then 0 else Suc ((m - n) div n))" by (simp add: div_geq) lemma div_mult_self_is_m [simp]: "0 (m*n) div n = (m::nat)" by simp lemma div_mult_self1_is_m [simp]: "0 (n*m) div n = (m::nat)" by simp subsubsection {* Remainder *} lemma mod_less_divisor [simp]: fixes m n :: nat assumes "n > 0" shows "m mod n < (n::nat)" using assms divmod_rel unfolding divmod_rel_def by auto lemma mod_less_eq_dividend [simp]: fixes m n :: nat shows "m mod n \ m" proof (rule add_leD2) from mod_div_equality have "m div n * n + m mod n = m" . then show "m div n * n + m mod n \ m" by auto qed lemma mod_geq: "\ m < (n\nat) \ m mod n = (m - n) mod n" by (simp add: le_mod_geq linorder_not_less) lemma mod_if: "m mod (n\nat) = (if m < n then m else (m - n) mod n)" by (simp add: le_mod_geq) lemma mod_1 [simp]: "m mod Suc 0 = 0" by (induct m) (simp_all add: mod_geq) lemma mod_mult_distrib: "(m mod n) * (k\nat) = (m * k) mod (n * k)" apply (cases "n = 0", simp) apply (cases "k = 0", simp) apply (induct m rule: nat_less_induct) apply (subst mod_if, simp) apply (simp add: mod_geq diff_mult_distrib) done lemma mod_mult_distrib2: "(k::nat) * (m mod n) = (k*m) mod (k*n)" by (simp add: mult_commute [of k] mod_mult_distrib) (* a simple rearrangement of mod_div_equality: *) lemma mult_div_cancel: "(n::nat) * (m div n) = m - (m mod n)" by (cut_tac a = m and b = n in mod_div_equality2, arith) lemma mod_le_divisor[simp]: "0 < n \ m mod n \ (n::nat)" apply (drule mod_less_divisor [where m = m]) apply simp done subsubsection {* Quotient and Remainder *} lemma divmod_rel_mult1_eq: "[| divmod_rel b c q r; c > 0 |] ==> divmod_rel (a*b) c (a*q + a*r div c) (a*r mod c)" by (auto simp add: split_ifs divmod_rel_def algebra_simps) lemma div_mult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::nat)" apply (cases "c = 0", simp) apply (blast intro: divmod_rel [THEN divmod_rel_mult1_eq, THEN div_eq]) done lemma divmod_rel_add1_eq: "[| divmod_rel a c aq ar; divmod_rel b c bq br; c > 0 |] ==> divmod_rel (a + b) c (aq + bq + (ar+br) div c) ((ar + br) mod c)" by (auto simp add: split_ifs divmod_rel_def algebra_simps) (*NOT suitable for rewriting: the RHS has an instance of the LHS*) lemma div_add1_eq: "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)" apply (cases "c = 0", simp) apply (blast intro: divmod_rel_add1_eq [THEN div_eq] divmod_rel) done lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c" apply (cut_tac m = q and n = c in mod_less_divisor) apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto) apply (erule_tac P = "%x. ?lhs < ?rhs x" in ssubst) apply (simp add: add_mult_distrib2) done lemma divmod_rel_mult2_eq: "[| divmod_rel a b q r; 0 < b; 0 < c |] ==> divmod_rel a (b*c) (q div c) (b*(q mod c) + r)" by (auto simp add: mult_ac divmod_rel_def add_mult_distrib2 [symmetric] mod_lemma) lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)" apply (cases "b = 0", simp) apply (cases "c = 0", simp) apply (force simp add: divmod_rel [THEN divmod_rel_mult2_eq, THEN div_eq]) done lemma mod_mult2_eq: "a mod (b*c) = b*(a div b mod c) + a mod (b::nat)" apply (cases "b = 0", simp) apply (cases "c = 0", simp) apply (auto simp add: mult_commute divmod_rel [THEN divmod_rel_mult2_eq, THEN mod_eq]) done subsubsection{*Cancellation of Common Factors in Division*} lemma div_mult_mult_lemma: "[| (0::nat) < b; 0 < c |] ==> (c*a) div (c*b) = a div b" by (auto simp add: div_mult2_eq) lemma div_mult_mult1 [simp]: "(0::nat) < c ==> (c*a) div (c*b) = a div b" apply (cases "b = 0") apply (auto simp add: linorder_neq_iff [of b] div_mult_mult_lemma) done lemma div_mult_mult2 [simp]: "(0::nat) < c ==> (a*c) div (b*c) = a div b" apply (drule div_mult_mult1) apply (auto simp add: mult_commute) done subsubsection{*Further Facts about Quotient and Remainder*} lemma div_1 [simp]: "m div Suc 0 = m" by (induct m) (simp_all add: div_geq) (* Monotonicity of div in first argument *) lemma div_le_mono [rule_format (no_asm)]: "\m::nat. m \ n --> (m div k) \ (n div k)" apply (case_tac "k=0", simp) apply (induct "n" rule: nat_less_induct, clarify) apply (case_tac "n= k *) apply (case_tac "m=k *) apply (simp add: div_geq diff_le_mono) done (* Antimonotonicity of div in second argument *) lemma div_le_mono2: "!!m::nat. [| 0n |] ==> (k div n) \ (k div m)" apply (subgoal_tac "0 (k-m) div n") prefer 2 apply (blast intro: div_le_mono diff_le_mono2) apply (rule le_trans, simp) apply (simp) done lemma div_le_dividend [simp]: "m div n \ (m::nat)" apply (case_tac "n=0", simp) apply (subgoal_tac "m div n \ m div 1", simp) apply (rule div_le_mono2) apply (simp_all (no_asm_simp)) done (* Similar for "less than" *) lemma div_less_dividend [rule_format]: "!!n::nat. 1 0 < m --> m div n < m" apply (induct_tac m rule: nat_less_induct) apply (rename_tac "m") apply (case_tac "m Suc(na) *) apply (simp add: linorder_not_less le_Suc_eq mod_geq) apply (auto simp add: Suc_diff_le le_mod_geq) done subsubsection {* The Divides Relation *} lemma dvd_1_left [iff]: "Suc 0 dvd k" unfolding dvd_def by simp lemma dvd_1_iff_1 [simp]: "(m dvd Suc 0) = (m = Suc 0)" by (simp add: dvd_def) lemma nat_dvd_1_iff_1 [simp]: "m dvd (1::nat) \ m = 1" by (simp add: dvd_def) lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m = (n::nat)" unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult_assoc mult_eq_1_iff) text {* @{term "op dvd"} is a partial order *} -interpretation dvd!: order "op dvd" "\n m \ nat. n dvd m \ \ m dvd n" +interpretation dvd: order "op dvd" "\n m \ nat. n dvd m \ \ m dvd n" proof qed (auto intro: dvd_refl dvd_trans dvd_anti_sym) lemma nat_dvd_diff[simp]: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)" unfolding dvd_def by (blast intro: diff_mult_distrib2 [symmetric]) lemma dvd_diffD: "[| k dvd m-n; k dvd n; n\m |] ==> k dvd (m::nat)" apply (erule linorder_not_less [THEN iffD2, THEN add_diff_inverse, THEN subst]) apply (blast intro: dvd_add) done lemma dvd_diffD1: "[| k dvd m-n; k dvd m; n\m |] ==> k dvd (n::nat)" by (drule_tac m = m in nat_dvd_diff, auto) lemma dvd_reduce: "(k dvd n + k) = (k dvd (n::nat))" apply (rule iffI) apply (erule_tac [2] dvd_add) apply (rule_tac [2] dvd_refl) apply (subgoal_tac "n = (n+k) -k") prefer 2 apply simp apply (erule ssubst) apply (erule nat_dvd_diff) apply (rule dvd_refl) done lemma dvd_mod: "!!n::nat. [| f dvd m; f dvd n |] ==> f dvd m mod n" unfolding dvd_def apply (case_tac "n = 0", auto) apply (blast intro: mod_mult_distrib2 [symmetric]) done lemma dvd_mod_iff: "k dvd n ==> ((k::nat) dvd m mod n) = (k dvd m)" by (blast intro: dvd_mod_imp_dvd dvd_mod) lemma dvd_mult_cancel: "!!k::nat. [| k*m dvd k*n; 0 m dvd n" unfolding dvd_def apply (erule exE) apply (simp add: mult_ac) done lemma dvd_mult_cancel1: "0 (m*n dvd m) = (n = (1::nat))" apply auto apply (subgoal_tac "m*n dvd m*1") apply (drule dvd_mult_cancel, auto) done lemma dvd_mult_cancel2: "0 (n*m dvd m) = (n = (1::nat))" apply (subst mult_commute) apply (erule dvd_mult_cancel1) done lemma dvd_imp_le: "[| k dvd n; 0 < n |] ==> k \ (n::nat)" apply (unfold dvd_def, clarify) apply (simp_all (no_asm_use) add: zero_less_mult_iff) apply (erule conjE) apply (rule le_trans) apply (rule_tac [2] le_refl [THEN mult_le_mono]) apply (erule_tac [2] Suc_leI, simp) done lemma dvd_mult_div_cancel: "n dvd m ==> n * (m div n) = (m::nat)" apply (subgoal_tac "m mod n = 0") apply (simp add: mult_div_cancel) apply (simp only: dvd_eq_mod_eq_0) done lemma nat_zero_less_power_iff [simp]: "(x^n > 0) = (x > (0::nat) | n=0)" by (induct n) auto lemma power_dvd_imp_le: "[|i^m dvd i^n; (1::nat) < i|] ==> m \ n" apply (rule power_le_imp_le_exp, assumption) apply (erule dvd_imp_le, simp) done lemma mod_eq_0_iff: "(m mod d = 0) = (\q::nat. m = d*q)" by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) lemmas mod_eq_0D [dest!] = mod_eq_0_iff [THEN iffD1] (*Loses information, namely we also have r \q::nat. m = r + q*d" apply (cut_tac a = m in mod_div_equality) apply (simp only: add_ac) apply (blast intro: sym) done lemma split_div: "P(n div k :: nat) = ((k = 0 \ P 0) \ (k \ 0 \ (!i. !j P i)))" (is "?P = ?Q" is "_ = (_ \ (_ \ ?R))") proof assume P: ?P show ?Q proof (cases) assume "k = 0" with P show ?Q by simp next assume not0: "k \ 0" thus ?Q proof (simp, intro allI impI) fix i j assume n: "n = k*i + j" and j: "j < k" show "P i" proof (cases) assume "i = 0" with n j P show "P i" by simp next assume "i \ 0" with not0 n j P show "P i" by(simp add:add_ac) qed qed qed next assume Q: ?Q show ?P proof (cases) assume "k = 0" with Q show ?P by simp next assume not0: "k \ 0" with Q have R: ?R by simp from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] show ?P by simp qed qed lemma split_div_lemma: assumes "0 < n" shows "n * q \ m \ m < n * Suc q \ q = ((m\nat) div n)" (is "?lhs \ ?rhs") proof assume ?rhs with mult_div_cancel have nq: "n * q = m - (m mod n)" by simp then have A: "n * q \ m" by simp have "n - (m mod n) > 0" using mod_less_divisor assms by auto then have "m < m + (n - (m mod n))" by simp then have "m < n + (m - (m mod n))" by simp with nq have "m < n + n * q" by simp then have B: "m < n * Suc q" by simp from A B show ?lhs .. next assume P: ?lhs then have "divmod_rel m n q (m - n * q)" unfolding divmod_rel_def by (auto simp add: mult_ac) then show ?rhs using divmod_rel by (rule divmod_rel_unique_div) qed theorem split_div': "P ((m::nat) div n) = ((n = 0 \ P 0) \ (\q. (n * q \ m \ m < n * (Suc q)) \ P q))" apply (case_tac "0 < n") apply (simp only: add: split_div_lemma) apply simp_all done lemma split_mod: "P(n mod k :: nat) = ((k = 0 \ P n) \ (k \ 0 \ (!i. !j P j)))" (is "?P = ?Q" is "_ = (_ \ (_ \ ?R))") proof assume P: ?P show ?Q proof (cases) assume "k = 0" with P show ?Q by simp next assume not0: "k \ 0" thus ?Q proof (simp, intro allI impI) fix i j assume "n = k*i + j" "j < k" thus "P j" using not0 P by(simp add:add_ac mult_ac) qed qed next assume Q: ?Q show ?P proof (cases) assume "k = 0" with Q show ?P by simp next assume not0: "k \ 0" with Q have R: ?R by simp from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] show ?P by simp qed qed theorem mod_div_equality': "(m::nat) mod n = m - (m div n) * n" apply (rule_tac P="%x. m mod n = x - (m div n) * n" in subst [OF mod_div_equality [of _ n]]) apply arith done lemma div_mod_equality': fixes m n :: nat shows "m div n * n = m - m mod n" proof - have "m mod n \ m mod n" .. from div_mod_equality have "m div n * n + m mod n - m mod n = m - m mod n" by simp with diff_add_assoc [OF `m mod n \ m mod n`, of "m div n * n"] have "m div n * n + (m mod n - m mod n) = m - m mod n" by simp then show ?thesis by simp qed subsubsection {*An ``induction'' law for modulus arithmetic.*} lemma mod_induct_0: assumes step: "\i P ((Suc i) mod p)" and base: "P i" and i: "i(P 0)" from i have p: "0k. 0 \ P (p-k)" (is "\k. ?A k") proof fix k show "?A k" proof (induct k) show "?A 0" by simp -- "by contradiction" next fix n assume ih: "?A n" show "?A (Suc n)" proof (clarsimp) assume y: "P (p - Suc n)" have n: "Suc n < p" proof (rule ccontr) assume "\(Suc n < p)" hence "p - Suc n = 0" by simp with y contra show "False" by simp qed hence n2: "Suc (p - Suc n) = p-n" by arith from p have "p - Suc n < p" by arith with y step have z: "P ((Suc (p - Suc n)) mod p)" by blast show "False" proof (cases "n=0") case True with z n2 contra show ?thesis by simp next case False with p have "p-n < p" by arith with z n2 False ih show ?thesis by simp qed qed qed qed moreover from i obtain k where "0 i+k=p" by (blast dest: less_imp_add_positive) hence "0 i=p-k" by auto moreover note base ultimately show "False" by blast qed lemma mod_induct: assumes step: "\i P ((Suc i) mod p)" and base: "P i" and i: "ij P j" (is "?A j") proof (induct j) from step base i show "?A 0" by (auto elim: mod_induct_0) next fix k assume ih: "?A k" show "?A (Suc k)" proof assume suc: "Suc k < p" hence k: "k m < n \ \ n dvd m" by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) end diff --git a/src/HOL/Finite_Set.thy b/src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy +++ b/src/HOL/Finite_Set.thy @@ -1,3169 +1,3169 @@ (* Title: HOL/Finite_Set.thy Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel with contributions by Jeremy Avigad *) header {* Finite sets *} theory Finite_Set imports Nat Product_Type Power begin subsection {* Definition and basic properties *} inductive finite :: "'a set => bool" where emptyI [simp, intro!]: "finite {}" | insertI [simp, intro!]: "finite A ==> finite (insert a A)" lemma ex_new_if_finite: -- "does not depend on def of finite at all" assumes "\ finite (UNIV :: 'a set)" and "finite A" shows "\a::'a. a \ A" proof - from assms have "A \ UNIV" by blast thus ?thesis by blast qed lemma finite_induct [case_names empty insert, induct set: finite]: "finite F ==> P {} ==> (!!x F. finite F ==> x \ F ==> P F ==> P (insert x F)) ==> P F" -- {* Discharging @{text "x \ F"} entails extra work. *} proof - assume "P {}" and insert: "!!x F. finite F ==> x \ F ==> P F ==> P (insert x F)" assume "finite F" thus "P F" proof induct show "P {}" by fact fix x F assume F: "finite F" and P: "P F" show "P (insert x F)" proof cases assume "x \ F" hence "insert x F = F" by (rule insert_absorb) with P show ?thesis by (simp only:) next assume "x \ F" from F this P show ?thesis by (rule insert) qed qed qed lemma finite_ne_induct[case_names singleton insert, consumes 2]: assumes fin: "finite F" shows "F \ {} \ \ \x. P{x}; \x F. \ finite F; F \ {}; x \ F; P F \ \ P (insert x F) \ \ P F" using fin proof induct case empty thus ?case by simp next case (insert x F) show ?case proof cases assume "F = {}" thus ?thesis using `P {x}` by simp next assume "F \ {}" thus ?thesis using insert by blast qed qed lemma finite_subset_induct [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "!!a F. finite F ==> a \ A ==> a \ F ==> P F ==> P (insert a F)" shows "P F" proof - from `finite F` and `F \ A` show ?thesis proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A ==> P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact qed qed qed text{* A finite choice principle. Does not need the SOME choice operator. *} lemma finite_set_choice: "finite A \ ALL x:A. (EX y. P x y) \ EX f. ALL x:A. P x (f x)" proof (induct set: finite) case empty thus ?case by simp next case (insert a A) then obtain f b where f: "ALL x:A. P x (f x)" and ab: "P a b" by auto show ?case (is "EX f. ?P f") proof show "?P(%x. if x = a then b else f x)" using f ab by auto qed qed text{* Finite sets are the images of initial segments of natural numbers: *} lemma finite_imp_nat_seg_image_inj_on: assumes fin: "finite A" shows "\ (n::nat) f. A = f ` {i. if. {} = f ` {i::nat. i < 0} & inj_on f {i. i<0}" by simp qed next case (insert a A) have notinA: "a \ A" by fact from insert.hyps obtain n f where "A = f ` {i::nat. i < n}" "inj_on f {i. i < n}" by blast hence "insert a A = f(n:=a) ` {i. i < Suc n}" "inj_on (f(n:=a)) {i. i < Suc n}" using notinA by (auto simp add: image_def Ball_def inj_on_def less_Suc_eq) thus ?case by blast qed lemma nat_seg_image_imp_finite: "!!f A. A = f ` {i::nat. i finite A" proof (induct n) case 0 thus ?case by simp next case (Suc n) let ?B = "f ` {i. i < n}" have finB: "finite ?B" by(rule Suc.hyps[OF refl]) show ?case proof cases assume "\k(\ k (n::nat) f. A = f ` {i::nat. i finite G ==> finite (F Un G)" by (induct set: finite) simp_all lemma finite_subset: "A \ B ==> finite B ==> finite A" -- {* Every subset of a finite set is finite. *} proof - assume "finite B" thus "!!A. A \ B ==> finite A" proof induct case empty thus ?case by simp next case (insert x F A) have A: "A \ insert x F" and r: "A - {x} \ F ==> finite (A - {x})" by fact+ show "finite A" proof cases assume x: "x \ A" with A have "A - {x} \ F" by (simp add: subset_insert_iff) with r have "finite (A - {x})" . hence "finite (insert x (A - {x}))" .. also have "insert x (A - {x}) = A" using x by (rule insert_Diff) finally show ?thesis . next show "A \ F ==> ?thesis" by fact assume "x \ A" with A show "A \ F" by (simp add: subset_insert_iff) qed qed qed lemma finite_Un [iff]: "finite (F Un G) = (finite F & finite G)" by (blast intro: finite_subset [of _ "X Un Y", standard] finite_UnI) lemma finite_Collect_disjI[simp]: "finite{x. P x | Q x} = (finite{x. P x} & finite{x. Q x})" by(simp add:Collect_disj_eq) lemma finite_Int [simp, intro]: "finite F | finite G ==> finite (F Int G)" -- {* The converse obviously fails. *} by (blast intro: finite_subset) lemma finite_Collect_conjI [simp, intro]: "finite{x. P x} | finite{x. Q x} ==> finite{x. P x & Q x}" -- {* The converse obviously fails. *} by(simp add:Collect_conj_eq) lemma finite_Collect_le_nat[iff]: "finite{n::nat. n<=k}" by(simp add: le_eq_less_or_eq) lemma finite_insert [simp]: "finite (insert a A) = finite A" apply (subst insert_is_Un) apply (simp only: finite_Un, blast) done lemma finite_Union[simp, intro]: "\ finite A; !!M. M \ A \ finite M \ \ finite(\A)" by (induct rule:finite_induct) simp_all lemma finite_empty_induct: assumes "finite A" and "P A" and "!!a A. finite A ==> a:A ==> P A ==> P (A - {a})" shows "P {}" proof - have "P (A - A)" proof - { fix c b :: "'a set" assume c: "finite c" and b: "finite b" and P1: "P b" and P2: "!!x y. finite y ==> x \ y ==> P y ==> P (y - {x})" have "c \ b ==> P (b - c)" using c proof induct case empty from P1 show ?case by simp next case (insert x F) have "P (b - F - {x})" proof (rule P2) from _ b show "finite (b - F)" by (rule finite_subset) blast from insert show "x \ b - F" by simp from insert show "P (b - F)" by simp qed also have "b - F - {x} = b - insert x F" by (rule Diff_insert [symmetric]) finally show ?case . qed } then show ?thesis by this (simp_all add: assms) qed then show ?thesis by simp qed lemma finite_Diff [simp]: "finite A ==> finite (A - B)" by (rule Diff_subset [THEN finite_subset]) lemma finite_Diff2 [simp]: assumes "finite B" shows "finite (A - B) = finite A" proof - have "finite A \ finite((A-B) Un (A Int B))" by(simp add: Un_Diff_Int) also have "\ \ finite(A-B)" using `finite B` by(simp) finally show ?thesis .. qed lemma finite_compl[simp]: "finite(A::'a set) \ finite(-A) = finite(UNIV::'a set)" by(simp add:Compl_eq_Diff_UNIV) lemma finite_Collect_not[simp]: "finite{x::'a. P x} \ finite{x. ~P x} = finite(UNIV::'a set)" by(simp add:Collect_neg_eq) lemma finite_Diff_insert [iff]: "finite (A - insert a B) = finite (A - B)" apply (subst Diff_insert) apply (case_tac "a : A - B") apply (rule finite_insert [symmetric, THEN trans]) apply (subst insert_Diff, simp_all) done text {* Image and Inverse Image over Finite Sets *} lemma finite_imageI[simp]: "finite F ==> finite (h ` F)" -- {* The image of a finite set is finite. *} by (induct set: finite) simp_all lemma finite_surj: "finite A ==> B <= f ` A ==> finite B" apply (frule finite_imageI) apply (erule finite_subset, assumption) done lemma finite_range_imageI: "finite (range g) ==> finite (range (%x. f (g x)))" apply (drule finite_imageI, simp add: range_composition) done lemma finite_imageD: "finite (f`A) ==> inj_on f A ==> finite A" proof - have aux: "!!A. finite (A - {}) = finite A" by simp fix B :: "'a set" assume "finite B" thus "!!A. f`A = B ==> inj_on f A ==> finite A" apply induct apply simp apply (subgoal_tac "EX y:A. f y = x & F = f ` (A - {y})") apply clarify apply (simp (no_asm_use) add: inj_on_def) apply (blast dest!: aux [THEN iffD1], atomize) apply (erule_tac V = "ALL A. ?PP (A)" in thin_rl) apply (frule subsetD [OF equalityD2 insertI1], clarify) apply (rule_tac x = xa in bexI) apply (simp_all add: inj_on_image_set_diff) done qed (rule refl) lemma inj_vimage_singleton: "inj f ==> f-`{a} \ {THE x. f x = a}" -- {* The inverse image of a singleton under an injective function is included in a singleton. *} apply (auto simp add: inj_on_def) apply (blast intro: the_equality [symmetric]) done lemma finite_vimageI: "[|finite F; inj h|] ==> finite (h -` F)" -- {* The inverse image of a finite set under an injective function is finite. *} apply (induct set: finite) apply simp_all apply (subst vimage_insert) apply (simp add: finite_Un finite_subset [OF inj_vimage_singleton]) done text {* The finite UNION of finite sets *} lemma finite_UN_I: "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (UN a:A. B a)" by (induct set: finite) simp_all text {* Strengthen RHS to @{prop "((ALL x:A. finite (B x)) & finite {x. x:A & B x \ {}})"}? We'd need to prove @{prop "finite C ==> ALL A B. (UNION A B) <= C --> finite {x. x:A & B x \ {}}"} by induction. *} lemma finite_UN [simp]: "finite A ==> finite (UNION A B) = (ALL x:A. finite (B x))" by (blast intro: finite_UN_I finite_subset) lemma finite_Collect_bex[simp]: "finite A \ finite{x. EX y:A. Q x y} = (ALL y:A. finite{x. Q x y})" apply(subgoal_tac "{x. EX y:A. Q x y} = UNION A (%y. {x. Q x y})") apply auto done lemma finite_Collect_bounded_ex[simp]: "finite{y. P y} \ finite{x. EX y. P y & Q x y} = (ALL y. P y \ finite{x. Q x y})" apply(subgoal_tac "{x. EX y. P y & Q x y} = UNION {y. P y} (%y. {x. Q x y})") apply auto done lemma finite_Plus: "[| finite A; finite B |] ==> finite (A <+> B)" by (simp add: Plus_def) text {* Sigma of finite sets *} lemma finite_SigmaI [simp]: "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (SIGMA a:A. B a)" by (unfold Sigma_def) (blast intro!: finite_UN_I) lemma finite_cartesian_product: "[| finite A; finite B |] ==> finite (A <*> B)" by (rule finite_SigmaI) lemma finite_Prod_UNIV: "finite (UNIV::'a set) ==> finite (UNIV::'b set) ==> finite (UNIV::('a * 'b) set)" apply (subgoal_tac "(UNIV:: ('a * 'b) set) = Sigma UNIV (%x. UNIV)") apply (erule ssubst) apply (erule finite_SigmaI, auto) done lemma finite_cartesian_productD1: "[| finite (A <*> B); B \ {} |] ==> finite A" apply (auto simp add: finite_conv_nat_seg_image) apply (drule_tac x=n in spec) apply (drule_tac x="fst o f" in spec) apply (auto simp add: o_def) prefer 2 apply (force dest!: equalityD2) apply (drule equalityD1) apply (rename_tac y x) apply (subgoal_tac "\k. k B); A \ {} |] ==> finite B" apply (auto simp add: finite_conv_nat_seg_image) apply (drule_tac x=n in spec) apply (drule_tac x="snd o f" in spec) apply (auto simp add: o_def) prefer 2 apply (force dest!: equalityD2) apply (drule equalityD1) apply (rename_tac x y) apply (subgoal_tac "\k. k finite{B. B \ A}" by(simp add: Pow_def[symmetric]) lemma finite_UnionD: "finite(\A) \ finite A" by(blast intro: finite_subset[OF subset_Pow_Union]) subsection {* Class @{text finite} *} setup {* Sign.add_path "finite" *} -- {*FIXME: name tweaking*} class finite = assumes finite_UNIV: "finite (UNIV \ 'a set)" setup {* Sign.parent_path *} hide const finite context finite begin lemma finite [simp]: "finite (A \ 'a set)" by (rule subset_UNIV finite_UNIV finite_subset)+ end lemma UNIV_unit [noatp]: "UNIV = {()}" by auto instance unit :: finite by default (simp add: UNIV_unit) lemma UNIV_bool [noatp]: "UNIV = {False, True}" by auto instance bool :: finite by default (simp add: UNIV_bool) instance * :: (finite, finite) finite by default (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite) lemma inj_graph: "inj (%f. {(x, y). y = f x})" by (rule inj_onI, auto simp add: expand_set_eq expand_fun_eq) instance "fun" :: (finite, finite) finite proof show "finite (UNIV :: ('a => 'b) set)" proof (rule finite_imageD) let ?graph = "%f::'a => 'b. {(x, y). y = f x}" have "range ?graph \ Pow UNIV" by simp moreover have "finite (Pow (UNIV :: ('a * 'b) set))" by (simp only: finite_Pow_iff finite) ultimately show "finite (range ?graph)" by (rule finite_subset) show "inj ?graph" by (rule inj_graph) qed qed instance "+" :: (finite, finite) finite by default (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite) subsection {* A fold functional for finite sets *} text {* The intended behaviour is @{text "fold f z {x\<^isub>1, ..., x\<^isub>n} = f x\<^isub>1 (\ (f x\<^isub>n z)\)"} if @{text f} is ``left-commutative'': *} locale fun_left_comm = fixes f :: "'a \ 'b \ 'b" assumes fun_left_comm: "f x (f y z) = f y (f x z)" begin text{* On a functional level it looks much nicer: *} lemma fun_comp_comm: "f x \ f y = f y \ f x" by (simp add: fun_left_comm expand_fun_eq) end inductive fold_graph :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b \ bool" for f :: "'a \ 'b \ 'b" and z :: 'b where emptyI [intro]: "fold_graph f z {} z" | insertI [intro]: "x \ A \ fold_graph f z A y \ fold_graph f z (insert x A) (f x y)" inductive_cases empty_fold_graphE [elim!]: "fold_graph f z {} x" definition fold :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" where [code del]: "fold f z A = (THE y. fold_graph f z A y)" text{*A tempting alternative for the definiens is @{term "if finite A then THE y. fold_graph f z A y else e"}. It allows the removal of finiteness assumptions from the theorems @{text fold_comm}, @{text fold_reindex} and @{text fold_distrib}. The proofs become ugly. It is not worth the effort. (???) *} lemma Diff1_fold_graph: "fold_graph f z (A - {x}) y \ x \ A \ fold_graph f z A (f x y)" by (erule insert_Diff [THEN subst], rule fold_graph.intros, auto) lemma fold_graph_imp_finite: "fold_graph f z A x \ finite A" by (induct set: fold_graph) auto lemma finite_imp_fold_graph: "finite A \ \x. fold_graph f z A x" by (induct set: finite) auto subsubsection{*From @{const fold_graph} to @{term fold}*} lemma image_less_Suc: "h ` {i. i < Suc m} = insert (h m) (h ` {i. i < m})" by (auto simp add: less_Suc_eq) lemma insert_image_inj_on_eq: "[|insert (h m) A = h ` {i. i < Suc m}; h m \ A; inj_on h {i. i < Suc m}|] ==> A = h ` {i. i < m}" apply (auto simp add: image_less_Suc inj_on_def) apply (blast intro: less_trans) done lemma insert_inj_onE: assumes aA: "insert a A = h`{i::nat. i A" and inj_on: "inj_on h {i::nat. ihm m. inj_on hm {i::nat. i A" by (simp add: swap_def hkeq anot) show "insert (?hm m) A = ?hm ` {i. i < Suc m}" using aA hkeq nSuc klessn by (auto simp add: swap_def image_less_Suc fun_upd_image less_Suc_eq inj_on_image_set_diff [OF inj_on]) qed qed qed context fun_left_comm begin lemma fold_graph_determ_aux: "A = h`{i::nat. i inj_on h {i. i fold_graph f z A x \ fold_graph f z A x' \ x' = x" proof (induct n arbitrary: A x x' h rule: less_induct) case (less n) have IH: "\m h A x x'. m < n \ A = h ` {i. i inj_on h {i. i fold_graph f z A x \ fold_graph f z A x' \ x' = x" by fact have Afoldx: "fold_graph f z A x" and Afoldx': "fold_graph f z A x'" and A: "A = h`{i. i B" and Bu: "fold_graph f z B u" show "x'=x" proof (rule fold_graph.cases [OF Afoldx']) assume "A = {}" and "x' = z" with AbB show "x' = x" by blast next fix C c v assume AcC: "A = insert c C" and x': "x' = f c v" and notinC: "c \ C" and Cv: "fold_graph f z C v" from A AbB have Beq: "insert b B = h`{i. i c" let ?D = "B - {c}" have B: "B = insert c ?D" and C: "C = insert b ?D" using AbB AcC notinB notinC diff by(blast elim!:equalityE)+ have "finite A" by(rule fold_graph_imp_finite [OF Afoldx]) with AbB have "finite ?D" by simp then obtain d where Dfoldd: "fold_graph f z ?D d" using finite_imp_fold_graph by iprover moreover have cinB: "c \ B" using B by auto ultimately have "fold_graph f z B (f c d)" by(rule Diff1_fold_graph) hence "f c d = u" by (rule IH [OF lessB Beq inj_onB Bu]) moreover have "f b d = v" proof (rule IH[OF lessC Ceq inj_onC Cv]) show "fold_graph f z C (f b d)" using C notinB Dfoldd by fastsimp qed ultimately show ?thesis using fun_left_comm [of c b] x x' by (auto simp add: o_def) qed qed qed qed lemma fold_graph_determ: "fold_graph f z A x \ fold_graph f z A y \ y = x" apply (frule fold_graph_imp_finite [THEN finite_imp_nat_seg_image_inj_on]) apply (blast intro: fold_graph_determ_aux [rule_format]) done lemma fold_equality: "fold_graph f z A y \ fold f z A = y" by (unfold fold_def) (blast intro: fold_graph_determ) text{* The base case for @{text fold}: *} lemma (in -) fold_empty [simp]: "fold f z {} = z" by (unfold fold_def) blast text{* The various recursion equations for @{const fold}: *} lemma fold_insert_aux: "x \ A \ fold_graph f z (insert x A) v \ (\y. fold_graph f z A y \ v = f x y)" apply auto apply (rule_tac A1 = A and f1 = f in finite_imp_fold_graph [THEN exE]) apply (fastsimp dest: fold_graph_imp_finite) apply (blast intro: fold_graph_determ) done lemma fold_insert [simp]: "finite A ==> x \ A ==> fold f z (insert x A) = f x (fold f z A)" apply (simp add: fold_def fold_insert_aux) apply (rule the_equality) apply (auto intro: finite_imp_fold_graph cong add: conj_cong simp add: fold_def[symmetric] fold_equality) done lemma fold_fun_comm: "finite A \ f x (fold f z A) = fold f (f x z) A" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert y A) then show ?case by (simp add: fun_left_comm[of x]) qed lemma fold_insert2: "finite A \ x \ A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_insert fold_fun_comm) lemma fold_rec: assumes "finite A" and "x \ A" shows "fold f z A = f x (fold f z (A - {x}))" proof - have A: "A = insert x (A - {x})" using `x \ A` by blast then have "fold f z A = fold f z (insert x (A - {x}))" by simp also have "\ = f x (fold f z (A - {x}))" by (rule fold_insert) (simp add: `finite A`)+ finally show ?thesis . qed lemma fold_insert_remove: assumes "finite A" shows "fold f z (insert x A) = f x (fold f z (A - {x}))" proof - from `finite A` have "finite (insert x A)" by auto moreover have "x \ insert x A" by auto ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))" by (rule fold_rec) then show ?thesis by simp qed end text{* A simplified version for idempotent functions: *} locale fun_left_comm_idem = fun_left_comm + assumes fun_left_idem: "f x (f x z) = f x z" begin text{* The nice version: *} lemma fun_comp_idem : "f x o f x = f x" by (simp add: fun_left_idem expand_fun_eq) lemma fold_insert_idem: assumes fin: "finite A" shows "fold f z (insert x A) = f x (fold f z A)" proof cases assume "x \ A" then obtain B where "A = insert x B" and "x \ B" by (rule set_insert) then show ?thesis using assms by (simp add:fun_left_idem) next assume "x \ A" then show ?thesis using assms by simp qed declare fold_insert[simp del] fold_insert_idem[simp] lemma fold_insert_idem2: "finite A \ fold f z (insert x A) = fold f (f x z) A" by(simp add:fold_fun_comm) end subsubsection{* The derived combinator @{text fold_image} *} definition fold_image :: "('b \ 'b \ 'b) \ ('a \ 'b) \ 'b \ 'a set \ 'b" where "fold_image f g = fold (%x y. f (g x) y)" lemma fold_image_empty[simp]: "fold_image f g z {} = z" by(simp add:fold_image_def) context ab_semigroup_mult begin lemma fold_image_insert[simp]: assumes "finite A" and "a \ A" shows "fold_image times g z (insert a A) = g a * (fold_image times g z A)" proof - interpret I: fun_left_comm "%x y. (g x) * y" by unfold_locales (simp add: mult_ac) show ?thesis using assms by(simp add:fold_image_def I.fold_insert) qed (* lemma fold_commute: "finite A ==> (!!z. x * (fold times g z A) = fold times g (x * z) A)" apply (induct set: finite) apply simp apply (simp add: mult_left_commute [of x]) done lemma fold_nest_Un_Int: "finite A ==> finite B ==> fold times g (fold times g z B) A = fold times g (fold times g z (A Int B)) (A Un B)" apply (induct set: finite) apply simp apply (simp add: fold_commute Int_insert_left insert_absorb) done lemma fold_nest_Un_disjoint: "finite A ==> finite B ==> A Int B = {} ==> fold times g z (A Un B) = fold times g (fold times g z B) A" by (simp add: fold_nest_Un_Int) *) lemma fold_image_reindex: assumes fin: "finite A" shows "inj_on h A \ fold_image times g z (h`A) = fold_image times (g\h) z A" using fin apply induct apply simp apply simp done (* text{* Fusion theorem, as described in Graham Hutton's paper, A Tutorial on the Universality and Expressiveness of Fold, JFP 9:4 (355-372), 1999. *} lemma fold_fusion: assumes "ab_semigroup_mult g" assumes fin: "finite A" and hyp: "\x y. h (g x y) = times x (h y)" shows "h (fold g j w A) = fold times j (h w) A" proof - class_interpret ab_semigroup_mult [g] by fact show ?thesis using fin hyp by (induct set: finite) simp_all qed *) lemma fold_image_cong: "finite A \ (!!x. x:A ==> g x = h x) ==> fold_image times g z A = fold_image times h z A" apply (subgoal_tac "ALL C. C <= A --> (ALL x:C. g x = h x) --> fold_image times g z C = fold_image times h z C") apply simp apply (erule finite_induct, simp) apply (simp add: subset_insert_iff, clarify) apply (subgoal_tac "finite C") prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl]) apply (subgoal_tac "C = insert x (C - {x})") prefer 2 apply blast apply (erule ssubst) apply (drule spec) apply (erule (1) notE impE) apply (simp add: Ball_def del: insert_Diff_single) done end context comm_monoid_mult begin lemma fold_image_Un_Int: "finite A ==> finite B ==> fold_image times g 1 A * fold_image times g 1 B = fold_image times g 1 (A Un B) * fold_image times g 1 (A Int B)" by (induct set: finite) (auto simp add: mult_ac insert_absorb Int_insert_left) corollary fold_Un_disjoint: "finite A ==> finite B ==> A Int B = {} ==> fold_image times g 1 (A Un B) = fold_image times g 1 A * fold_image times g 1 B" by (simp add: fold_image_Un_Int) lemma fold_image_UN_disjoint: "\ finite I; ALL i:I. finite (A i); ALL i:I. ALL j:I. i \ j --> A i Int A j = {} \ \ fold_image times g 1 (UNION I A) = fold_image times (%i. fold_image times g 1 (A i)) 1 I" apply (induct set: finite, simp, atomize) apply (subgoal_tac "ALL i:F. x \ i") prefer 2 apply blast apply (subgoal_tac "A x Int UNION F A = {}") prefer 2 apply blast apply (simp add: fold_Un_disjoint) done lemma fold_image_Sigma: "finite A ==> ALL x:A. finite (B x) ==> fold_image times (%x. fold_image times (g x) 1 (B x)) 1 A = fold_image times (split g) 1 (SIGMA x:A. B x)" apply (subst Sigma_def) apply (subst fold_image_UN_disjoint, assumption, simp) apply blast apply (erule fold_image_cong) apply (subst fold_image_UN_disjoint, simp, simp) apply blast apply simp done lemma fold_image_distrib: "finite A \ fold_image times (%x. g x * h x) 1 A = fold_image times g 1 A * fold_image times h 1 A" by (erule finite_induct) (simp_all add: mult_ac) lemma fold_image_related: assumes Re: "R e e" and Rop: "\x1 y1 x2 y2. R x1 x2 \ R y1 y2 \ R (x1 * y1) (x2 * y2)" and fS: "finite S" and Rfg: "\x\S. R (h x) (g x)" shows "R (fold_image (op *) h e S) (fold_image (op *) g e S)" using fS by (rule finite_subset_induct) (insert assms, auto) lemma fold_image_eq_general: assumes fS: "finite S" and h: "\y\S'. \!x. x\ S \ h(x) = y" and f12: "\x\S. h x \ S' \ f2(h x) = f1 x" shows "fold_image (op *) f1 e S = fold_image (op *) f2 e S'" proof- from h f12 have hS: "h ` S = S'" by auto {fix x y assume H: "x \ S" "y \ S" "h x = h y" from f12 h H have "x = y" by auto } hence hinj: "inj_on h S" unfolding inj_on_def Ex1_def by blast from f12 have th: "\x. x \ S \ (f2 \ h) x = f1 x" by auto from hS have "fold_image (op *) f2 e S' = fold_image (op *) f2 e (h ` S)" by simp also have "\ = fold_image (op *) (f2 o h) e S" using fold_image_reindex[OF fS hinj, of f2 e] . also have "\ = fold_image (op *) f1 e S " using th fold_image_cong[OF fS, of "f2 o h" f1 e] by blast finally show ?thesis .. qed lemma fold_image_eq_general_inverses: assumes fS: "finite S" and kh: "\y. y \ T \ k y \ S \ h (k y) = y" and hk: "\x. x \ S \ h x \ T \ k (h x) = x \ g (h x) = f x" shows "fold_image (op *) f e S = fold_image (op *) g e T" (* metis solves it, but not yet available here *) apply (rule fold_image_eq_general[OF fS, of T h g f e]) apply (rule ballI) apply (frule kh) apply (rule ex1I[]) apply blast apply clarsimp apply (drule hk) apply simp apply (rule sym) apply (erule conjunct1[OF conjunct2[OF hk]]) apply (rule ballI) apply (drule hk) apply blast done end subsection {* Generalized summation over a set *} -interpretation comm_monoid_add!: comm_monoid_mult "0::'a::comm_monoid_add" "op +" +interpretation comm_monoid_add: comm_monoid_mult "0::'a::comm_monoid_add" "op +" proof qed (auto intro: add_assoc add_commute) definition setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add" where "setsum f A == if finite A then fold_image (op +) f 0 A else 0" abbreviation Setsum ("\_" [1000] 999) where "\A == setsum (%x. x) A" text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is written @{text"\x\A. e"}. *} syntax "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_. _)" [0, 51, 10] 10) syntax (xsymbols) "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\_\_. _)" [0, 51, 10] 10) syntax (HTML output) "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\_\_. _)" [0, 51, 10] 10) translations -- {* Beware of argument permutation! *} "SUM i:A. b" == "CONST setsum (%i. b) A" "\i\A. b" == "CONST setsum (%i. b) A" text{* Instead of @{term"\x\{x. P}. e"} we introduce the shorter @{text"\x|P. e"}. *} syntax "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10) syntax (xsymbols) "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) syntax (HTML output) "_qsetsum" :: "pttrn \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) translations "SUM x|P. t" => "CONST setsum (%x. t) {x. P}" "\x|P. t" => "CONST setsum (%x. t) {x. P}" print_translation {* let fun setsum_tr' [Abs(x,Tx,t), Const ("Collect",_) $ Abs(y,Ty,P)] = if x<>y then raise Match else let val x' = Syntax.mark_bound x val t' = subst_bound(x',t) val P' = subst_bound(x',P) in Syntax.const "_qsetsum" $ Syntax.mark_bound x $ P' $ t' end in [("setsum", setsum_tr')] end *} lemma setsum_empty [simp]: "setsum f {} = 0" by (simp add: setsum_def) lemma setsum_insert [simp]: "finite F ==> a \ F ==> setsum f (insert a F) = f a + setsum f F" by (simp add: setsum_def) lemma setsum_infinite [simp]: "~ finite A ==> setsum f A = 0" by (simp add: setsum_def) lemma setsum_reindex: "inj_on f B ==> setsum h (f ` B) = setsum (h \ f) B" by(auto simp add: setsum_def comm_monoid_add.fold_image_reindex dest!:finite_imageD) lemma setsum_reindex_id: "inj_on f B ==> setsum f B = setsum id (f ` B)" by (auto simp add: setsum_reindex) lemma setsum_reindex_nonzero: assumes fS: "finite S" and nz: "\ x y. x \ S \ y \ S \ x \ y \ f x = f y \ h (f x) = 0" shows "setsum h (f ` S) = setsum (h o f) S" using nz proof(induct rule: finite_induct[OF fS]) case 1 thus ?case by simp next case (2 x F) {assume fxF: "f x \ f ` F" hence "\y \ F . f y = f x" by auto then obtain y where y: "y \ F" "f x = f y" by auto from "2.hyps" y have xy: "x \ y" by auto from "2.prems"[of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto also have "\ = setsum (h o f) (insert x F)" unfolding setsum_insert[OF `finite F` `x\F`] using h0 apply simp apply (rule "2.hyps"(3)) apply (rule_tac y="y" in "2.prems") apply simp_all done finally have ?case .} moreover {assume fxF: "f x \ f ` F" have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)" using fxF "2.hyps" by simp also have "\ = setsum (h o f) (insert x F)" unfolding setsum_insert[OF `finite F` `x\F`] apply simp apply (rule cong[OF refl[of "op + (h (f x))"]]) apply (rule "2.hyps"(3)) apply (rule_tac y="y" in "2.prems") apply simp_all done finally have ?case .} ultimately show ?case by blast qed lemma setsum_cong: "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B" by(fastsimp simp: setsum_def intro: comm_monoid_add.fold_image_cong) lemma strong_setsum_cong[cong]: "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setsum (%x. f x) A = setsum (%x. g x) B" by(fastsimp simp: simp_implies_def setsum_def intro: comm_monoid_add.fold_image_cong) lemma setsum_cong2: "\\x. x \ A \ f x = g x\ \ setsum f A = setsum g A"; by (rule setsum_cong[OF refl], auto); lemma setsum_reindex_cong: "[|inj_on f A; B = f ` A; !!a. a:A \ g a = h (f a)|] ==> setsum h B = setsum g A" by (simp add: setsum_reindex cong: setsum_cong) lemma setsum_0[simp]: "setsum (%i. 0) A = 0" apply (clarsimp simp: setsum_def) apply (erule finite_induct, auto) done lemma setsum_0': "ALL a:A. f a = 0 ==> setsum f A = 0" by(simp add:setsum_cong) lemma setsum_Un_Int: "finite A ==> finite B ==> setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B" -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *} by(simp add: setsum_def comm_monoid_add.fold_image_Un_Int [symmetric]) lemma setsum_Un_disjoint: "finite A ==> finite B ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B" by (subst setsum_Un_Int [symmetric], auto) lemma setsum_mono_zero_left: assumes fT: "finite T" and ST: "S \ T" and z: "\i \ T - S. f i = 0" shows "setsum f S = setsum f T" proof- have eq: "T = S \ (T - S)" using ST by blast have d: "S \ (T - S) = {}" using ST by blast from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset) show ?thesis by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z]) qed lemma setsum_mono_zero_right: assumes fT: "finite T" and ST: "S \ T" and z: "\i \ T - S. f i = 0" shows "setsum f T = setsum f S" using setsum_mono_zero_left[OF fT ST z] by simp lemma setsum_mono_zero_cong_left: assumes fT: "finite T" and ST: "S \ T" and z: "\i \ T - S. g i = 0" and fg: "\x. x \ S \ f x = g x" shows "setsum f S = setsum g T" proof- have eq: "T = S \ (T - S)" using ST by blast have d: "S \ (T - S) = {}" using ST by blast from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset) show ?thesis using fg by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z]) qed lemma setsum_mono_zero_cong_right: assumes fT: "finite T" and ST: "S \ T" and z: "\i \ T - S. f i = 0" and fg: "\x. x \ S \ f x = g x" shows "setsum f T = setsum g S" using setsum_mono_zero_cong_left[OF fT ST z] fg[symmetric] by auto lemma setsum_delta: assumes fS: "finite S" shows "setsum (\k. if k=a then b k else 0) S = (if a \ S then b a else 0)" proof- let ?f = "(\k. if k=a then b k else 0)" {assume a: "a \ S" hence "\ k\ S. ?f k = 0" by simp hence ?thesis using a by simp} moreover {assume a: "a \ S" let ?A = "S - {a}" let ?B = "{a}" have eq: "S = ?A \ ?B" using a by blast have dj: "?A \ ?B = {}" by simp from fS have fAB: "finite ?A" "finite ?B" by auto have "setsum ?f S = setsum ?f ?A + setsum ?f ?B" using setsum_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]] by simp then have ?thesis using a by simp} ultimately show ?thesis by blast qed lemma setsum_delta': assumes fS: "finite S" shows "setsum (\k. if a = k then b k else 0) S = (if a\ S then b a else 0)" using setsum_delta[OF fS, of a b, symmetric] by (auto intro: setsum_cong) lemma setsum_restrict_set: assumes fA: "finite A" shows "setsum f (A \ B) = setsum (\x. if x \ B then f x else 0) A" proof- from fA have fab: "finite (A \ B)" by auto have aba: "A \ B \ A" by blast let ?g = "\x. if x \ A\B then f x else 0" from setsum_mono_zero_left[OF fA aba, of ?g] show ?thesis by simp qed lemma setsum_cases: assumes fA: "finite A" shows "setsum (\x. if x \ B then f x else g x) A = setsum f (A \ B) + setsum g (A \ - B)" proof- have a: "A = A \ B \ A \ -B" "(A \ B) \ (A \ -B) = {}" by blast+ from fA have f: "finite (A \ B)" "finite (A \ -B)" by auto let ?g = "\x. if x \ B then f x else g x" from setsum_Un_disjoint[OF f a(2), of ?g] a(1) show ?thesis by simp qed (*But we can't get rid of finite I. If infinite, although the rhs is 0, the lhs need not be, since UNION I A could still be finite.*) lemma setsum_UN_disjoint: "finite I ==> (ALL i:I. finite (A i)) ==> (ALL i:I. ALL j:I. i \ j --> A i Int A j = {}) ==> setsum f (UNION I A) = (\i\I. setsum f (A i))" by(simp add: setsum_def comm_monoid_add.fold_image_UN_disjoint cong: setsum_cong) text{*No need to assume that @{term C} is finite. If infinite, the rhs is directly 0, and @{term "Union C"} is also infinite, hence the lhs is also 0.*} lemma setsum_Union_disjoint: "[| (ALL A:C. finite A); (ALL A:C. ALL B:C. A \ B --> A Int B = {}) |] ==> setsum f (Union C) = setsum (setsum f) C" apply (cases "finite C") prefer 2 apply (force dest: finite_UnionD simp add: setsum_def) apply (frule setsum_UN_disjoint [of C id f]) apply (unfold Union_def id_def, assumption+) done (*But we can't get rid of finite A. If infinite, although the lhs is 0, the rhs need not be, since SIGMA A B could still be finite.*) lemma setsum_Sigma: "finite A ==> ALL x:A. finite (B x) ==> (\x\A. (\y\B x. f x y)) = (\(x,y)\(SIGMA x:A. B x). f x y)" by(simp add:setsum_def comm_monoid_add.fold_image_Sigma split_def cong:setsum_cong) text{*Here we can eliminate the finiteness assumptions, by cases.*} lemma setsum_cartesian_product: "(\x\A. (\y\B. f x y)) = (\(x,y) \ A <*> B. f x y)" apply (cases "finite A") apply (cases "finite B") apply (simp add: setsum_Sigma) apply (cases "A={}", simp) apply (simp) apply (auto simp add: setsum_def dest: finite_cartesian_productD1 finite_cartesian_productD2) done lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)" by(simp add:setsum_def comm_monoid_add.fold_image_distrib) subsubsection {* Properties in more restricted classes of structures *} lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a" apply (case_tac "finite A") prefer 2 apply (simp add: setsum_def) apply (erule rev_mp) apply (erule finite_induct, auto) done lemma setsum_eq_0_iff [simp]: "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))" by (induct set: finite) auto lemma setsum_Un_nat: "finite A ==> finite B ==> (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)" -- {* For the natural numbers, we have subtraction. *} by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps) lemma setsum_Un: "finite A ==> finite B ==> (setsum f (A Un B) :: 'a :: ab_group_add) = setsum f A + setsum f B - setsum f (A Int B)" by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps) lemma (in comm_monoid_mult) fold_image_1: "finite S \ (\x\S. f x = 1) \ fold_image op * f 1 S = 1" apply (induct set: finite) apply simp by (auto simp add: fold_image_insert) lemma (in comm_monoid_mult) fold_image_Un_one: assumes fS: "finite S" and fT: "finite T" and I0: "\x \ S\T. f x = 1" shows "fold_image (op *) f 1 (S \ T) = fold_image (op *) f 1 S * fold_image (op *) f 1 T" proof- have "fold_image op * f 1 (S \ T) = 1" apply (rule fold_image_1) using fS fT I0 by auto with fold_image_Un_Int[OF fS fT] show ?thesis by simp qed lemma setsum_eq_general_reverses: assumes fS: "finite S" and fT: "finite T" and kh: "\y. y \ T \ k y \ S \ h (k y) = y" and hk: "\x. x \ S \ h x \ T \ k (h x) = x \ g (h x) = f x" shows "setsum f S = setsum g T" apply (simp add: setsum_def fS fT) apply (rule comm_monoid_add.fold_image_eq_general_inverses[OF fS]) apply (erule kh) apply (erule hk) done lemma setsum_Un_zero: assumes fS: "finite S" and fT: "finite T" and I0: "\x \ S\T. f x = 0" shows "setsum f (S \ T) = setsum f S + setsum f T" using fS fT apply (simp add: setsum_def) apply (rule comm_monoid_add.fold_image_Un_one) using I0 by auto lemma setsum_UNION_zero: assumes fS: "finite S" and fSS: "\T \ S. finite T" and f0: "\T1 T2 x. T1\S \ T2\S \ T1 \ T2 \ x \ T1 \ x \ T2 \ f x = 0" shows "setsum f (\S) = setsum (\T. setsum f T) S" using fSS f0 proof(induct rule: finite_induct[OF fS]) case 1 thus ?case by simp next case (2 T F) then have fTF: "finite T" "\T\F. finite T" "finite F" and TF: "T \ F" and H: "setsum f (\ F) = setsum (setsum f) F" by (auto simp add: finite_insert) from fTF have fUF: "finite (\F)" by (auto intro: finite_Union) from "2.prems" TF fTF show ?case by (auto simp add: H[symmetric] intro: setsum_Un_zero[OF fTF(1) fUF, of f]) qed lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) = (if a:A then setsum f A - f a else setsum f A)" apply (case_tac "finite A") prefer 2 apply (simp add: setsum_def) apply (erule finite_induct) apply (auto simp add: insert_Diff_if) apply (drule_tac a = a in mk_disjoint_insert, auto) done lemma setsum_diff1: "finite A \ (setsum f (A - {a}) :: ('a::ab_group_add)) = (if a:A then setsum f A - f a else setsum f A)" by (erule finite_induct) (auto simp add: insert_Diff_if) lemma setsum_diff1'[rule_format]: "finite A \ a \ A \ (\ x \ A. f x) = f a + (\ x \ (A - {a}). f x)" apply (erule finite_induct[where F=A and P="% A. (a \ A \ (\ x \ A. f x) = f a + (\ x \ (A - {a}). f x))"]) apply (auto simp add: insert_Diff_if add_ac) done (* By Jeremy Siek: *) lemma setsum_diff_nat: assumes "finite B" and "B \ A" shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)" using assms proof induct show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp next fix F x assume finF: "finite F" and xnotinF: "x \ F" and xFinA: "insert x F \ A" and IH: "F \ A \ setsum f (A - F) = setsum f A - setsum f F" from xnotinF xFinA have xinAF: "x \ (A - F)" by simp from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x" by (simp add: setsum_diff1_nat) from xFinA have "F \ A" by simp with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x" by simp from xnotinF have "A - insert x F = (A - F) - {x}" by auto with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x" by simp from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp qed lemma setsum_diff: assumes le: "finite A" "B \ A" shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))" proof - from le have finiteB: "finite B" using finite_subset by auto show ?thesis using finiteB le proof induct case empty thus ?case by auto next case (insert x F) thus ?case using le finiteB by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb) qed qed lemma setsum_mono: assumes le: "\i. i\K \ f (i::'a) \ ((g i)::('b::{comm_monoid_add, pordered_ab_semigroup_add}))" shows "(\i\K. f i) \ (\i\K. g i)" proof (cases "finite K") case True thus ?thesis using le proof induct case empty thus ?case by simp next case insert thus ?case using add_mono by fastsimp qed next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_strict_mono: fixes f :: "'a \ 'b::{pordered_cancel_ab_semigroup_add,comm_monoid_add}" assumes "finite A" "A \ {}" and "!!x. x:A \ f x < g x" shows "setsum f A < setsum g A" using prems proof (induct rule: finite_ne_induct) case singleton thus ?case by simp next case insert thus ?case by (auto simp: add_strict_mono) qed lemma setsum_negf: "setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A" proof (cases "finite A") case True thus ?thesis by (induct set: finite) auto next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_subtractf: "setsum (%x. ((f x)::'a::ab_group_add) - g x) A = setsum f A - setsum g A" proof (cases "finite A") case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf) next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_nonneg: assumes nn: "\x\A. (0::'a::{pordered_ab_semigroup_add,comm_monoid_add}) \ f x" shows "0 \ setsum f A" proof (cases "finite A") case True thus ?thesis using nn proof induct case empty then show ?case by simp next case (insert x F) then have "0 + 0 \ f x + setsum f F" by (blast intro: add_mono) with insert show ?case by simp qed next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_nonpos: assumes np: "\x\A. f x \ (0::'a::{pordered_ab_semigroup_add,comm_monoid_add})" shows "setsum f A \ 0" proof (cases "finite A") case True thus ?thesis using np proof induct case empty then show ?case by simp next case (insert x F) then have "f x + setsum f F \ 0 + 0" by (blast intro: add_mono) with insert show ?case by simp qed next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_mono2: fixes f :: "'a \ 'b :: {pordered_ab_semigroup_add_imp_le,comm_monoid_add}" assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 0 \ f b" shows "setsum f A \ setsum f B" proof - have "setsum f A \ setsum f A + setsum f (B-A)" by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def) also have "\ = setsum f (A \ (B-A))" using fin finite_subset[OF sub fin] by (simp add:setsum_Un_disjoint del:Un_Diff_cancel) also have "A \ (B-A) = B" using sub by blast finally show ?thesis . qed lemma setsum_mono3: "finite B ==> A <= B ==> ALL x: B - A. 0 <= ((f x)::'a::{comm_monoid_add,pordered_ab_semigroup_add}) ==> setsum f A <= setsum f B" apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)") apply (erule ssubst) apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)") apply simp apply (rule add_left_mono) apply (erule setsum_nonneg) apply (subst setsum_Un_disjoint [THEN sym]) apply (erule finite_subset, assumption) apply (rule finite_subset) prefer 2 apply assumption apply auto apply (rule setsum_cong) apply auto done lemma setsum_right_distrib: fixes f :: "'a => ('b::semiring_0)" shows "r * setsum f A = setsum (%n. r * f n) A" proof (cases "finite A") case True thus ?thesis proof induct case empty thus ?case by simp next case (insert x A) thus ?case by (simp add: right_distrib) qed next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_left_distrib: "setsum f A * (r::'a::semiring_0) = (\n\A. f n * r)" proof (cases "finite A") case True then show ?thesis proof induct case empty thus ?case by simp next case (insert x A) thus ?case by (simp add: left_distrib) qed next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_divide_distrib: "setsum f A / (r::'a::field) = (\n\A. f n / r)" proof (cases "finite A") case True then show ?thesis proof induct case empty thus ?case by simp next case (insert x A) thus ?case by (simp add: add_divide_distrib) qed next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_abs[iff]: fixes f :: "'a => ('b::pordered_ab_group_add_abs)" shows "abs (setsum f A) \ setsum (%i. abs(f i)) A" proof (cases "finite A") case True thus ?thesis proof induct case empty thus ?case by simp next case (insert x A) thus ?case by (auto intro: abs_triangle_ineq order_trans) qed next case False thus ?thesis by (simp add: setsum_def) qed lemma setsum_abs_ge_zero[iff]: fixes f :: "'a => ('b::pordered_ab_group_add_abs)" shows "0 \ setsum (%i. abs(f i)) A" proof (cases "finite A") case True thus ?thesis proof induct case empty thus ?case by simp next case (insert x A) thus ?case by (auto simp: add_nonneg_nonneg) qed next case False thus ?thesis by (simp add: setsum_def) qed lemma abs_setsum_abs[simp]: fixes f :: "'a => ('b::pordered_ab_group_add_abs)" shows "abs (\a\A. abs(f a)) = (\a\A. abs(f a))" proof (cases "finite A") case True thus ?thesis proof induct case empty thus ?case by simp next case (insert a A) hence "\\a\insert a A. \f a\\ = \\f a\ + (\a\A. \f a\)\" by simp also have "\ = \\f a\ + \\a\A. \f a\\\" using insert by simp also have "\ = \f a\ + \\a\A. \f a\\" by (simp del: abs_of_nonneg) also have "\ = (\a\insert a A. \f a\)" using insert by simp finally show ?case . qed next case False thus ?thesis by (simp add: setsum_def) qed text {* Commuting outer and inner summation *} lemma swap_inj_on: "inj_on (%(i, j). (j, i)) (A \ B)" by (unfold inj_on_def) fast lemma swap_product: "(%(i, j). (j, i)) ` (A \ B) = B \ A" by (simp add: split_def image_def) blast lemma setsum_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" proof (simp add: setsum_cartesian_product) have "(\(x,y) \ A <*> B. f x y) = (\(y,x) \ (%(i, j). (j, i)) ` (A \ B). f x y)" (is "?s = _") apply (simp add: setsum_reindex [where f = "%(i, j). (j, i)"] swap_inj_on) apply (simp add: split_def) done also have "... = (\(y,x)\B \ A. f x y)" (is "_ = ?t") apply (simp add: swap_product) done finally show "?s = ?t" . qed lemma setsum_product: fixes f :: "'a => ('b::semiring_0)" shows "setsum f A * setsum g B = (\i\A. \j\B. f i * g j)" by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute) subsection {* Generalized product over a set *} definition setprod :: "('a => 'b) => 'a set => 'b::comm_monoid_mult" where "setprod f A == if finite A then fold_image (op *) f 1 A else 1" abbreviation Setprod ("\_" [1000] 999) where "\A == setprod (%x. x) A" syntax "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3PROD _:_. _)" [0, 51, 10] 10) syntax (xsymbols) "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\_\_. _)" [0, 51, 10] 10) syntax (HTML output) "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\_\_. _)" [0, 51, 10] 10) translations -- {* Beware of argument permutation! *} "PROD i:A. b" == "CONST setprod (%i. b) A" "\i\A. b" == "CONST setprod (%i. b) A" text{* Instead of @{term"\x\{x. P}. e"} we introduce the shorter @{text"\x|P. e"}. *} syntax "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10) syntax (xsymbols) "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) syntax (HTML output) "_qsetprod" :: "pttrn \ bool \ 'a \ 'a" ("(3\_ | (_)./ _)" [0,0,10] 10) translations "PROD x|P. t" => "CONST setprod (%x. t) {x. P}" "\x|P. t" => "CONST setprod (%x. t) {x. P}" lemma setprod_empty [simp]: "setprod f {} = 1" by (auto simp add: setprod_def) lemma setprod_insert [simp]: "[| finite A; a \ A |] ==> setprod f (insert a A) = f a * setprod f A" by (simp add: setprod_def) lemma setprod_infinite [simp]: "~ finite A ==> setprod f A = 1" by (simp add: setprod_def) lemma setprod_reindex: "inj_on f B ==> setprod h (f ` B) = setprod (h \ f) B" by(auto simp: setprod_def fold_image_reindex dest!:finite_imageD) lemma setprod_reindex_id: "inj_on f B ==> setprod f B = setprod id (f ` B)" by (auto simp add: setprod_reindex) lemma setprod_cong: "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B" by(fastsimp simp: setprod_def intro: fold_image_cong) lemma strong_setprod_cong: "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B" by(fastsimp simp: simp_implies_def setprod_def intro: fold_image_cong) lemma setprod_reindex_cong: "inj_on f A ==> B = f ` A ==> g = h \ f ==> setprod h B = setprod g A" by (frule setprod_reindex, simp) lemma strong_setprod_reindex_cong: assumes i: "inj_on f A" and B: "B = f ` A" and eq: "\x. x \ A \ g x = (h \ f) x" shows "setprod h B = setprod g A" proof- have "setprod h B = setprod (h o f) A" by (simp add: B setprod_reindex[OF i, of h]) then show ?thesis apply simp apply (rule setprod_cong) apply simp by (erule eq[symmetric]) qed lemma setprod_Un_one: assumes fS: "finite S" and fT: "finite T" and I0: "\x \ S\T. f x = 1" shows "setprod f (S \ T) = setprod f S * setprod f T" using fS fT apply (simp add: setprod_def) apply (rule fold_image_Un_one) using I0 by auto lemma setprod_1: "setprod (%i. 1) A = 1" apply (case_tac "finite A") apply (erule finite_induct, auto simp add: mult_ac) done lemma setprod_1': "ALL a:F. f a = 1 ==> setprod f F = 1" apply (subgoal_tac "setprod f F = setprod (%x. 1) F") apply (erule ssubst, rule setprod_1) apply (rule setprod_cong, auto) done lemma setprod_Un_Int: "finite A ==> finite B ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B" by(simp add: setprod_def fold_image_Un_Int[symmetric]) lemma setprod_Un_disjoint: "finite A ==> finite B ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B" by (subst setprod_Un_Int [symmetric], auto) lemma setprod_delta: assumes fS: "finite S" shows "setprod (\k. if k=a then b k else 1) S = (if a \ S then b a else 1)" proof- let ?f = "(\k. if k=a then b k else 1)" {assume a: "a \ S" hence "\ k\ S. ?f k = 1" by simp hence ?thesis using a by (simp add: setprod_1 cong add: setprod_cong) } moreover {assume a: "a \ S" let ?A = "S - {a}" let ?B = "{a}" have eq: "S = ?A \ ?B" using a by blast have dj: "?A \ ?B = {}" by simp from fS have fAB: "finite ?A" "finite ?B" by auto have fA1: "setprod ?f ?A = 1" apply (rule setprod_1') by auto have "setprod ?f ?A * setprod ?f ?B = setprod ?f S" using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]] by simp then have ?thesis using a by (simp add: fA1 cong add: setprod_cong cong del: if_weak_cong)} ultimately show ?thesis by blast qed lemma setprod_delta': assumes fS: "finite S" shows "setprod (\k. if a = k then b k else 1) S = (if a\ S then b a else 1)" using setprod_delta[OF fS, of a b, symmetric] by (auto intro: setprod_cong) lemma setprod_UN_disjoint: "finite I ==> (ALL i:I. finite (A i)) ==> (ALL i:I. ALL j:I. i \ j --> A i Int A j = {}) ==> setprod f (UNION I A) = setprod (%i. setprod f (A i)) I" by(simp add: setprod_def fold_image_UN_disjoint cong: setprod_cong) lemma setprod_Union_disjoint: "[| (ALL A:C. finite A); (ALL A:C. ALL B:C. A \ B --> A Int B = {}) |] ==> setprod f (Union C) = setprod (setprod f) C" apply (cases "finite C") prefer 2 apply (force dest: finite_UnionD simp add: setprod_def) apply (frule setprod_UN_disjoint [of C id f]) apply (unfold Union_def id_def, assumption+) done lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==> (\x\A. (\y\ B x. f x y)) = (\(x,y)\(SIGMA x:A. B x). f x y)" by(simp add:setprod_def fold_image_Sigma split_def cong:setprod_cong) text{*Here we can eliminate the finiteness assumptions, by cases.*} lemma setprod_cartesian_product: "(\x\A. (\y\ B. f x y)) = (\(x,y)\(A <*> B). f x y)" apply (cases "finite A") apply (cases "finite B") apply (simp add: setprod_Sigma) apply (cases "A={}", simp) apply (simp add: setprod_1) apply (auto simp add: setprod_def dest: finite_cartesian_productD1 finite_cartesian_productD2) done lemma setprod_timesf: "setprod (%x. f x * g x) A = (setprod f A * setprod g A)" by(simp add:setprod_def fold_image_distrib) subsubsection {* Properties in more restricted classes of structures *} lemma setprod_eq_1_iff [simp]: "finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))" by (induct set: finite) auto lemma setprod_zero: "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0" apply (induct set: finite, force, clarsimp) apply (erule disjE, auto) done lemma setprod_nonneg [rule_format]: "(ALL x: A. (0::'a::ordered_idom) \ f x) --> 0 \ setprod f A" apply (case_tac "finite A") apply (induct set: finite, force, clarsimp) apply (subgoal_tac "0 * 0 \ f x * setprod f F", force) apply (rule mult_mono, assumption+) apply (auto simp add: setprod_def) done lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::ordered_idom) < f x) --> 0 < setprod f A" apply (case_tac "finite A") apply (induct set: finite, force, clarsimp) apply (subgoal_tac "0 * 0 < f x * setprod f F", force) apply (rule mult_strict_mono, assumption+) apply (auto simp add: setprod_def) done lemma setprod_nonzero [rule_format]: "(ALL x y. (x::'a::comm_semiring_1) * y = 0 --> x = 0 | y = 0) ==> finite A ==> (ALL x: A. f x \ (0::'a)) --> setprod f A \ 0" by (erule finite_induct, auto) lemma setprod_zero_eq: "(ALL x y. (x::'a::comm_semiring_1) * y = 0 --> x = 0 | y = 0) ==> finite A ==> (setprod f A = (0::'a)) = (EX x: A. f x = 0)" by (insert setprod_zero [of A f] setprod_nonzero [of A f], blast) lemma setprod_nonzero_field: "finite A ==> (ALL x: A. f x \ (0::'a::idom)) ==> setprod f A \ 0" by (rule setprod_nonzero, auto) lemma setprod_zero_eq_field: "finite A ==> (setprod f A = (0::'a::idom)) = (EX x: A. f x = 0)" by (rule setprod_zero_eq, auto) lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \ 0) ==> (setprod f (A Un B) :: 'a ::{field}) = setprod f A * setprod f B / setprod f (A Int B)" apply (subst setprod_Un_Int [symmetric], auto) apply (subgoal_tac "finite (A Int B)") apply (frule setprod_nonzero_field [of "A Int B" f], assumption) apply (subst times_divide_eq_right [THEN sym], auto) done lemma setprod_diff1: "finite A ==> f a \ 0 ==> (setprod f (A - {a}) :: 'a :: {field}) = (if a:A then setprod f A / f a else setprod f A)" by (erule finite_induct) (auto simp add: insert_Diff_if) lemma setprod_inversef: "finite A ==> ALL x: A. f x \ (0::'a::{field,division_by_zero}) ==> setprod (inverse \ f) A = inverse (setprod f A)" by (erule finite_induct) auto lemma setprod_dividef: "[|finite A; \x \ A. g x \ (0::'a::{field,division_by_zero})|] ==> setprod (%x. f x / g x) A = setprod f A / setprod g A" apply (subgoal_tac "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \ g) x) A") apply (erule ssubst) apply (subst divide_inverse) apply (subst setprod_timesf) apply (subst setprod_inversef, assumption+, rule refl) apply (rule setprod_cong, rule refl) apply (subst divide_inverse, auto) done lemma setprod_dvd_setprod [rule_format]: "(ALL x : A. f x dvd g x) \ setprod f A dvd setprod g A" apply (cases "finite A") apply (induct set: finite) apply (auto simp add: dvd_def) apply (rule_tac x = "k * ka" in exI) apply (simp add: algebra_simps) done lemma setprod_dvd_setprod_subset: "finite B \ A <= B \ setprod f A dvd setprod f B" apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)") apply (unfold dvd_def, blast) apply (subst setprod_Un_disjoint [symmetric]) apply (auto elim: finite_subset intro: setprod_cong) done lemma setprod_dvd_setprod_subset2: "finite B \ A <= B \ ALL x : A. (f x::'a::comm_semiring_1) dvd g x \ setprod f A dvd setprod g B" apply (rule dvd_trans) apply (rule setprod_dvd_setprod, erule (1) bspec) apply (erule (1) setprod_dvd_setprod_subset) done lemma dvd_setprod: "finite A \ i:A \ (f i ::'a::comm_semiring_1) dvd setprod f A" by (induct set: finite) (auto intro: dvd_mult) lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \ (d::'a::comm_semiring_1) dvd (SUM x : A. f x)" apply (cases "finite A") apply (induct set: finite) apply auto done subsection {* Finite cardinality *} text {* This definition, although traditional, is ugly to work with: @{text "card A == LEAST n. EX f. A = {f i | i. i < n}"}. But now that we have @{text setsum} things are easy: *} definition card :: "'a set \ nat" where "card A = setsum (\x. 1) A" lemma card_empty [simp]: "card {} = 0" by (simp add: card_def) lemma card_infinite [simp]: "~ finite A ==> card A = 0" by (simp add: card_def) lemma card_eq_setsum: "card A = setsum (%x. 1) A" by (simp add: card_def) lemma card_insert_disjoint [simp]: "finite A ==> x \ A ==> card (insert x A) = Suc(card A)" by(simp add: card_def) lemma card_insert_if: "finite A ==> card (insert x A) = (if x:A then card A else Suc(card(A)))" by (simp add: insert_absorb) lemma card_0_eq [simp,noatp]: "finite A ==> (card A = 0) = (A = {})" apply auto apply (drule_tac a = x in mk_disjoint_insert, clarify, auto) done lemma card_eq_0_iff: "(card A = 0) = (A = {} | ~ finite A)" by auto lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A" apply(rule_tac t = A in insert_Diff [THEN subst], assumption) apply(simp del:insert_Diff_single) done lemma card_Diff_singleton: "finite A ==> x: A ==> card (A - {x}) = card A - 1" by (simp add: card_Suc_Diff1 [symmetric]) lemma card_Diff_singleton_if: "finite A ==> card (A-{x}) = (if x : A then card A - 1 else card A)" by (simp add: card_Diff_singleton) lemma card_Diff_insert[simp]: assumes "finite A" and "a:A" and "a ~: B" shows "card(A - insert a B) = card(A - B) - 1" proof - have "A - insert a B = (A - B) - {a}" using assms by blast then show ?thesis using assms by(simp add:card_Diff_singleton) qed lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))" by (simp add: card_insert_if card_Suc_Diff1 del:card_Diff_insert) lemma card_insert_le: "finite A ==> card A <= card (insert x A)" by (simp add: card_insert_if) lemma card_mono: "\ finite B; A \ B \ \ card A \ card B" by (simp add: card_def setsum_mono2) lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)" apply (induct set: finite, simp, clarify) apply (subgoal_tac "finite A & A - {x} <= F") prefer 2 apply (blast intro: finite_subset, atomize) apply (drule_tac x = "A - {x}" in spec) apply (simp add: card_Diff_singleton_if split add: split_if_asm) apply (case_tac "card A", auto) done lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B" apply (simp add: psubset_eq linorder_not_le [symmetric]) apply (blast dest: card_seteq) done lemma card_Un_Int: "finite A ==> finite B ==> card A + card B = card (A Un B) + card (A Int B)" by(simp add:card_def setsum_Un_Int) lemma card_Un_disjoint: "finite A ==> finite B ==> A Int B = {} ==> card (A Un B) = card A + card B" by (simp add: card_Un_Int) lemma card_Diff_subset: "finite B ==> B <= A ==> card (A - B) = card A - card B" by(simp add:card_def setsum_diff_nat) lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A" apply (rule Suc_less_SucD) apply (simp add: card_Suc_Diff1 del:card_Diff_insert) done lemma card_Diff2_less: "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A" apply (case_tac "x = y") apply (simp add: card_Diff1_less del:card_Diff_insert) apply (rule less_trans) prefer 2 apply (auto intro!: card_Diff1_less simp del:card_Diff_insert) done lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A" apply (case_tac "x : A") apply (simp_all add: card_Diff1_less less_imp_le) done lemma card_psubset: "finite B ==> A \ B ==> card A < card B ==> A < B" by (erule psubsetI, blast) lemma insert_partition: "\ x \ F; \c1 \ insert x F. \c2 \ insert x F. c1 \ c2 \ c1 \ c2 = {} \ \ x \ \ F = {}" by auto text{* main cardinality theorem *} lemma card_partition [rule_format]: "finite C ==> finite (\ C) --> (\c\C. card c = k) --> (\c1 \ C. \c2 \ C. c1 \ c2 --> c1 \ c2 = {}) --> k * card(C) = card (\ C)" apply (erule finite_induct, simp) apply (simp add: card_insert_disjoint card_Un_disjoint insert_partition finite_subset [of _ "\ (insert x F)"]) done text{*The form of a finite set of given cardinality*} lemma card_eq_SucD: assumes "card A = Suc k" shows "\b B. A = insert b B & b \ B & card B = k & (k=0 \ B={})" proof - have fin: "finite A" using assms by (auto intro: ccontr) moreover have "card A \ 0" using assms by auto ultimately obtain b where b: "b \ A" by auto show ?thesis proof (intro exI conjI) show "A = insert b (A-{b})" using b by blast show "b \ A - {b}" by blast show "card (A - {b}) = k" and "k = 0 \ A - {b} = {}" using assms b fin by(fastsimp dest:mk_disjoint_insert)+ qed qed lemma card_Suc_eq: "(card A = Suc k) = (\b B. A = insert b B & b \ B & card B = k & (k=0 \ B={}))" apply(rule iffI) apply(erule card_eq_SucD) apply(auto) apply(subst card_insert) apply(auto intro:ccontr) done lemma setsum_constant [simp]: "(\x \ A. y) = of_nat(card A) * y" apply (cases "finite A") apply (erule finite_induct) apply (auto simp add: algebra_simps) done lemma setprod_constant: "finite A ==> (\x\ A. (y::'a::{recpower, comm_monoid_mult})) = y^(card A)" apply (erule finite_induct) apply (auto simp add: power_Suc) done lemma setprod_gen_delta: assumes fS: "finite S" shows "setprod (\k. if k=a then b k else c) S = (if a \ S then (b a ::'a::{comm_monoid_mult, recpower}) * c^ (card S - 1) else c^ card S)" proof- let ?f = "(\k. if k=a then b k else c)" {assume a: "a \ S" hence "\ k\ S. ?f k = c" by simp hence ?thesis using a setprod_constant[OF fS, of c] by (simp add: setprod_1 cong add: setprod_cong) } moreover {assume a: "a \ S" let ?A = "S - {a}" let ?B = "{a}" have eq: "S = ?A \ ?B" using a by blast have dj: "?A \ ?B = {}" by simp from fS have fAB: "finite ?A" "finite ?B" by auto have fA0:"setprod ?f ?A = setprod (\i. c) ?A" apply (rule setprod_cong) by auto have cA: "card ?A = card S - 1" using fS a by auto have fA1: "setprod ?f ?A = c ^ card ?A" unfolding fA0 apply (rule setprod_constant) using fS by auto have "setprod ?f ?A * setprod ?f ?B = setprod ?f S" using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]] by simp then have ?thesis using a cA by (simp add: fA1 ring_simps cong add: setprod_cong cong del: if_weak_cong)} ultimately show ?thesis by blast qed lemma setsum_bounded: assumes le: "\i. i\A \ f i \ (K::'a::{semiring_1, pordered_ab_semigroup_add})" shows "setsum f A \ of_nat(card A) * K" proof (cases "finite A") case True thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp next case False thus ?thesis by (simp add: setsum_def) qed subsubsection {* Cardinality of unions *} lemma card_UN_disjoint: "finite I ==> (ALL i:I. finite (A i)) ==> (ALL i:I. ALL j:I. i \ j --> A i Int A j = {}) ==> card (UNION I A) = (\i\I. card(A i))" apply (simp add: card_def del: setsum_constant) apply (subgoal_tac "setsum (%i. card (A i)) I = setsum (%i. (setsum (%x. 1) (A i))) I") apply (simp add: setsum_UN_disjoint del: setsum_constant) apply (simp cong: setsum_cong) done lemma card_Union_disjoint: "finite C ==> (ALL A:C. finite A) ==> (ALL A:C. ALL B:C. A \ B --> A Int B = {}) ==> card (Union C) = setsum card C" apply (frule card_UN_disjoint [of C id]) apply (unfold Union_def id_def, assumption+) done subsubsection {* Cardinality of image *} text{*The image of a finite set can be expressed using @{term fold_image}.*} lemma image_eq_fold_image: "finite A ==> f ` A = fold_image (op Un) (%x. {f x}) {} A" proof (induct rule: finite_induct) case empty then show ?case by simp next interpret ab_semigroup_mult "op Un" proof qed auto case insert then show ?case by simp qed lemma card_image_le: "finite A ==> card (f ` A) <= card A" apply (induct set: finite) apply simp apply (simp add: le_SucI finite_imageI card_insert_if) done lemma card_image: "inj_on f A ==> card (f ` A) = card A" by(simp add:card_def setsum_reindex o_def del:setsum_constant) lemma endo_inj_surj: "finite A ==> f ` A \ A ==> inj_on f A ==> f ` A = A" by (simp add: card_seteq card_image) lemma eq_card_imp_inj_on: "[| finite A; card(f ` A) = card A |] ==> inj_on f A" apply (induct rule:finite_induct) apply simp apply(frule card_image_le[where f = f]) apply(simp add:card_insert_if split:if_splits) done lemma inj_on_iff_eq_card: "finite A ==> inj_on f A = (card(f ` A) = card A)" by(blast intro: card_image eq_card_imp_inj_on) lemma card_inj_on_le: "[|inj_on f A; f ` A \ B; finite B |] ==> card A \ card B" apply (subgoal_tac "finite A") apply (force intro: card_mono simp add: card_image [symmetric]) apply (blast intro: finite_imageD dest: finite_subset) done lemma card_bij_eq: "[|inj_on f A; f ` A \ B; inj_on g B; g ` B \ A; finite A; finite B |] ==> card A = card B" by (auto intro: le_anti_sym card_inj_on_le) subsubsection {* Cardinality of products *} (* lemma SigmaI_insert: "y \ A ==> (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \ (SIGMA x: A. B x))" by auto *) lemma card_SigmaI [simp]: "\ finite A; ALL a:A. finite (B a) \ \ card (SIGMA x: A. B x) = (\a\A. card (B a))" by(simp add:card_def setsum_Sigma del:setsum_constant) lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)" apply (cases "finite A") apply (cases "finite B") apply (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2) done lemma card_cartesian_product_singleton: "card({x} <*> A) = card(A)" by (simp add: card_cartesian_product) subsubsection {* Cardinality of sums *} lemma card_Plus: assumes "finite A" and "finite B" shows "card (A <+> B) = card A + card B" proof - have "Inl`A \ Inr`B = {}" by fast with assms show ?thesis unfolding Plus_def by (simp add: card_Un_disjoint card_image) qed subsubsection {* Cardinality of the Powerset *} lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A" (* FIXME numeral 2 (!?) *) apply (induct set: finite) apply (simp_all add: Pow_insert) apply (subst card_Un_disjoint, blast) apply (blast intro: finite_imageI, blast) apply (subgoal_tac "inj_on (insert x) (Pow F)") apply (simp add: card_image Pow_insert) apply (unfold inj_on_def) apply (blast elim!: equalityE) done text {* Relates to equivalence classes. Based on a theorem of F. Kammüller. *} lemma dvd_partition: "finite (Union C) ==> ALL c : C. k dvd card c ==> (ALL c1: C. ALL c2: C. c1 \ c2 --> c1 Int c2 = {}) ==> k dvd card (Union C)" apply(frule finite_UnionD) apply(rotate_tac -1) apply (induct set: finite, simp_all, clarify) apply (subst card_Un_disjoint) apply (auto simp add: dvd_add disjoint_eq_subset_Compl) done subsubsection {* Relating injectivity and surjectivity *} lemma finite_surj_inj: "finite(A) \ A <= f`A \ inj_on f A" apply(rule eq_card_imp_inj_on, assumption) apply(frule finite_imageI) apply(drule (1) card_seteq) apply(erule card_image_le) apply simp done lemma finite_UNIV_surj_inj: fixes f :: "'a \ 'a" shows "finite(UNIV:: 'a set) \ surj f \ inj f" by (blast intro: finite_surj_inj subset_UNIV dest:surj_range) lemma finite_UNIV_inj_surj: fixes f :: "'a \ 'a" shows "finite(UNIV:: 'a set) \ inj f \ surj f" by(fastsimp simp:surj_def dest!: endo_inj_surj) corollary infinite_UNIV_nat: "~finite(UNIV::nat set)" proof assume "finite(UNIV::nat set)" with finite_UNIV_inj_surj[of Suc] show False by simp (blast dest: Suc_neq_Zero surjD) qed lemma infinite_UNIV_char_0: "\ finite (UNIV::'a::semiring_char_0 set)" proof assume "finite (UNIV::'a set)" with subset_UNIV have "finite (range of_nat::'a set)" by (rule finite_subset) moreover have "inj (of_nat::nat \ 'a)" by (simp add: inj_on_def) ultimately have "finite (UNIV::nat set)" by (rule finite_imageD) then show "False" by (simp add: infinite_UNIV_nat) qed subsection{* A fold functional for non-empty sets *} text{* Does not require start value. *} inductive fold1Set :: "('a => 'a => 'a) => 'a set => 'a => bool" for f :: "'a => 'a => 'a" where fold1Set_insertI [intro]: "\ fold_graph f a A x; a \ A \ \ fold1Set f (insert a A) x" constdefs fold1 :: "('a => 'a => 'a) => 'a set => 'a" "fold1 f A == THE x. fold1Set f A x" lemma fold1Set_nonempty: "fold1Set f A x \ A \ {}" by(erule fold1Set.cases, simp_all) inductive_cases empty_fold1SetE [elim!]: "fold1Set f {} x" inductive_cases insert_fold1SetE [elim!]: "fold1Set f (insert a X) x" lemma fold1Set_sing [iff]: "(fold1Set f {a} b) = (a = b)" by (blast intro: fold_graph.intros elim: fold_graph.cases) lemma fold1_singleton [simp]: "fold1 f {a} = a" by (unfold fold1_def) blast lemma finite_nonempty_imp_fold1Set: "\ finite A; A \ {} \ \ EX x. fold1Set f A x" apply (induct A rule: finite_induct) apply (auto dest: finite_imp_fold_graph [of _ f]) done text{*First, some lemmas about @{const fold_graph}.*} context ab_semigroup_mult begin lemma fun_left_comm: "fun_left_comm(op *)" by unfold_locales (simp add: mult_ac) lemma fold_graph_insert_swap: assumes fold: "fold_graph times (b::'a) A y" and "b \ A" shows "fold_graph times z (insert b A) (z * y)" proof - interpret fun_left_comm "op *::'a \ 'a \ 'a" by (rule fun_left_comm) from assms show ?thesis proof (induct rule: fold_graph.induct) case emptyI thus ?case by (force simp add: fold_insert_aux mult_commute) next case (insertI x A y) have "fold_graph times z (insert x (insert b A)) (x * (z * y))" using insertI by force --{*how does @{term id} get unfolded?*} thus ?case by (simp add: insert_commute mult_ac) qed qed lemma fold_graph_permute_diff: assumes fold: "fold_graph times b A x" shows "!!a. \a \ A; b \ A\ \ fold_graph times a (insert b (A-{a})) x" using fold proof (induct rule: fold_graph.induct) case emptyI thus ?case by simp next case (insertI x A y) have "a = x \ a \ A" using insertI by simp thus ?case proof assume "a = x" with insertI show ?thesis by (simp add: id_def [symmetric], blast intro: fold_graph_insert_swap) next assume ainA: "a \ A" hence "fold_graph times a (insert x (insert b (A - {a}))) (x * y)" using insertI by force moreover have "insert x (insert b (A - {a})) = insert b (insert x A - {a})" using ainA insertI by blast ultimately show ?thesis by simp qed qed lemma fold1_eq_fold: assumes "finite A" "a \ A" shows "fold1 times (insert a A) = fold times a A" proof - interpret fun_left_comm "op *::'a \ 'a \ 'a" by (rule fun_left_comm) from assms show ?thesis apply (simp add: fold1_def fold_def) apply (rule the_equality) apply (best intro: fold_graph_determ theI dest: finite_imp_fold_graph [of _ times]) apply (rule sym, clarify) apply (case_tac "Aa=A") apply (best intro: the_equality fold_graph_determ) apply (subgoal_tac "fold_graph times a A x") apply (best intro: the_equality fold_graph_determ) apply (subgoal_tac "insert aa (Aa - {a}) = A") prefer 2 apply (blast elim: equalityE) apply (auto dest: fold_graph_permute_diff [where a=a]) done qed lemma nonempty_iff: "(A \ {}) = (\x B. A = insert x B & x \ B)" apply safe apply simp apply (drule_tac x=x in spec) apply (drule_tac x="A-{x}" in spec, auto) done lemma fold1_insert: assumes nonempty: "A \ {}" and A: "finite A" "x \ A" shows "fold1 times (insert x A) = x * fold1 times A" proof - interpret fun_left_comm "op *::'a \ 'a \ 'a" by (rule fun_left_comm) from nonempty obtain a A' where "A = insert a A' & a ~: A'" by (auto simp add: nonempty_iff) with A show ?thesis by (simp add: insert_commute [of x] fold1_eq_fold eq_commute) qed end context ab_semigroup_idem_mult begin lemma fun_left_comm_idem: "fun_left_comm_idem(op *)" apply unfold_locales apply (simp add: mult_ac) apply (simp add: mult_idem mult_assoc[symmetric]) done lemma fold1_insert_idem [simp]: assumes nonempty: "A \ {}" and A: "finite A" shows "fold1 times (insert x A) = x * fold1 times A" proof - interpret fun_left_comm_idem "op *::'a \ 'a \ 'a" by (rule fun_left_comm_idem) from nonempty obtain a A' where A': "A = insert a A' & a ~: A'" by (auto simp add: nonempty_iff) show ?thesis proof cases assume "a = x" thus ?thesis proof cases assume "A' = {}" with prems show ?thesis by (simp add: mult_idem) next assume "A' \ {}" with prems show ?thesis by (simp add: fold1_insert mult_assoc [symmetric] mult_idem) qed next assume "a \ x" with prems show ?thesis by (simp add: insert_commute fold1_eq_fold fold_insert_idem) qed qed lemma hom_fold1_commute: assumes hom: "!!x y. h (x * y) = h x * h y" and N: "finite N" "N \ {}" shows "h (fold1 times N) = fold1 times (h ` N)" using N proof (induct rule: finite_ne_induct) case singleton thus ?case by simp next case (insert n N) then have "h (fold1 times (insert n N)) = h (n * fold1 times N)" by simp also have "\ = h n * h (fold1 times N)" by(rule hom) also have "h (fold1 times N) = fold1 times (h ` N)" by(rule insert) also have "times (h n) \ = fold1 times (insert (h n) (h ` N))" using insert by(simp) also have "insert (h n) (h ` N) = h ` insert n N" by simp finally show ?case . qed end text{* Now the recursion rules for definitions: *} lemma fold1_singleton_def: "g = fold1 f \ g {a} = a" by(simp add:fold1_singleton) lemma (in ab_semigroup_mult) fold1_insert_def: "\ g = fold1 times; finite A; x \ A; A \ {} \ \ g (insert x A) = x * g A" by (simp add:fold1_insert) lemma (in ab_semigroup_idem_mult) fold1_insert_idem_def: "\ g = fold1 times; finite A; A \ {} \ \ g (insert x A) = x * g A" by simp subsubsection{* Determinacy for @{term fold1Set} *} (*Not actually used!!*) (* context ab_semigroup_mult begin lemma fold_graph_permute: "[|fold_graph times id b (insert a A) x; a \ A; b \ A|] ==> fold_graph times id a (insert b A) x" apply (cases "a=b") apply (auto dest: fold_graph_permute_diff) done lemma fold1Set_determ: "fold1Set times A x ==> fold1Set times A y ==> y = x" proof (clarify elim!: fold1Set.cases) fix A x B y a b assume Ax: "fold_graph times id a A x" assume By: "fold_graph times id b B y" assume anotA: "a \ A" assume bnotB: "b \ B" assume eq: "insert a A = insert b B" show "y=x" proof cases assume same: "a=b" hence "A=B" using anotA bnotB eq by (blast elim!: equalityE) thus ?thesis using Ax By same by (blast intro: fold_graph_determ) next assume diff: "a\b" let ?D = "B - {a}" have B: "B = insert a ?D" and A: "A = insert b ?D" and aB: "a \ B" and bA: "b \ A" using eq anotA bnotB diff by (blast elim!:equalityE)+ with aB bnotB By have "fold_graph times id a (insert b ?D) y" by (auto intro: fold_graph_permute simp add: insert_absorb) moreover have "fold_graph times id a (insert b ?D) x" by (simp add: A [symmetric] Ax) ultimately show ?thesis by (blast intro: fold_graph_determ) qed qed lemma fold1Set_equality: "fold1Set times A y ==> fold1 times A = y" by (unfold fold1_def) (blast intro: fold1Set_determ) end *) declare empty_fold_graphE [rule del] fold_graph.intros [rule del] empty_fold1SetE [rule del] insert_fold1SetE [rule del] -- {* No more proofs involve these relations. *} subsubsection {* Lemmas about @{text fold1} *} context ab_semigroup_mult begin lemma fold1_Un: assumes A: "finite A" "A \ {}" shows "finite B \ B \ {} \ A Int B = {} \ fold1 times (A Un B) = fold1 times A * fold1 times B" using A by (induct rule: finite_ne_induct) (simp_all add: fold1_insert mult_assoc) lemma fold1_in: assumes A: "finite (A)" "A \ {}" and elem: "\x y. x * y \ {x,y}" shows "fold1 times A \ A" using A proof (induct rule:finite_ne_induct) case singleton thus ?case by simp next case insert thus ?case using elem by (force simp add:fold1_insert) qed end lemma (in ab_semigroup_idem_mult) fold1_Un2: assumes A: "finite A" "A \ {}" shows "finite B \ B \ {} \ fold1 times (A Un B) = fold1 times A * fold1 times B" using A proof(induct rule:finite_ne_induct) case singleton thus ?case by simp next case insert thus ?case by (simp add: mult_assoc) qed subsubsection {* Fold1 in lattices with @{const inf} and @{const sup} *} text{* As an application of @{text fold1} we define infimum and supremum in (not necessarily complete!) lattices over (non-empty) sets by means of @{text fold1}. *} context lower_semilattice begin lemma ab_semigroup_idem_mult_inf: "ab_semigroup_idem_mult inf" proof qed (rule inf_assoc inf_commute inf_idem)+ lemma below_fold1_iff: assumes "finite A" "A \ {}" shows "x \ fold1 inf A \ (\a\A. x \ a)" proof - interpret ab_semigroup_idem_mult inf by (rule ab_semigroup_idem_mult_inf) show ?thesis using assms by (induct rule: finite_ne_induct) simp_all qed lemma fold1_belowI: assumes "finite A" and "a \ A" shows "fold1 inf A \ a" proof - from assms have "A \ {}" by auto from `finite A` `A \ {}` `a \ A` show ?thesis proof (induct rule: finite_ne_induct) case singleton thus ?case by simp next interpret ab_semigroup_idem_mult inf by (rule ab_semigroup_idem_mult_inf) case (insert x F) from insert(5) have "a = x \ a \ F" by simp thus ?case proof assume "a = x" thus ?thesis using insert by (simp add: mult_ac) next assume "a \ F" hence bel: "fold1 inf F \ a" by (rule insert) have "inf (fold1 inf (insert x F)) a = inf x (inf (fold1 inf F) a)" using insert by (simp add: mult_ac) also have "inf (fold1 inf F) a = fold1 inf F" using bel by (auto intro: antisym) also have "inf x \ = fold1 inf (insert x F)" using insert by (simp add: mult_ac) finally have aux: "inf (fold1 inf (insert x F)) a = fold1 inf (insert x F)" . moreover have "inf (fold1 inf (insert x F)) a \ a" by simp ultimately show ?thesis by simp qed qed qed end lemma (in upper_semilattice) ab_semigroup_idem_mult_sup: "ab_semigroup_idem_mult sup" by (rule lower_semilattice.ab_semigroup_idem_mult_inf) (rule dual_lattice) context lattice begin definition Inf_fin :: "'a set \ 'a" ("\\<^bsub>fin\<^esub>_" [900] 900) where "Inf_fin = fold1 inf" definition Sup_fin :: "'a set \ 'a" ("\\<^bsub>fin\<^esub>_" [900] 900) where "Sup_fin = fold1 sup" lemma Inf_le_Sup [simp]: "\ finite A; A \ {} \ \ \\<^bsub>fin\<^esub>A \ \\<^bsub>fin\<^esub>A" apply(unfold Sup_fin_def Inf_fin_def) apply(subgoal_tac "EX a. a:A") prefer 2 apply blast apply(erule exE) apply(rule order_trans) apply(erule (1) fold1_belowI) apply(erule (1) lower_semilattice.fold1_belowI [OF dual_lattice]) done lemma sup_Inf_absorb [simp]: "finite A \ a \ A \ sup a (\\<^bsub>fin\<^esub>A) = a" apply(subst sup_commute) apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI) done lemma inf_Sup_absorb [simp]: "finite A \ a \ A \ inf a (\\<^bsub>fin\<^esub>A) = a" by (simp add: Sup_fin_def inf_absorb1 lower_semilattice.fold1_belowI [OF dual_lattice]) end context distrib_lattice begin lemma sup_Inf1_distrib: assumes "finite A" and "A \ {}" shows "sup x (\\<^bsub>fin\<^esub>A) = \\<^bsub>fin\<^esub>{sup x a|a. a \ A}" proof - interpret ab_semigroup_idem_mult inf by (rule ab_semigroup_idem_mult_inf) from assms show ?thesis by (simp add: Inf_fin_def image_def hom_fold1_commute [where h="sup x", OF sup_inf_distrib1]) (rule arg_cong [where f="fold1 inf"], blast) qed lemma sup_Inf2_distrib: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" shows "sup (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B) = \\<^bsub>fin\<^esub>{sup a b|a b. a \ A \ b \ B}" using A proof (induct rule: finite_ne_induct) case singleton thus ?case by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def]) next interpret ab_semigroup_idem_mult inf by (rule ab_semigroup_idem_mult_inf) case (insert x A) have finB: "finite {sup x b |b. b \ B}" by(rule finite_surj[where f = "sup x", OF B(1)], auto) have finAB: "finite {sup a b |a b. a \ A \ b \ B}" proof - have "{sup a b |a b. a \ A \ b \ B} = (UN a:A. UN b:B. {sup a b})" by blast thus ?thesis by(simp add: insert(1) B(1)) qed have ne: "{sup a b |a b. a \ A \ b \ B} \ {}" using insert B by blast have "sup (\\<^bsub>fin\<^esub>(insert x A)) (\\<^bsub>fin\<^esub>B) = sup (inf x (\\<^bsub>fin\<^esub>A)) (\\<^bsub>fin\<^esub>B)" using insert by (simp add: fold1_insert_idem_def [OF Inf_fin_def]) also have "\ = inf (sup x (\\<^bsub>fin\<^esub>B)) (sup (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2) also have "\ = inf (\\<^bsub>fin\<^esub>{sup x b|b. b \ B}) (\\<^bsub>fin\<^esub>{sup a b|a b. a \ A \ b \ B})" using insert by(simp add:sup_Inf1_distrib[OF B]) also have "\ = \\<^bsub>fin\<^esub>({sup x b |b. b \ B} \ {sup a b |a b. a \ A \ b \ B})" (is "_ = \\<^bsub>fin\<^esub>?M") using B insert by (simp add: Inf_fin_def fold1_Un2 [OF finB _ finAB ne]) also have "?M = {sup a b |a b. a \ insert x A \ b \ B}" by blast finally show ?case . qed lemma inf_Sup1_distrib: assumes "finite A" and "A \ {}" shows "inf x (\\<^bsub>fin\<^esub>A) = \\<^bsub>fin\<^esub>{inf x a|a. a \ A}" proof - interpret ab_semigroup_idem_mult sup by (rule ab_semigroup_idem_mult_sup) from assms show ?thesis by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1]) (rule arg_cong [where f="fold1 sup"], blast) qed lemma inf_Sup2_distrib: assumes A: "finite A" "A \ {}" and B: "finite B" "B \ {}" shows "inf (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B) = \\<^bsub>fin\<^esub>{inf a b|a b. a \ A \ b \ B}" using A proof (induct rule: finite_ne_induct) case singleton thus ?case by(simp add: inf_Sup1_distrib [OF B] fold1_singleton_def [OF Sup_fin_def]) next case (insert x A) have finB: "finite {inf x b |b. b \ B}" by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto) have finAB: "finite {inf a b |a b. a \ A \ b \ B}" proof - have "{inf a b |a b. a \ A \ b \ B} = (UN a:A. UN b:B. {inf a b})" by blast thus ?thesis by(simp add: insert(1) B(1)) qed have ne: "{inf a b |a b. a \ A \ b \ B} \ {}" using insert B by blast interpret ab_semigroup_idem_mult sup by (rule ab_semigroup_idem_mult_sup) have "inf (\\<^bsub>fin\<^esub>(insert x A)) (\\<^bsub>fin\<^esub>B) = inf (sup x (\\<^bsub>fin\<^esub>A)) (\\<^bsub>fin\<^esub>B)" using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def]) also have "\ = sup (inf x (\\<^bsub>fin\<^esub>B)) (inf (\\<^bsub>fin\<^esub>A) (\\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2) also have "\ = sup (\\<^bsub>fin\<^esub>{inf x b|b. b \ B}) (\\<^bsub>fin\<^esub>{inf a b|a b. a \ A \ b \ B})" using insert by(simp add:inf_Sup1_distrib[OF B]) also have "\ = \\<^bsub>fin\<^esub>({inf x b |b. b \ B} \ {inf a b |a b. a \ A \ b \ B})" (is "_ = \\<^bsub>fin\<^esub>?M") using B insert by (simp add: Sup_fin_def fold1_Un2 [OF finB _ finAB ne]) also have "?M = {inf a b |a b. a \ insert x A \ b \ B}" by blast finally show ?case . qed end context complete_lattice begin text {* Coincidence on finite sets in complete lattices: *} lemma Inf_fin_Inf: assumes "finite A" and "A \ {}" shows "\\<^bsub>fin\<^esub>A = Inf A" proof - interpret ab_semigroup_idem_mult inf by (rule ab_semigroup_idem_mult_inf) from assms show ?thesis unfolding Inf_fin_def by (induct A set: finite) (simp_all add: Inf_insert_simp) qed lemma Sup_fin_Sup: assumes "finite A" and "A \ {}" shows "\\<^bsub>fin\<^esub>A = Sup A" proof - interpret ab_semigroup_idem_mult sup by (rule ab_semigroup_idem_mult_sup) from assms show ?thesis unfolding Sup_fin_def by (induct A set: finite) (simp_all add: Sup_insert_simp) qed end subsubsection {* Fold1 in linear orders with @{const min} and @{const max} *} text{* As an application of @{text fold1} we define minimum and maximum in (not necessarily complete!) linear orders over (non-empty) sets by means of @{text fold1}. *} context linorder begin lemma ab_semigroup_idem_mult_min: "ab_semigroup_idem_mult min" proof qed (auto simp add: min_def) lemma ab_semigroup_idem_mult_max: "ab_semigroup_idem_mult max" proof qed (auto simp add: max_def) lemma min_lattice: "lower_semilattice (op \) (op <) min" proof qed (auto simp add: min_def) lemma max_lattice: "lower_semilattice (op \) (op >) max" proof qed (auto simp add: max_def) lemma dual_max: "ord.max (op \) = min" by (auto simp add: ord.max_def_raw min_def_raw expand_fun_eq) lemma dual_min: "ord.min (op \) = max" by (auto simp add: ord.min_def_raw max_def_raw expand_fun_eq) lemma strict_below_fold1_iff: assumes "finite A" and "A \ {}" shows "x < fold1 min A \ (\a\A. x < a)" proof - interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) from assms show ?thesis by (induct rule: finite_ne_induct) (simp_all add: fold1_insert) qed lemma fold1_below_iff: assumes "finite A" and "A \ {}" shows "fold1 min A \ x \ (\a\A. a \ x)" proof - interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) from assms show ?thesis by (induct rule: finite_ne_induct) (simp_all add: fold1_insert min_le_iff_disj) qed lemma fold1_strict_below_iff: assumes "finite A" and "A \ {}" shows "fold1 min A < x \ (\a\A. a < x)" proof - interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) from assms show ?thesis by (induct rule: finite_ne_induct) (simp_all add: fold1_insert min_less_iff_disj) qed lemma fold1_antimono: assumes "A \ {}" and "A \ B" and "finite B" shows "fold1 min B \ fold1 min A" proof cases assume "A = B" thus ?thesis by simp next interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) assume "A \ B" have B: "B = A \ (B-A)" using `A \ B` by blast have "fold1 min B = fold1 min (A \ (B-A))" by(subst B)(rule refl) also have "\ = min (fold1 min A) (fold1 min (B-A))" proof - have "finite A" by(rule finite_subset[OF `A \ B` `finite B`]) moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`]) (* by(blast intro:finite_Diff prems) fails *) moreover have "(B-A) \ {}" using prems by blast moreover have "A Int (B-A) = {}" using prems by blast ultimately show ?thesis using `A \ {}` by (rule_tac fold1_Un) qed also have "\ \ fold1 min A" by (simp add: min_le_iff_disj) finally show ?thesis . qed definition Min :: "'a set \ 'a" where "Min = fold1 min" definition Max :: "'a set \ 'a" where "Max = fold1 max" lemmas Min_singleton [simp] = fold1_singleton_def [OF Min_def] lemmas Max_singleton [simp] = fold1_singleton_def [OF Max_def] lemma Min_insert [simp]: assumes "finite A" and "A \ {}" shows "Min (insert x A) = min x (Min A)" proof - interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) from assms show ?thesis by (rule fold1_insert_idem_def [OF Min_def]) qed lemma Max_insert [simp]: assumes "finite A" and "A \ {}" shows "Max (insert x A) = max x (Max A)" proof - interpret ab_semigroup_idem_mult max by (rule ab_semigroup_idem_mult_max) from assms show ?thesis by (rule fold1_insert_idem_def [OF Max_def]) qed lemma Min_in [simp]: assumes "finite A" and "A \ {}" shows "Min A \ A" proof - interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) from assms fold1_in show ?thesis by (fastsimp simp: Min_def min_def) qed lemma Max_in [simp]: assumes "finite A" and "A \ {}" shows "Max A \ A" proof - interpret ab_semigroup_idem_mult max by (rule ab_semigroup_idem_mult_max) from assms fold1_in [of A] show ?thesis by (fastsimp simp: Max_def max_def) qed lemma Min_Un: assumes "finite A" and "A \ {}" and "finite B" and "B \ {}" shows "Min (A \ B) = min (Min A) (Min B)" proof - interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) from assms show ?thesis by (simp add: Min_def fold1_Un2) qed lemma Max_Un: assumes "finite A" and "A \ {}" and "finite B" and "B \ {}" shows "Max (A \ B) = max (Max A) (Max B)" proof - interpret ab_semigroup_idem_mult max by (rule ab_semigroup_idem_mult_max) from assms show ?thesis by (simp add: Max_def fold1_Un2) qed lemma hom_Min_commute: assumes "\x y. h (min x y) = min (h x) (h y)" and "finite N" and "N \ {}" shows "h (Min N) = Min (h ` N)" proof - interpret ab_semigroup_idem_mult min by (rule ab_semigroup_idem_mult_min) from assms show ?thesis by (simp add: Min_def hom_fold1_commute) qed lemma hom_Max_commute: assumes "\x y. h (max x y) = max (h x) (h y)" and "finite N" and "N \ {}" shows "h (Max N) = Max (h ` N)" proof - interpret ab_semigroup_idem_mult max by (rule ab_semigroup_idem_mult_max) from assms show ?thesis by (simp add: Max_def hom_fold1_commute [of h]) qed lemma Min_le [simp]: assumes "finite A" and "x \ A" shows "Min A \ x" proof - interpret lower_semilattice "op \" "op <" min by (rule min_lattice) from assms show ?thesis by (simp add: Min_def fold1_belowI) qed lemma Max_ge [simp]: assumes "finite A" and "x \ A" shows "x \ Max A" proof - interpret lower_semilattice "op \" "op >" max by (rule max_lattice) from assms show ?thesis by (simp add: Max_def fold1_belowI) qed lemma Min_ge_iff [simp, noatp]: assumes "finite A" and "A \ {}" shows "x \ Min A \ (\a\A. x \ a)" proof - interpret lower_semilattice "op \" "op <" min by (rule min_lattice) from assms show ?thesis by (simp add: Min_def below_fold1_iff) qed lemma Max_le_iff [simp, noatp]: assumes "finite A" and "A \ {}" shows "Max A \ x \ (\a\A. a \ x)" proof - interpret lower_semilattice "op \" "op >" max by (rule max_lattice) from assms show ?thesis by (simp add: Max_def below_fold1_iff) qed lemma Min_gr_iff [simp, noatp]: assumes "finite A" and "A \ {}" shows "x < Min A \ (\a\A. x < a)" proof - interpret lower_semilattice "op \" "op <" min by (rule min_lattice) from assms show ?thesis by (simp add: Min_def strict_below_fold1_iff) qed lemma Max_less_iff [simp, noatp]: assumes "finite A" and "A \ {}" shows "Max A < x \ (\a\A. a < x)" proof - note Max = Max_def interpret linorder "op \" "op >" by (rule dual_linorder) from assms show ?thesis by (simp add: Max strict_below_fold1_iff [folded dual_max]) qed lemma Min_le_iff [noatp]: assumes "finite A" and "A \ {}" shows "Min A \ x \ (\a\A. a \ x)" proof - interpret lower_semilattice "op \" "op <" min by (rule min_lattice) from assms show ?thesis by (simp add: Min_def fold1_below_iff) qed lemma Max_ge_iff [noatp]: assumes "finite A" and "A \ {}" shows "x \ Max A \ (\a\A. x \ a)" proof - note Max = Max_def interpret linorder "op \" "op >" by (rule dual_linorder) from assms show ?thesis by (simp add: Max fold1_below_iff [folded dual_max]) qed lemma Min_less_iff [noatp]: assumes "finite A" and "A \ {}" shows "Min A < x \ (\a\A. a < x)" proof - interpret lower_semilattice "op \" "op <" min by (rule min_lattice) from assms show ?thesis by (simp add: Min_def fold1_strict_below_iff) qed lemma Max_gr_iff [noatp]: assumes "finite A" and "A \ {}" shows "x < Max A \ (\a\A. x < a)" proof - note Max = Max_def interpret linorder "op \" "op >" by (rule dual_linorder) from assms show ?thesis by (simp add: Max fold1_strict_below_iff [folded dual_max]) qed lemma Min_eqI: assumes "finite A" assumes "\y. y \ A \ y \ x" and "x \ A" shows "Min A = x" proof (rule antisym) from `x \ A` have "A \ {}" by auto with assms show "Min A \ x" by simp next from assms show "x \ Min A" by simp qed lemma Max_eqI: assumes "finite A" assumes "\y. y \ A \ y \ x" and "x \ A" shows "Max A = x" proof (rule antisym) from `x \ A` have "A \ {}" by auto with assms show "Max A \ x" by simp next from assms show "x \ Max A" by simp qed lemma Min_antimono: assumes "M \ N" and "M \ {}" and "finite N" shows "Min N \ Min M" proof - interpret distrib_lattice "op \" "op <" min max by (rule distrib_lattice_min_max) from assms show ?thesis by (simp add: Min_def fold1_antimono) qed lemma Max_mono: assumes "M \ N" and "M \ {}" and "finite N" shows "Max M \ Max N" proof - note Max = Max_def interpret linorder "op \" "op >" by (rule dual_linorder) from assms show ?thesis by (simp add: Max fold1_antimono [folded dual_max]) qed lemma finite_linorder_induct[consumes 1, case_names empty insert]: "finite A \ P {} \ (!!A b. finite A \ ALL a:A. a < b \ P A \ P(insert b A)) \ P A" proof (induct A rule: measure_induct_rule[where f=card]) fix A :: "'a set" assume IH: "!! B. card B < card A \ finite B \ P {} \ (!!A b. finite A \ (\a\A. a P A \ P (insert b A)) \ P B" and "finite A" and "P {}" and step: "!!A b. \finite A; \a\A. a < b; P A\ \ P (insert b A)" show "P A" proof (cases "A = {}") assume "A = {}" thus "P A" using `P {}` by simp next let ?B = "A - {Max A}" let ?A = "insert (Max A) ?B" assume "A \ {}" with `finite A` have "Max A : A" by auto hence A: "?A = A" using insert_Diff_single insert_absorb by auto note card_Diff1_less[OF `finite A` `Max A : A`] moreover have "finite ?B" using `finite A` by simp ultimately have "P ?B" using `P {}` step IH by blast moreover have "\a\?B. a < Max A" using Max_ge [OF `finite A`] by fastsimp ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastsimp qed qed end context ordered_ab_semigroup_add begin lemma add_Min_commute: fixes k assumes "finite N" and "N \ {}" shows "k + Min N = Min {k + m | m. m \ N}" proof - have "\x y. k + min x y = min (k + x) (k + y)" by (simp add: min_def not_le) (blast intro: antisym less_imp_le add_left_mono) with assms show ?thesis using hom_Min_commute [of "plus k" N] by simp (blast intro: arg_cong [where f = Min]) qed lemma add_Max_commute: fixes k assumes "finite N" and "N \ {}" shows "k + Max N = Max {k + m | m. m \ N}" proof - have "\x y. k + max x y = max (k + x) (k + y)" by (simp add: max_def not_le) (blast intro: antisym less_imp_le add_left_mono) with assms show ?thesis using hom_Max_commute [of "plus k" N] by simp (blast intro: arg_cong [where f = Max]) qed end end diff --git a/src/HOL/Groebner_Basis.thy b/src/HOL/Groebner_Basis.thy --- a/src/HOL/Groebner_Basis.thy +++ b/src/HOL/Groebner_Basis.thy @@ -1,671 +1,671 @@ (* Title: HOL/Groebner_Basis.thy Author: Amine Chaieb, TU Muenchen *) header {* Semiring normalization and Groebner Bases *} theory Groebner_Basis imports NatBin uses "Tools/Groebner_Basis/misc.ML" "Tools/Groebner_Basis/normalizer_data.ML" ("Tools/Groebner_Basis/normalizer.ML") ("Tools/Groebner_Basis/groebner.ML") begin subsection {* Semiring normalization *} setup NormalizerData.setup locale gb_semiring = fixes add mul pwr r0 r1 assumes add_a:"(add x (add y z) = add (add x y) z)" and add_c: "add x y = add y x" and add_0:"add r0 x = x" and mul_a:"mul x (mul y z) = mul (mul x y) z" and mul_c:"mul x y = mul y x" and mul_1:"mul r1 x = x" and mul_0:"mul r0 x = r0" and mul_d:"mul x (add y z) = add (mul x y) (mul x z)" and pwr_0:"pwr x 0 = r1" and pwr_Suc:"pwr x (Suc n) = mul x (pwr x n)" begin lemma mul_pwr:"mul (pwr x p) (pwr x q) = pwr x (p + q)" proof (induct p) case 0 then show ?case by (auto simp add: pwr_0 mul_1) next case Suc from this [symmetric] show ?case by (auto simp add: pwr_Suc mul_1 mul_a) qed lemma pwr_mul: "pwr (mul x y) q = mul (pwr x q) (pwr y q)" proof (induct q arbitrary: x y, auto simp add:pwr_0 pwr_Suc mul_1) fix q x y assume "\x y. pwr (mul x y) q = mul (pwr x q) (pwr y q)" have "mul (mul x y) (mul (pwr x q) (pwr y q)) = mul x (mul y (mul (pwr x q) (pwr y q)))" by (simp add: mul_a) also have "\ = (mul (mul y (mul (pwr y q) (pwr x q))) x)" by (simp add: mul_c) also have "\ = (mul (mul y (pwr y q)) (mul (pwr x q) x))" by (simp add: mul_a) finally show "mul (mul x y) (mul (pwr x q) (pwr y q)) = mul (mul x (pwr x q)) (mul y (pwr y q))" by (simp add: mul_c) qed lemma pwr_pwr: "pwr (pwr x p) q = pwr x (p * q)" proof (induct p arbitrary: q) case 0 show ?case using pwr_Suc mul_1 pwr_0 by (induct q) auto next case Suc thus ?case by (auto simp add: mul_pwr [symmetric] pwr_mul pwr_Suc) qed subsubsection {* Declaring the abstract theory *} lemma semiring_ops: shows "TERM (add x y)" and "TERM (mul x y)" and "TERM (pwr x n)" and "TERM r0" and "TERM r1" . lemma semiring_rules: "add (mul a m) (mul b m) = mul (add a b) m" "add (mul a m) m = mul (add a r1) m" "add m (mul a m) = mul (add a r1) m" "add m m = mul (add r1 r1) m" "add r0 a = a" "add a r0 = a" "mul a b = mul b a" "mul (add a b) c = add (mul a c) (mul b c)" "mul r0 a = r0" "mul a r0 = r0" "mul r1 a = a" "mul a r1 = a" "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)" "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))" "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)" "mul (mul lx ly) rx = mul (mul lx rx) ly" "mul (mul lx ly) rx = mul lx (mul ly rx)" "mul lx (mul rx ry) = mul (mul lx rx) ry" "mul lx (mul rx ry) = mul rx (mul lx ry)" "add (add a b) (add c d) = add (add a c) (add b d)" "add (add a b) c = add a (add b c)" "add a (add c d) = add c (add a d)" "add (add a b) c = add (add a c) b" "add a c = add c a" "add a (add c d) = add (add a c) d" "mul (pwr x p) (pwr x q) = pwr x (p + q)" "mul x (pwr x q) = pwr x (Suc q)" "mul (pwr x q) x = pwr x (Suc q)" "mul x x = pwr x 2" "pwr (mul x y) q = mul (pwr x q) (pwr y q)" "pwr (pwr x p) q = pwr x (p * q)" "pwr x 0 = r1" "pwr x 1 = x" "mul x (add y z) = add (mul x y) (mul x z)" "pwr x (Suc q) = mul x (pwr x q)" "pwr x (2*n) = mul (pwr x n) (pwr x n)" "pwr x (Suc (2*n)) = mul x (mul (pwr x n) (pwr x n))" proof - show "add (mul a m) (mul b m) = mul (add a b) m" using mul_d mul_c by simp next show"add (mul a m) m = mul (add a r1) m" using mul_d mul_c mul_1 by simp next show "add m (mul a m) = mul (add a r1) m" using mul_c mul_d mul_1 add_c by simp next show "add m m = mul (add r1 r1) m" using mul_c mul_d mul_1 by simp next show "add r0 a = a" using add_0 by simp next show "add a r0 = a" using add_0 add_c by simp next show "mul a b = mul b a" using mul_c by simp next show "mul (add a b) c = add (mul a c) (mul b c)" using mul_c mul_d by simp next show "mul r0 a = r0" using mul_0 by simp next show "mul a r0 = r0" using mul_0 mul_c by simp next show "mul r1 a = a" using mul_1 by simp next show "mul a r1 = a" using mul_1 mul_c by simp next show "mul (mul lx ly) (mul rx ry) = mul (mul lx rx) (mul ly ry)" using mul_c mul_a by simp next show "mul (mul lx ly) (mul rx ry) = mul lx (mul ly (mul rx ry))" using mul_a by simp next have "mul (mul lx ly) (mul rx ry) = mul (mul rx ry) (mul lx ly)" by (rule mul_c) also have "\ = mul rx (mul ry (mul lx ly))" using mul_a by simp finally show "mul (mul lx ly) (mul rx ry) = mul rx (mul (mul lx ly) ry)" using mul_c by simp next show "mul (mul lx ly) rx = mul (mul lx rx) ly" using mul_c mul_a by simp next show "mul (mul lx ly) rx = mul lx (mul ly rx)" by (simp add: mul_a) next show "mul lx (mul rx ry) = mul (mul lx rx) ry" by (simp add: mul_a ) next show "mul lx (mul rx ry) = mul rx (mul lx ry)" by (simp add: mul_a,simp add: mul_c) next show "add (add a b) (add c d) = add (add a c) (add b d)" using add_c add_a by simp next show "add (add a b) c = add a (add b c)" using add_a by simp next show "add a (add c d) = add c (add a d)" apply (simp add: add_a) by (simp only: add_c) next show "add (add a b) c = add (add a c) b" using add_a add_c by simp next show "add a c = add c a" by (rule add_c) next show "add a (add c d) = add (add a c) d" using add_a by simp next show "mul (pwr x p) (pwr x q) = pwr x (p + q)" by (rule mul_pwr) next show "mul x (pwr x q) = pwr x (Suc q)" using pwr_Suc by simp next show "mul (pwr x q) x = pwr x (Suc q)" using pwr_Suc mul_c by simp next show "mul x x = pwr x 2" by (simp add: nat_number pwr_Suc pwr_0 mul_1 mul_c) next show "pwr (mul x y) q = mul (pwr x q) (pwr y q)" by (rule pwr_mul) next show "pwr (pwr x p) q = pwr x (p * q)" by (rule pwr_pwr) next show "pwr x 0 = r1" using pwr_0 . next show "pwr x 1 = x" unfolding One_nat_def by (simp add: nat_number pwr_Suc pwr_0 mul_1 mul_c) next show "mul x (add y z) = add (mul x y) (mul x z)" using mul_d by simp next show "pwr x (Suc q) = mul x (pwr x q)" using pwr_Suc by simp next show "pwr x (2 * n) = mul (pwr x n) (pwr x n)" by (simp add: nat_number mul_pwr) next show "pwr x (Suc (2 * n)) = mul x (mul (pwr x n) (pwr x n))" by (simp add: nat_number pwr_Suc mul_pwr) qed lemmas gb_semiring_axioms' = gb_semiring_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules] end -interpretation class_semiring!: gb_semiring +interpretation class_semiring: gb_semiring "op +" "op *" "op ^" "0::'a::{comm_semiring_1, recpower}" "1" proof qed (auto simp add: algebra_simps power_Suc) lemmas nat_arith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of lemma not_iszero_Numeral1: "\ iszero (Numeral1::'a::number_ring)" by (simp add: numeral_1_eq_1) lemmas comp_arith = Let_def arith_simps nat_arith rel_simps neg_simps if_False if_True add_0 add_Suc add_number_of_left mult_number_of_left numeral_1_eq_1[symmetric] Suc_eq_add_numeral_1 numeral_0_eq_0[symmetric] numerals[symmetric] iszero_simps not_iszero_Numeral1 lemmas semiring_norm = comp_arith ML {* local open Conv; fun numeral_is_const ct = can HOLogic.dest_number (Thm.term_of ct); fun int_of_rat x = (case Rat.quotient_of_rat x of (i, 1) => i | _ => error "int_of_rat: bad int"); val numeral_conv = Simplifier.rewrite (HOL_basic_ss addsimps @{thms semiring_norm}) then_conv Simplifier.rewrite (HOL_basic_ss addsimps (@{thms numeral_1_eq_1} @ @{thms numeral_0_eq_0} @ @{thms numerals(1-2)})); in fun normalizer_funs key = NormalizerData.funs key {is_const = fn phi => numeral_is_const, dest_const = fn phi => fn ct => Rat.rat_of_int (snd (HOLogic.dest_number (Thm.term_of ct) handle TERM _ => error "ring_dest_const")), mk_const = fn phi => fn cT => fn x => Numeral.mk_cnumber cT (int_of_rat x), conv = fn phi => K numeral_conv} end *} declaration {* normalizer_funs @{thm class_semiring.gb_semiring_axioms'} *} locale gb_ring = gb_semiring + fixes sub :: "'a \ 'a \ 'a" and neg :: "'a \ 'a" assumes neg_mul: "neg x = mul (neg r1) x" and sub_add: "sub x y = add x (neg y)" begin lemma ring_ops: shows "TERM (sub x y)" and "TERM (neg x)" . lemmas ring_rules = neg_mul sub_add lemmas gb_ring_axioms' = gb_ring_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules ring ops: ring_ops ring rules: ring_rules] end -interpretation class_ring!: gb_ring "op +" "op *" "op ^" +interpretation class_ring: gb_ring "op +" "op *" "op ^" "0::'a::{comm_semiring_1,recpower,number_ring}" 1 "op -" "uminus" proof qed simp_all declaration {* normalizer_funs @{thm class_ring.gb_ring_axioms'} *} use "Tools/Groebner_Basis/normalizer.ML" method_setup sring_norm = {* Scan.succeed (SIMPLE_METHOD' o Normalizer.semiring_normalize_tac) *} "semiring normalizer" locale gb_field = gb_ring + fixes divide :: "'a \ 'a \ 'a" and inverse:: "'a \ 'a" assumes divide: "divide x y = mul x (inverse y)" and inverse: "inverse x = divide r1 x" begin lemmas gb_field_axioms' = gb_field_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules ring ops: ring_ops ring rules: ring_rules] end subsection {* Groebner Bases *} locale semiringb = gb_semiring + assumes add_cancel: "add (x::'a) y = add x z \ y = z" and add_mul_solve: "add (mul w y) (mul x z) = add (mul w z) (mul x y) \ w = x \ y = z" begin lemma noteq_reduce: "a \ b \ c \ d \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" proof- have "a \ b \ c \ d \ \ (a = b \ c = d)" by simp also have "\ \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" using add_mul_solve by blast finally show "a \ b \ c \ d \ add (mul a c) (mul b d) \ add (mul a d) (mul b c)" by simp qed lemma add_scale_eq_noteq: "\r \ r0 ; (a = b) \ ~(c = d)\ \ add a (mul r c) \ add b (mul r d)" proof(clarify) assume nz: "r\ r0" and cnd: "c\d" and eq: "add b (mul r c) = add b (mul r d)" hence "mul r c = mul r d" using cnd add_cancel by simp hence "add (mul r0 d) (mul r c) = add (mul r0 c) (mul r d)" using mul_0 add_cancel by simp thus "False" using add_mul_solve nz cnd by simp qed lemma add_r0_iff: " x = add x a \ a = r0" proof- have "a = r0 \ add x a = add x r0" by (simp add: add_cancel) thus "x = add x a \ a = r0" by (auto simp add: add_c add_0) qed declare gb_semiring_axioms' [normalizer del] lemmas semiringb_axioms' = semiringb_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules idom rules: noteq_reduce add_scale_eq_noteq] end locale ringb = semiringb + gb_ring + assumes subr0_iff: "sub x y = r0 \ x = y" begin declare gb_ring_axioms' [normalizer del] lemmas ringb_axioms' = ringb_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules ring ops: ring_ops ring rules: ring_rules idom rules: noteq_reduce add_scale_eq_noteq ideal rules: subr0_iff add_r0_iff] end lemma no_zero_divirors_neq0: assumes az: "(a::'a::no_zero_divisors) \ 0" and ab: "a*b = 0" shows "b = 0" proof - { assume bz: "b \ 0" from no_zero_divisors [OF az bz] ab have False by blast } thus "b = 0" by blast qed -interpretation class_ringb!: ringb +interpretation class_ringb: ringb "op +" "op *" "op ^" "0::'a::{idom,recpower,number_ring}" "1" "op -" "uminus" proof(unfold_locales, simp add: algebra_simps power_Suc, auto) fix w x y z ::"'a::{idom,recpower,number_ring}" assume p: "w * y + x * z = w * z + x * y" and ynz: "y \ z" hence ynz': "y - z \ 0" by simp from p have "w * y + x* z - w*z - x*y = 0" by simp hence "w* (y - z) - x * (y - z) = 0" by (simp add: algebra_simps) hence "(y - z) * (w - x) = 0" by (simp add: algebra_simps) with no_zero_divirors_neq0 [OF ynz'] have "w - x = 0" by blast thus "w = x" by simp qed declaration {* normalizer_funs @{thm class_ringb.ringb_axioms'} *} -interpretation natgb!: semiringb +interpretation natgb: semiringb "op +" "op *" "op ^" "0::nat" "1" proof (unfold_locales, simp add: algebra_simps power_Suc) fix w x y z ::"nat" { assume p: "w * y + x * z = w * z + x * y" and ynz: "y \ z" hence "y < z \ y > z" by arith moreover { assume lt:"y k. z = y + k \ k > 0" by (rule_tac x="z - y" in exI, auto) then obtain k where kp: "k>0" and yz:"z = y + k" by blast from p have "(w * y + x *y) + x*k = (w * y + x*y) + w*k" by (simp add: yz algebra_simps) hence "x*k = w*k" by simp hence "w = x" using kp by (simp add: mult_cancel2) } moreover { assume lt: "y >z" hence "\k. y = z + k \ k>0" by (rule_tac x="y - z" in exI, auto) then obtain k where kp: "k>0" and yz:"y = z + k" by blast from p have "(w * z + x *z) + w*k = (w * z + x*z) + x*k" by (simp add: yz algebra_simps) hence "w*k = x*k" by simp hence "w = x" using kp by (simp add: mult_cancel2)} ultimately have "w=x" by blast } thus "(w * y + x * z = w * z + x * y) = (w = x \ y = z)" by auto qed declaration {* normalizer_funs @{thm natgb.semiringb_axioms'} *} locale fieldgb = ringb + gb_field begin declare gb_field_axioms' [normalizer del] lemmas fieldgb_axioms' = fieldgb_axioms [normalizer semiring ops: semiring_ops semiring rules: semiring_rules ring ops: ring_ops ring rules: ring_rules idom rules: noteq_reduce add_scale_eq_noteq ideal rules: subr0_iff add_r0_iff] end lemmas bool_simps = simp_thms(1-34) lemma dnf: "(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))" "(P \ Q) = (Q \ P)" "(P \ Q) = (Q \ P)" by blast+ lemmas weak_dnf_simps = dnf bool_simps lemma nnf_simps: "(\(P \ Q)) = (\P \ \Q)" "(\(P \ Q)) = (\P \ \Q)" "(P \ Q) = (\P \ Q)" "(P = Q) = ((P \ Q) \ (\P \ \ Q))" "(\ \(P)) = P" by blast+ lemma PFalse: "P \ False \ \ P" "\ P \ (P \ False)" by auto use "Tools/Groebner_Basis/groebner.ML" method_setup algebra = {* let fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () val addN = "add" val delN = "del" val any_keyword = keyword addN || keyword delN val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; in ((Scan.optional (keyword addN |-- thms) []) -- (Scan.optional (keyword delN |-- thms) [])) >> (fn (add_ths, del_ths) => fn ctxt => SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt)) end *} "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases" declare dvd_def[algebra] declare dvd_eq_mod_eq_0[symmetric, algebra] declare mod_div_trivial[algebra] declare mod_mod_trivial[algebra] declare conjunct1[OF DIVISION_BY_ZERO, algebra] declare conjunct2[OF DIVISION_BY_ZERO, algebra] declare zmod_zdiv_equality[symmetric,algebra] declare zdiv_zmod_equality[symmetric, algebra] declare zdiv_zminus_zminus[algebra] declare zmod_zminus_zminus[algebra] declare zdiv_zminus2[algebra] declare zmod_zminus2[algebra] declare zdiv_zero[algebra] declare zmod_zero[algebra] declare mod_by_1[algebra] declare div_by_1[algebra] declare zmod_minus1_right[algebra] declare zdiv_minus1_right[algebra] declare mod_div_trivial[algebra] declare mod_mod_trivial[algebra] declare mod_mult_self2_is_0[algebra] declare mod_mult_self1_is_0[algebra] declare zmod_eq_0_iff[algebra] declare dvd_0_left_iff[algebra] declare zdvd1_eq[algebra] declare zmod_eq_dvd_iff[algebra] declare nat_mod_eq_iff[algebra] subsection{* Groebner Bases for fields *} -interpretation class_fieldgb!: +interpretation class_fieldgb: fieldgb "op +" "op *" "op ^" "0::'a::{field,recpower,number_ring}" "1" "op -" "uminus" "op /" "inverse" apply (unfold_locales) by (simp_all add: divide_inverse) lemma divide_Numeral1: "(x::'a::{field,number_ring}) / Numeral1 = x" by simp lemma divide_Numeral0: "(x::'a::{field,number_ring, division_by_zero}) / Numeral0 = 0" by simp lemma mult_frac_frac: "((x::'a::{field,division_by_zero}) / y) * (z / w) = (x*z) / (y*w)" by simp lemma mult_frac_num: "((x::'a::{field, division_by_zero}) / y) * z = (x*z) / y" by simp lemma mult_num_frac: "((x::'a::{field, division_by_zero}) / y) * z = (x*z) / y" by simp lemma Numeral1_eq1_nat: "(1::nat) = Numeral1" by simp lemma add_frac_num: "y\ 0 \ (x::'a::{field, division_by_zero}) / y + z = (x + z*y) / y" by (simp add: add_divide_distrib) lemma add_num_frac: "y\ 0 \ z + (x::'a::{field, division_by_zero}) / y = (x + z*y) / y" by (simp add: add_divide_distrib) ML{* local val zr = @{cpat "0"} val zT = ctyp_of_term zr val geq = @{cpat "op ="} val eqT = Thm.dest_ctyp (ctyp_of_term geq) |> hd val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"} val add_frac_num = mk_meta_eq @{thm "add_frac_num"} val add_num_frac = mk_meta_eq @{thm "add_num_frac"} fun prove_nz ss T t = let val z = instantiate_cterm ([(zT,T)],[]) zr val eq = instantiate_cterm ([(eqT,T)],[]) geq val th = Simplifier.rewrite (ss addsimps simp_thms) (Thm.capply @{cterm "Trueprop"} (Thm.capply @{cterm "Not"} (Thm.capply (Thm.capply eq t) z))) in equal_elim (symmetric th) TrueI end fun proc phi ss ct = let val ((x,y),(w,z)) = (Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct val _ = map (HOLogic.dest_number o term_of) [x,y,z,w] val T = ctyp_of_term x val [y_nz, z_nz] = map (prove_nz ss T) [y, z] val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq in SOME (implies_elim (implies_elim th y_nz) z_nz) end handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE fun proc2 phi ss ct = let val (l,r) = Thm.dest_binop ct val T = ctyp_of_term l in (case (term_of l, term_of r) of (Const(@{const_name "HOL.divide"},_)$_$_, _) => let val (x,y) = Thm.dest_binop l val z = r val _ = map (HOLogic.dest_number o term_of) [x,y,z] val ynz = prove_nz ss T y in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz) end | (_, Const (@{const_name "HOL.divide"},_)$_$_) => let val (x,y) = Thm.dest_binop r val z = l val _ = map (HOLogic.dest_number o term_of) [x,y,z] val ynz = prove_nz ss T y in SOME (implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz) end | _ => NONE) end handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE fun is_number (Const(@{const_name "HOL.divide"},_)$a$b) = is_number a andalso is_number b | is_number t = can HOLogic.dest_number t val is_number = is_number o term_of fun proc3 phi ss ct = (case term_of ct of Const(@{const_name HOL.less},_)$(Const(@{const_name "HOL.divide"},_)$_$_)$_ => let val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop val _ = map is_number [a,b,c] val T = ctyp_of_term c val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_less_eq"} in SOME (mk_meta_eq th) end | Const(@{const_name HOL.less_eq},_)$(Const(@{const_name "HOL.divide"},_)$_$_)$_ => let val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop val _ = map is_number [a,b,c] val T = ctyp_of_term c val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"} in SOME (mk_meta_eq th) end | Const("op =",_)$(Const(@{const_name "HOL.divide"},_)$_$_)$_ => let val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop val _ = map is_number [a,b,c] val T = ctyp_of_term c val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_eq_eq"} in SOME (mk_meta_eq th) end | Const(@{const_name HOL.less},_)$_$(Const(@{const_name "HOL.divide"},_)$_$_) => let val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop val _ = map is_number [a,b,c] val T = ctyp_of_term c val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "less_divide_eq"} in SOME (mk_meta_eq th) end | Const(@{const_name HOL.less_eq},_)$_$(Const(@{const_name "HOL.divide"},_)$_$_) => let val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop val _ = map is_number [a,b,c] val T = ctyp_of_term c val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"} in SOME (mk_meta_eq th) end | Const("op =",_)$_$(Const(@{const_name "HOL.divide"},_)$_$_) => let val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop val _ = map is_number [a,b,c] val T = ctyp_of_term c val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "eq_divide_eq"} in SOME (mk_meta_eq th) end | _ => NONE) handle TERM _ => NONE | CTERM _ => NONE | THM _ => NONE val add_frac_frac_simproc = make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + (?w::?'a::field)/?z"}], name = "add_frac_frac_simproc", proc = proc, identifier = []} val add_frac_num_simproc = make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + ?z"}, @{cpat "?z + (?x::?'a::field)/?y"}], name = "add_frac_num_simproc", proc = proc2, identifier = []} val ord_frac_simproc = make_simproc {lhss = [@{cpat "(?a::(?'a::{field, ord}))/?b < ?c"}, @{cpat "(?a::(?'a::{field, ord}))/?b \ ?c"}, @{cpat "?c < (?a::(?'a::{field, ord}))/?b"}, @{cpat "?c \ (?a::(?'a::{field, ord}))/?b"}, @{cpat "?c = ((?a::(?'a::{field, ord}))/?b)"}, @{cpat "((?a::(?'a::{field, ord}))/ ?b) = ?c"}], name = "ord_frac_simproc", proc = proc3, identifier = []} val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, @{thm "divide_Numeral1"}, @{thm "Ring_and_Field.divide_zero"}, @{thm "divide_Numeral0"}, @{thm "divide_divide_eq_left"}, @{thm "mult_frac_frac"}, @{thm "mult_num_frac"}, @{thm "mult_frac_num"}, @{thm "mult_frac_frac"}, @{thm "times_divide_eq_right"}, @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"}, @{thm "diff_def"}, @{thm "minus_divide_left"}, @{thm "Numeral1_eq1_nat"}, @{thm "add_divide_distrib"} RS sym] local open Conv in val comp_conv = (Simplifier.rewrite (HOL_basic_ss addsimps @{thms "Groebner_Basis.comp_arith"} addsimps ths addsimps simp_thms addsimprocs field_cancel_numeral_factors addsimprocs [add_frac_frac_simproc, add_frac_num_simproc, ord_frac_simproc] addcongs [@{thm "if_weak_cong"}])) then_conv (Simplifier.rewrite (HOL_basic_ss addsimps [@{thm numeral_1_eq_1},@{thm numeral_0_eq_0}] @ @{thms numerals(1-2)})) end fun numeral_is_const ct = case term_of ct of Const (@{const_name "HOL.divide"},_) $ a $ b => numeral_is_const (Thm.dest_arg1 ct) andalso numeral_is_const (Thm.dest_arg ct) | Const (@{const_name "HOL.uminus"},_)$t => numeral_is_const (Thm.dest_arg ct) | t => can HOLogic.dest_number t fun dest_const ct = ((case term_of ct of Const (@{const_name "HOL.divide"},_) $ a $ b=> Rat.rat_of_quotient (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b)) | t => Rat.rat_of_int (snd (HOLogic.dest_number t))) handle TERM _ => error "ring_dest_const") fun mk_const phi cT x = let val (a, b) = Rat.quotient_of_rat x in if b = 1 then Numeral.mk_cnumber cT a else Thm.capply (Thm.capply (Drule.cterm_rule (instantiate' [SOME cT] []) @{cpat "op /"}) (Numeral.mk_cnumber cT a)) (Numeral.mk_cnumber cT b) end in val field_comp_conv = comp_conv; val fieldgb_declaration = NormalizerData.funs @{thm class_fieldgb.fieldgb_axioms'} {is_const = K numeral_is_const, dest_const = K dest_const, mk_const = mk_const, conv = K (K comp_conv)} end; *} declaration fieldgb_declaration end diff --git a/src/HOL/HahnBanach/Subspace.thy b/src/HOL/HahnBanach/Subspace.thy --- a/src/HOL/HahnBanach/Subspace.thy +++ b/src/HOL/HahnBanach/Subspace.thy @@ -1,513 +1,513 @@ (* Title: HOL/Real/HahnBanach/Subspace.thy Author: Gertrud Bauer, TU Munich *) header {* Subspaces *} theory Subspace imports VectorSpace begin subsection {* Definition *} text {* A non-empty subset @{text U} of a vector space @{text V} is a \emph{subspace} of @{text V}, iff @{text U} is closed under addition and scalar multiplication. *} locale subspace = fixes U :: "'a\{minus, plus, zero, uminus} set" and V assumes non_empty [iff, intro]: "U \ {}" and subset [iff]: "U \ V" and add_closed [iff]: "x \ U \ y \ U \ x + y \ U" and mult_closed [iff]: "x \ U \ a \ x \ U" notation (symbols) subspace (infix "\" 50) declare vectorspace.intro [intro?] subspace.intro [intro?] lemma subspace_subset [elim]: "U \ V \ U \ V" by (rule subspace.subset) lemma (in subspace) subsetD [iff]: "x \ U \ x \ V" using subset by blast lemma subspaceD [elim]: "U \ V \ x \ U \ x \ V" by (rule subspace.subsetD) lemma rev_subspaceD [elim?]: "x \ U \ U \ V \ x \ V" by (rule subspace.subsetD) lemma (in subspace) diff_closed [iff]: assumes "vectorspace V" assumes x: "x \ U" and y: "y \ U" shows "x - y \ U" proof - interpret vectorspace V by fact from x y show ?thesis by (simp add: diff_eq1 negate_eq1) qed text {* \medskip Similar as for linear spaces, the existence of the zero element in every subspace follows from the non-emptiness of the carrier set and by vector space laws. *} lemma (in subspace) zero [intro]: assumes "vectorspace V" shows "0 \ U" proof - - interpret V!: vectorspace V by fact + interpret V: vectorspace V by fact have "U \ {}" by (rule non_empty) then obtain x where x: "x \ U" by blast then have "x \ V" .. then have "0 = x - x" by simp also from `vectorspace V` x x have "\ \ U" by (rule diff_closed) finally show ?thesis . qed lemma (in subspace) neg_closed [iff]: assumes "vectorspace V" assumes x: "x \ U" shows "- x \ U" proof - interpret vectorspace V by fact from x show ?thesis by (simp add: negate_eq1) qed text {* \medskip Further derived laws: every subspace is a vector space. *} lemma (in subspace) vectorspace [iff]: assumes "vectorspace V" shows "vectorspace U" proof - interpret vectorspace V by fact show ?thesis proof show "U \ {}" .. fix x y z assume x: "x \ U" and y: "y \ U" and z: "z \ U" fix a b :: real from x y show "x + y \ U" by simp from x show "a \ x \ U" by simp from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac) from x y show "x + y = y + x" by (simp add: add_ac) from x show "x - x = 0" by simp from x show "0 + x = x" by simp from x y show "a \ (x + y) = a \ x + a \ y" by (simp add: distrib) from x show "(a + b) \ x = a \ x + b \ x" by (simp add: distrib) from x show "(a * b) \ x = a \ b \ x" by (simp add: mult_assoc) from x show "1 \ x = x" by simp from x show "- x = - 1 \ x" by (simp add: negate_eq1) from x y show "x - y = x + - y" by (simp add: diff_eq1) qed qed text {* The subspace relation is reflexive. *} lemma (in vectorspace) subspace_refl [intro]: "V \ V" proof show "V \ {}" .. show "V \ V" .. fix x y assume x: "x \ V" and y: "y \ V" fix a :: real from x y show "x + y \ V" by simp from x show "a \ x \ V" by simp qed text {* The subspace relation is transitive. *} lemma (in vectorspace) subspace_trans [trans]: "U \ V \ V \ W \ U \ W" proof assume uv: "U \ V" and vw: "V \ W" from uv show "U \ {}" by (rule subspace.non_empty) show "U \ W" proof - from uv have "U \ V" by (rule subspace.subset) also from vw have "V \ W" by (rule subspace.subset) finally show ?thesis . qed fix x y assume x: "x \ U" and y: "y \ U" from uv and x y show "x + y \ U" by (rule subspace.add_closed) from uv and x show "\a. a \ x \ U" by (rule subspace.mult_closed) qed subsection {* Linear closure *} text {* The \emph{linear closure} of a vector @{text x} is the set of all scalar multiples of @{text x}. *} definition lin :: "('a::{minus, plus, zero}) \ 'a set" where "lin x = {a \ x | a. True}" lemma linI [intro]: "y = a \ x \ y \ lin x" unfolding lin_def by blast lemma linI' [iff]: "a \ x \ lin x" unfolding lin_def by blast lemma linE [elim]: "x \ lin v \ (\a::real. x = a \ v \ C) \ C" unfolding lin_def by blast text {* Every vector is contained in its linear closure. *} lemma (in vectorspace) x_lin_x [iff]: "x \ V \ x \ lin x" proof - assume "x \ V" then have "x = 1 \ x" by simp also have "\ \ lin x" .. finally show ?thesis . qed lemma (in vectorspace) "0_lin_x" [iff]: "x \ V \ 0 \ lin x" proof assume "x \ V" then show "0 = 0 \ x" by simp qed text {* Any linear closure is a subspace. *} lemma (in vectorspace) lin_subspace [intro]: "x \ V \ lin x \ V" proof assume x: "x \ V" then show "lin x \ {}" by (auto simp add: x_lin_x) show "lin x \ V" proof fix x' assume "x' \ lin x" then obtain a where "x' = a \ x" .. with x show "x' \ V" by simp qed fix x' x'' assume x': "x' \ lin x" and x'': "x'' \ lin x" show "x' + x'' \ lin x" proof - from x' obtain a' where "x' = a' \ x" .. moreover from x'' obtain a'' where "x'' = a'' \ x" .. ultimately have "x' + x'' = (a' + a'') \ x" using x by (simp add: distrib) also have "\ \ lin x" .. finally show ?thesis . qed fix a :: real show "a \ x' \ lin x" proof - from x' obtain a' where "x' = a' \ x" .. with x have "a \ x' = (a * a') \ x" by (simp add: mult_assoc) also have "\ \ lin x" .. finally show ?thesis . qed qed text {* Any linear closure is a vector space. *} lemma (in vectorspace) lin_vectorspace [intro]: assumes "x \ V" shows "vectorspace (lin x)" proof - from `x \ V` have "subspace (lin x) V" by (rule lin_subspace) from this and vectorspace_axioms show ?thesis by (rule subspace.vectorspace) qed subsection {* Sum of two vectorspaces *} text {* The \emph{sum} of two vectorspaces @{text U} and @{text V} is the set of all sums of elements from @{text U} and @{text V}. *} instantiation "fun" :: (type, type) plus begin definition sum_def: "plus_fun U V = {u + v | u v. u \ U \ v \ V}" (* FIXME not fully general!? *) instance .. end lemma sumE [elim]: "x \ U + V \ (\u v. x = u + v \ u \ U \ v \ V \ C) \ C" unfolding sum_def by blast lemma sumI [intro]: "u \ U \ v \ V \ x = u + v \ x \ U + V" unfolding sum_def by blast lemma sumI' [intro]: "u \ U \ v \ V \ u + v \ U + V" unfolding sum_def by blast text {* @{text U} is a subspace of @{text "U + V"}. *} lemma subspace_sum1 [iff]: assumes "vectorspace U" "vectorspace V" shows "U \ U + V" proof - interpret vectorspace U by fact interpret vectorspace V by fact show ?thesis proof show "U \ {}" .. show "U \ U + V" proof fix x assume x: "x \ U" moreover have "0 \ V" .. ultimately have "x + 0 \ U + V" .. with x show "x \ U + V" by simp qed fix x y assume x: "x \ U" and "y \ U" then show "x + y \ U" by simp from x show "\a. a \ x \ U" by simp qed qed text {* The sum of two subspaces is again a subspace. *} lemma sum_subspace [intro?]: assumes "subspace U E" "vectorspace E" "subspace V E" shows "U + V \ E" proof - interpret subspace U E by fact interpret vectorspace E by fact interpret subspace V E by fact show ?thesis proof have "0 \ U + V" proof show "0 \ U" using `vectorspace E` .. show "0 \ V" using `vectorspace E` .. show "(0::'a) = 0 + 0" by simp qed then show "U + V \ {}" by blast show "U + V \ E" proof fix x assume "x \ U + V" then obtain u v where "x = u + v" and "u \ U" and "v \ V" .. then show "x \ E" by simp qed fix x y assume x: "x \ U + V" and y: "y \ U + V" show "x + y \ U + V" proof - from x obtain ux vx where "x = ux + vx" and "ux \ U" and "vx \ V" .. moreover from y obtain uy vy where "y = uy + vy" and "uy \ U" and "vy \ V" .. ultimately have "ux + uy \ U" and "vx + vy \ V" and "x + y = (ux + uy) + (vx + vy)" using x y by (simp_all add: add_ac) then show ?thesis .. qed fix a show "a \ x \ U + V" proof - from x obtain u v where "x = u + v" and "u \ U" and "v \ V" .. then have "a \ u \ U" and "a \ v \ V" and "a \ x = (a \ u) + (a \ v)" by (simp_all add: distrib) then show ?thesis .. qed qed qed text{* The sum of two subspaces is a vectorspace. *} lemma sum_vs [intro?]: "U \ E \ V \ E \ vectorspace E \ vectorspace (U + V)" by (rule subspace.vectorspace) (rule sum_subspace) subsection {* Direct sums *} text {* The sum of @{text U} and @{text V} is called \emph{direct}, iff the zero element is the only common element of @{text U} and @{text V}. For every element @{text x} of the direct sum of @{text U} and @{text V} the decomposition in @{text "x = u + v"} with @{text "u \ U"} and @{text "v \ V"} is unique. *} lemma decomp: assumes "vectorspace E" "subspace U E" "subspace V E" assumes direct: "U \ V = {0}" and u1: "u1 \ U" and u2: "u2 \ U" and v1: "v1 \ V" and v2: "v2 \ V" and sum: "u1 + v1 = u2 + v2" shows "u1 = u2 \ v1 = v2" proof - interpret vectorspace E by fact interpret subspace U E by fact interpret subspace V E by fact show ?thesis proof have U: "vectorspace U" (* FIXME: use interpret *) using `subspace U E` `vectorspace E` by (rule subspace.vectorspace) have V: "vectorspace V" using `subspace V E` `vectorspace E` by (rule subspace.vectorspace) from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1" by (simp add: add_diff_swap) from u1 u2 have u: "u1 - u2 \ U" by (rule vectorspace.diff_closed [OF U]) with eq have v': "v2 - v1 \ U" by (simp only:) from v2 v1 have v: "v2 - v1 \ V" by (rule vectorspace.diff_closed [OF V]) with eq have u': " u1 - u2 \ V" by (simp only:) show "u1 = u2" proof (rule add_minus_eq) from u1 show "u1 \ E" .. from u2 show "u2 \ E" .. from u u' and direct show "u1 - u2 = 0" by blast qed show "v1 = v2" proof (rule add_minus_eq [symmetric]) from v1 show "v1 \ E" .. from v2 show "v2 \ E" .. from v v' and direct show "v2 - v1 = 0" by blast qed qed qed text {* An application of the previous lemma will be used in the proof of the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any element @{text "y + a \ x\<^sub>0"} of the direct sum of a vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"} the components @{text "y \ H"} and @{text a} are uniquely determined. *} lemma decomp_H': assumes "vectorspace E" "subspace H E" assumes y1: "y1 \ H" and y2: "y2 \ H" and x': "x' \ H" "x' \ E" "x' \ 0" and eq: "y1 + a1 \ x' = y2 + a2 \ x'" shows "y1 = y2 \ a1 = a2" proof - interpret vectorspace E by fact interpret subspace H E by fact show ?thesis proof have c: "y1 = y2 \ a1 \ x' = a2 \ x'" proof (rule decomp) show "a1 \ x' \ lin x'" .. show "a2 \ x' \ lin x'" .. show "H \ lin x' = {0}" proof show "H \ lin x' \ {0}" proof fix x assume x: "x \ H \ lin x'" then obtain a where xx': "x = a \ x'" by blast have "x = 0" proof cases assume "a = 0" with xx' and x' show ?thesis by simp next assume a: "a \ 0" from x have "x \ H" .. with xx' have "inverse a \ a \ x' \ H" by simp with a and x' have "x' \ H" by (simp add: mult_assoc2) with `x' \ H` show ?thesis by contradiction qed then show "x \ {0}" .. qed show "{0} \ H \ lin x'" proof - have "0 \ H" using `vectorspace E` .. moreover have "0 \ lin x'" using `x' \ E` .. ultimately show ?thesis by blast qed qed show "lin x' \ E" using `x' \ E` .. qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq) then show "y1 = y2" .. from c have "a1 \ x' = a2 \ x'" .. with x' show "a1 = a2" by (simp add: mult_right_cancel) qed qed text {* Since for any element @{text "y + a \ x'"} of the direct sum of a vectorspace @{text H} and the linear closure of @{text x'} the components @{text "y \ H"} and @{text a} are unique, it follows from @{text "y \ H"} that @{text "a = 0"}. *} lemma decomp_H'_H: assumes "vectorspace E" "subspace H E" assumes t: "t \ H" and x': "x' \ H" "x' \ E" "x' \ 0" shows "(SOME (y, a). t = y + a \ x' \ y \ H) = (t, 0)" proof - interpret vectorspace E by fact interpret subspace H E by fact show ?thesis proof (rule, simp_all only: split_paired_all split_conv) from t x' show "t = t + 0 \ x' \ t \ H" by simp fix y and a assume ya: "t = y + a \ x' \ y \ H" have "y = t \ a = 0" proof (rule decomp_H') from ya x' show "y + a \ x' = t + 0 \ x'" by simp from ya show "y \ H" .. qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+) with t x' show "(y, a) = (y + a \ x', 0)" by simp qed qed text {* The components @{text "y \ H"} and @{text a} in @{text "y + a \ x'"} are unique, so the function @{text h'} defined by @{text "h' (y + a \ x') = h y + a \ \"} is definite. *} lemma h'_definite: fixes H assumes h'_def: "h' \ (\x. let (y, a) = SOME (y, a). (x = y + a \ x' \ y \ H) in (h y) + a * xi)" and x: "x = y + a \ x'" assumes "vectorspace E" "subspace H E" assumes y: "y \ H" and x': "x' \ H" "x' \ E" "x' \ 0" shows "h' x = h y + a * xi" proof - interpret vectorspace E by fact interpret subspace H E by fact from x y x' have "x \ H + lin x'" by auto have "\!p. (\(y, a). x = y + a \ x' \ y \ H) p" (is "\!p. ?P p") proof (rule ex_ex1I) from x y show "\p. ?P p" by blast fix p q assume p: "?P p" and q: "?P q" show "p = q" proof - from p have xp: "x = fst p + snd p \ x' \ fst p \ H" by (cases p) simp from q have xq: "x = fst q + snd q \ x' \ fst q \ H" by (cases q) simp have "fst p = fst q \ snd p = snd q" proof (rule decomp_H') from xp show "fst p \ H" .. from xq show "fst q \ H" .. from xp and xq show "fst p + snd p \ x' = fst q + snd q \ x'" by simp qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+) then show ?thesis by (cases p, cases q) simp qed qed then have eq: "(SOME (y, a). x = y + a \ x' \ y \ H) = (y, a)" by (rule some1_equality) (simp add: x y) with h'_def show "h' x = h y + a * xi" by (simp add: Let_def) qed end diff --git a/src/HOL/Lattices.thy b/src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy +++ b/src/HOL/Lattices.thy @@ -1,386 +1,386 @@ (* Title: HOL/Lattices.thy Author: Tobias Nipkow *) header {* Abstract lattices *} theory Lattices imports Orderings begin subsection {* Lattices *} notation less_eq (infix "\" 50) and less (infix "\" 50) class lower_semilattice = order + fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) assumes inf_le1 [simp]: "x \ y \ x" and inf_le2 [simp]: "x \ y \ y" and inf_greatest: "x \ y \ x \ z \ x \ y \ z" class upper_semilattice = order + fixes sup :: "'a \ 'a \ 'a" (infixl "\" 65) assumes sup_ge1 [simp]: "x \ x \ y" and sup_ge2 [simp]: "y \ x \ y" and sup_least: "y \ x \ z \ x \ y \ z \ x" begin text {* Dual lattice *} lemma dual_lattice: "lower_semilattice (op \) (op >) sup" by (rule lower_semilattice.intro, rule dual_order) (unfold_locales, simp_all add: sup_least) end class lattice = lower_semilattice + upper_semilattice subsubsection {* Intro and elim rules*} context lower_semilattice begin lemma le_infI1[intro]: assumes "a \ x" shows "a \ b \ x" proof (rule order_trans) from assms show "a \ x" . show "a \ b \ a" by simp qed lemmas (in -) [rule del] = le_infI1 lemma le_infI2[intro]: assumes "b \ x" shows "a \ b \ x" proof (rule order_trans) from assms show "b \ x" . show "a \ b \ b" by simp qed lemmas (in -) [rule del] = le_infI2 lemma le_infI[intro!]: "x \ a \ x \ b \ x \ a \ b" by(blast intro: inf_greatest) lemmas (in -) [rule del] = le_infI lemma le_infE [elim!]: "x \ a \ b \ (x \ a \ x \ b \ P) \ P" by (blast intro: order_trans) lemmas (in -) [rule del] = le_infE lemma le_inf_iff [simp]: "x \ y \ z = (x \ y \ x \ z)" by blast lemma le_iff_inf: "(x \ y) = (x \ y = x)" by (blast intro: antisym dest: eq_iff [THEN iffD1]) lemma mono_inf: fixes f :: "'a \ 'b\lower_semilattice" shows "mono f \ f (A \ B) \ f A \ f B" by (auto simp add: mono_def intro: Lattices.inf_greatest) end context upper_semilattice begin lemma le_supI1[intro]: "x \ a \ x \ a \ b" by (rule order_trans) auto lemmas (in -) [rule del] = le_supI1 lemma le_supI2[intro]: "x \ b \ x \ a \ b" by (rule order_trans) auto lemmas (in -) [rule del] = le_supI2 lemma le_supI[intro!]: "a \ x \ b \ x \ a \ b \ x" by (blast intro: sup_least) lemmas (in -) [rule del] = le_supI lemma le_supE[elim!]: "a \ b \ x \ (a \ x \ b \ x \ P) \ P" by (blast intro: order_trans) lemmas (in -) [rule del] = le_supE lemma ge_sup_conv[simp]: "x \ y \ z = (x \ z \ y \ z)" by blast lemma le_iff_sup: "(x \ y) = (x \ y = y)" by (blast intro: antisym dest: eq_iff [THEN iffD1]) lemma mono_sup: fixes f :: "'a \ 'b\upper_semilattice" shows "mono f \ f A \ f B \ f (A \ B)" by (auto simp add: mono_def intro: Lattices.sup_least) end subsubsection{* Equational laws *} context lower_semilattice begin lemma inf_commute: "(x \ y) = (y \ x)" by (blast intro: antisym) lemma inf_assoc: "(x \ y) \ z = x \ (y \ z)" by (blast intro: antisym) lemma inf_idem[simp]: "x \ x = x" by (blast intro: antisym) lemma inf_left_idem[simp]: "x \ (x \ y) = x \ y" by (blast intro: antisym) lemma inf_absorb1: "x \ y \ x \ y = x" by (blast intro: antisym) lemma inf_absorb2: "y \ x \ x \ y = y" by (blast intro: antisym) lemma inf_left_commute: "x \ (y \ z) = y \ (x \ z)" by (blast intro: antisym) lemmas inf_ACI = inf_commute inf_assoc inf_left_commute inf_left_idem end context upper_semilattice begin lemma sup_commute: "(x \ y) = (y \ x)" by (blast intro: antisym) lemma sup_assoc: "(x \ y) \ z = x \ (y \ z)" by (blast intro: antisym) lemma sup_idem[simp]: "x \ x = x" by (blast intro: antisym) lemma sup_left_idem[simp]: "x \ (x \ y) = x \ y" by (blast intro: antisym) lemma sup_absorb1: "y \ x \ x \ y = x" by (blast intro: antisym) lemma sup_absorb2: "x \ y \ x \ y = y" by (blast intro: antisym) lemma sup_left_commute: "x \ (y \ z) = y \ (x \ z)" by (blast intro: antisym) lemmas sup_ACI = sup_commute sup_assoc sup_left_commute sup_left_idem end context lattice begin lemma inf_sup_absorb: "x \ (x \ y) = x" by (blast intro: antisym inf_le1 inf_greatest sup_ge1) lemma sup_inf_absorb: "x \ (x \ y) = x" by (blast intro: antisym sup_ge1 sup_least inf_le1) lemmas ACI = inf_ACI sup_ACI lemmas inf_sup_ord = inf_le1 inf_le2 sup_ge1 sup_ge2 text{* Towards distributivity *} lemma distrib_sup_le: "x \ (y \ z) \ (x \ y) \ (x \ z)" by blast lemma distrib_inf_le: "(x \ y) \ (x \ z) \ x \ (y \ z)" by blast text{* If you have one of them, you have them all. *} lemma distrib_imp1: assumes D: "!!x y z. x \ (y \ z) = (x \ y) \ (x \ z)" shows "x \ (y \ z) = (x \ y) \ (x \ z)" proof- have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by(simp add:sup_inf_absorb) also have "\ = x \ (z \ (x \ y))" by(simp add:D inf_commute sup_assoc) also have "\ = ((x \ y) \ x) \ ((x \ y) \ z)" by(simp add:inf_sup_absorb inf_commute) also have "\ = (x \ y) \ (x \ z)" by(simp add:D) finally show ?thesis . qed lemma distrib_imp2: assumes D: "!!x y z. x \ (y \ z) = (x \ y) \ (x \ z)" shows "x \ (y \ z) = (x \ y) \ (x \ z)" proof- have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by(simp add:inf_sup_absorb) also have "\ = x \ (z \ (x \ y))" by(simp add:D sup_commute inf_assoc) also have "\ = ((x \ y) \ x) \ ((x \ y) \ z)" by(simp add:sup_inf_absorb sup_commute) also have "\ = (x \ y) \ (x \ z)" by(simp add:D) finally show ?thesis . qed (* seems unused *) lemma modular_le: "x \ z \ x \ (y \ z) \ (x \ y) \ z" by blast end subsection {* Distributive lattices *} class distrib_lattice = lattice + assumes sup_inf_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" context distrib_lattice begin lemma sup_inf_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" by(simp add:ACI sup_inf_distrib1) lemma inf_sup_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" by(rule distrib_imp2[OF sup_inf_distrib1]) lemma inf_sup_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" by(simp add:ACI inf_sup_distrib1) lemmas distrib = sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 end subsection {* Uniqueness of inf and sup *} lemma (in lower_semilattice) inf_unique: fixes f (infixl "\" 70) assumes le1: "\x y. x \ y \ x" and le2: "\x y. x \ y \ y" and greatest: "\x y z. x \ y \ x \ z \ x \ y \ z" shows "x \ y = x \ y" proof (rule antisym) show "x \ y \ x \ y" by (rule le_infI) (rule le1, rule le2) next have leI: "\x y z. x \ y \ x \ z \ x \ y \ z" by (blast intro: greatest) show "x \ y \ x \ y" by (rule leI) simp_all qed lemma (in upper_semilattice) sup_unique: fixes f (infixl "\" 70) assumes ge1 [simp]: "\x y. x \ x \ y" and ge2: "\x y. y \ x \ y" and least: "\x y z. y \ x \ z \ x \ y \ z \ x" shows "x \ y = x \ y" proof (rule antisym) show "x \ y \ x \ y" by (rule le_supI) (rule ge1, rule ge2) next have leI: "\x y z. x \ z \ y \ z \ x \ y \ z" by (blast intro: least) show "x \ y \ x \ y" by (rule leI) simp_all qed subsection {* @{const min}/@{const max} on linear orders as special case of @{const inf}/@{const sup} *} lemma (in linorder) distrib_lattice_min_max: "distrib_lattice (op \) (op <) min max" proof have aux: "\x y \ 'a. x < y \ y \ x \ x = y" by (auto simp add: less_le antisym) fix x y z show "max x (min y z) = min (max x y) (max x z)" unfolding min_def max_def by auto qed (auto simp add: min_def max_def not_le less_imp_le) -interpretation min_max!: distrib_lattice "op \ :: 'a::linorder \ 'a \ bool" "op <" min max +interpretation min_max: distrib_lattice "op \ :: 'a::linorder \ 'a \ bool" "op <" min max by (rule distrib_lattice_min_max) lemma inf_min: "inf = (min \ 'a\{lower_semilattice, linorder} \ 'a \ 'a)" by (rule ext)+ (auto intro: antisym) lemma sup_max: "sup = (max \ 'a\{upper_semilattice, linorder} \ 'a \ 'a)" by (rule ext)+ (auto intro: antisym) lemmas le_maxI1 = min_max.sup_ge1 lemmas le_maxI2 = min_max.sup_ge2 lemmas max_ac = min_max.sup_assoc min_max.sup_commute mk_left_commute [of max, OF min_max.sup_assoc min_max.sup_commute] lemmas min_ac = min_max.inf_assoc min_max.inf_commute mk_left_commute [of min, OF min_max.inf_assoc min_max.inf_commute] text {* Now we have inherited antisymmetry as an intro-rule on all linear orders. This is a problem because it applies to bool, which is undesirable. *} lemmas [rule del] = min_max.le_infI min_max.le_supI min_max.le_supE min_max.le_infE min_max.le_supI1 min_max.le_supI2 min_max.le_infI1 min_max.le_infI2 subsection {* Bool as lattice *} instantiation bool :: distrib_lattice begin definition inf_bool_eq: "P \ Q \ P \ Q" definition sup_bool_eq: "P \ Q \ P \ Q" instance by intro_classes (auto simp add: inf_bool_eq sup_bool_eq le_bool_def) end subsection {* Fun as lattice *} instantiation "fun" :: (type, lattice) lattice begin definition inf_fun_eq [code del]: "f \ g = (\x. f x \ g x)" definition sup_fun_eq [code del]: "f \ g = (\x. f x \ g x)" instance apply intro_classes unfolding inf_fun_eq sup_fun_eq apply (auto intro: le_funI) apply (rule le_funI) apply (auto dest: le_funD) apply (rule le_funI) apply (auto dest: le_funD) done end instance "fun" :: (type, distrib_lattice) distrib_lattice by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) text {* redundant bindings *} lemmas inf_aci = inf_ACI lemmas sup_aci = sup_ACI no_notation less_eq (infix "\" 50) and less (infix "\" 50) and inf (infixl "\" 70) and sup (infixl "\" 65) end diff --git a/src/HOL/Library/FrechetDeriv.thy b/src/HOL/Library/FrechetDeriv.thy --- a/src/HOL/Library/FrechetDeriv.thy +++ b/src/HOL/Library/FrechetDeriv.thy @@ -1,502 +1,502 @@ (* Title : FrechetDeriv.thy Author : Brian Huffman *) header {* Frechet Derivative *} theory FrechetDeriv imports Lim Complex_Main begin definition fderiv :: "['a::real_normed_vector \ 'b::real_normed_vector, 'a, 'a \ 'b] \ bool" -- {* Frechet derivative: D is derivative of function f at x *} ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "FDERIV f x :> D = (bounded_linear D \ (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)" lemma FDERIV_I: "\bounded_linear D; (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0\ \ FDERIV f x :> D" by (simp add: fderiv_def) lemma FDERIV_D: "FDERIV f x :> D \ (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0" by (simp add: fderiv_def) lemma FDERIV_bounded_linear: "FDERIV f x :> D \ bounded_linear D" by (simp add: fderiv_def) lemma bounded_linear_zero: "bounded_linear (\x::'a::real_normed_vector. 0::'b::real_normed_vector)" proof show "(0::'b) = 0 + 0" by simp fix r show "(0::'b) = scaleR r 0" by simp have "\x::'a. norm (0::'b) \ norm x * 0" by simp thus "\K. \x::'a. norm (0::'b) \ norm x * K" .. qed lemma FDERIV_const: "FDERIV (\x. k) x :> (\h. 0)" by (simp add: fderiv_def bounded_linear_zero) lemma bounded_linear_ident: "bounded_linear (\x::'a::real_normed_vector. x)" proof fix x y :: 'a show "x + y = x + y" by simp fix r and x :: 'a show "scaleR r x = scaleR r x" by simp have "\x::'a. norm x \ norm x * 1" by simp thus "\K. \x::'a. norm x \ norm x * K" .. qed lemma FDERIV_ident: "FDERIV (\x. x) x :> (\h. h)" by (simp add: fderiv_def bounded_linear_ident) subsection {* Addition *} lemma add_diff_add: fixes a b c d :: "'a::ab_group_add" shows "(a + c) - (b + d) = (a - b) + (c - d)" by simp lemma bounded_linear_add: assumes "bounded_linear f" assumes "bounded_linear g" shows "bounded_linear (\x. f x + g x)" proof - interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact show ?thesis apply (unfold_locales) apply (simp only: f.add g.add add_ac) apply (simp only: f.scaleR g.scaleR scaleR_right_distrib) apply (rule f.pos_bounded [THEN exE], rename_tac Kf) apply (rule g.pos_bounded [THEN exE], rename_tac Kg) apply (rule_tac x="Kf + Kg" in exI, safe) apply (subst right_distrib) apply (rule order_trans [OF norm_triangle_ineq]) apply (rule add_mono, erule spec, erule spec) done qed lemma norm_ratio_ineq: fixes x y :: "'a::real_normed_vector" fixes h :: "'b::real_normed_vector" shows "norm (x + y) / norm h \ norm x / norm h + norm y / norm h" apply (rule ord_le_eq_trans) apply (rule divide_right_mono) apply (rule norm_triangle_ineq) apply (rule norm_ge_zero) apply (rule add_divide_distrib) done lemma FDERIV_add: assumes f: "FDERIV f x :> F" assumes g: "FDERIV g x :> G" shows "FDERIV (\x. f x + g x) x :> (\h. F h + G h)" proof (rule FDERIV_I) show "bounded_linear (\h. F h + G h)" apply (rule bounded_linear_add) apply (rule FDERIV_bounded_linear [OF f]) apply (rule FDERIV_bounded_linear [OF g]) done next have f': "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" using f by (rule FDERIV_D) have g': "(\h. norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" using g by (rule FDERIV_D) from f' g' have "(\h. norm (f (x + h) - f x - F h) / norm h + norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" by (rule LIM_add_zero) thus "(\h. norm (f (x + h) + g (x + h) - (f x + g x) - (F h + G h)) / norm h) -- 0 --> 0" apply (rule real_LIM_sandwich_zero) apply (simp add: divide_nonneg_pos) apply (simp only: add_diff_add) apply (rule norm_ratio_ineq) done qed subsection {* Subtraction *} lemma bounded_linear_minus: assumes "bounded_linear f" shows "bounded_linear (\x. - f x)" proof - interpret f: bounded_linear f by fact show ?thesis apply (unfold_locales) apply (simp add: f.add) apply (simp add: f.scaleR) apply (simp add: f.bounded) done qed lemma FDERIV_minus: "FDERIV f x :> F \ FDERIV (\x. - f x) x :> (\h. - F h)" apply (rule FDERIV_I) apply (rule bounded_linear_minus) apply (erule FDERIV_bounded_linear) apply (simp only: fderiv_def minus_diff_minus norm_minus_cancel) done lemma FDERIV_diff: "\FDERIV f x :> F; FDERIV g x :> G\ \ FDERIV (\x. f x - g x) x :> (\h. F h - G h)" by (simp only: diff_minus FDERIV_add FDERIV_minus) subsection {* Continuity *} lemma FDERIV_isCont: assumes f: "FDERIV f x :> F" shows "isCont f x" proof - from f interpret F: bounded_linear "F" by (rule FDERIV_bounded_linear) have "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" by (rule FDERIV_D [OF f]) hence "(\h. norm (f (x + h) - f x - F h) / norm h * norm h) -- 0 --> 0" by (intro LIM_mult_zero LIM_norm_zero LIM_ident) hence "(\h. norm (f (x + h) - f x - F h)) -- 0 --> 0" by (simp cong: LIM_cong) hence "(\h. f (x + h) - f x - F h) -- 0 --> 0" by (rule LIM_norm_zero_cancel) hence "(\h. f (x + h) - f x - F h + F h) -- 0 --> 0" by (intro LIM_add_zero F.LIM_zero LIM_ident) hence "(\h. f (x + h) - f x) -- 0 --> 0" by simp thus "isCont f x" unfolding isCont_iff by (rule LIM_zero_cancel) qed subsection {* Composition *} lemma real_divide_cancel_lemma: fixes a b c :: real shows "(b = 0 \ a = 0) \ (a / b) * (b / c) = a / c" by simp lemma bounded_linear_compose: assumes "bounded_linear f" assumes "bounded_linear g" shows "bounded_linear (\x. f (g x))" proof - interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact show ?thesis proof (unfold_locales) fix x y show "f (g (x + y)) = f (g x) + f (g y)" by (simp only: f.add g.add) next fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))" by (simp only: f.scaleR g.scaleR) next from f.pos_bounded obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by fast from g.pos_bounded obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by fast show "\K. \x. norm (f (g x)) \ norm x * K" proof (intro exI allI) fix x have "norm (f (g x)) \ norm (g x) * Kf" using f . also have "\ \ (norm x * Kg) * Kf" using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" by (rule mult_assoc) finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . qed qed qed lemma FDERIV_compose: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" fixes g :: "'b::real_normed_vector \ 'c::real_normed_vector" assumes f: "FDERIV f x :> F" assumes g: "FDERIV g (f x) :> G" shows "FDERIV (\x. g (f x)) x :> (\h. G (F h))" proof (rule FDERIV_I) from FDERIV_bounded_linear [OF g] FDERIV_bounded_linear [OF f] show "bounded_linear (\h. G (F h))" by (rule bounded_linear_compose) next let ?Rf = "\h. f (x + h) - f x - F h" let ?Rg = "\k. g (f x + k) - g (f x) - G k" let ?k = "\h. f (x + h) - f x" let ?Nf = "\h. norm (?Rf h) / norm h" let ?Ng = "\h. norm (?Rg (?k h)) / norm (?k h)" - from f interpret F!: bounded_linear "F" by (rule FDERIV_bounded_linear) - from g interpret G!: bounded_linear "G" by (rule FDERIV_bounded_linear) + from f interpret F: bounded_linear "F" by (rule FDERIV_bounded_linear) + from g interpret G: bounded_linear "G" by (rule FDERIV_bounded_linear) from F.bounded obtain kF where kF: "\x. norm (F x) \ norm x * kF" by fast from G.bounded obtain kG where kG: "\x. norm (G x) \ norm x * kG" by fast let ?fun2 = "\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)" show "(\h. norm (g (f (x + h)) - g (f x) - G (F h)) / norm h) -- 0 --> 0" proof (rule real_LIM_sandwich_zero) have Nf: "?Nf -- 0 --> 0" using FDERIV_D [OF f] . have Ng1: "isCont (\k. norm (?Rg k) / norm k) 0" by (simp add: isCont_def FDERIV_D [OF g]) have Ng2: "?k -- 0 --> 0" apply (rule LIM_zero) apply (fold isCont_iff) apply (rule FDERIV_isCont [OF f]) done have Ng: "?Ng -- 0 --> 0" using isCont_LIM_compose [OF Ng1 Ng2] by simp have "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) -- 0 --> 0 * kG + 0 * (0 + kF)" by (intro LIM_add LIM_mult LIM_const Nf Ng) thus "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) -- 0 --> 0" by simp next fix h::'a assume h: "h \ 0" thus "0 \ norm (g (f (x + h)) - g (f x) - G (F h)) / norm h" by (simp add: divide_nonneg_pos) next fix h::'a assume h: "h \ 0" have "g (f (x + h)) - g (f x) - G (F h) = G (?Rf h) + ?Rg (?k h)" by (simp add: G.diff) hence "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h = norm (G (?Rf h) + ?Rg (?k h)) / norm h" by (rule arg_cong) also have "\ \ norm (G (?Rf h)) / norm h + norm (?Rg (?k h)) / norm h" by (rule norm_ratio_ineq) also have "\ \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" proof (rule add_mono) show "norm (G (?Rf h)) / norm h \ ?Nf h * kG" apply (rule ord_le_eq_trans) apply (rule divide_right_mono [OF kG norm_ge_zero]) apply simp done next have "norm (?Rg (?k h)) / norm h = ?Ng h * (norm (?k h) / norm h)" apply (rule real_divide_cancel_lemma [symmetric]) apply (simp add: G.zero) done also have "\ \ ?Ng h * (?Nf h + kF)" proof (rule mult_left_mono) have "norm (?k h) / norm h = norm (?Rf h + F h) / norm h" by simp also have "\ \ ?Nf h + norm (F h) / norm h" by (rule norm_ratio_ineq) also have "\ \ ?Nf h + kF" apply (rule add_left_mono) apply (subst pos_divide_le_eq, simp add: h) apply (subst mult_commute) apply (rule kF) done finally show "norm (?k h) / norm h \ ?Nf h + kF" . next show "0 \ ?Ng h" apply (case_tac "f (x + h) - f x = 0", simp) apply (rule divide_nonneg_pos [OF norm_ge_zero]) apply simp done qed finally show "norm (?Rg (?k h)) / norm h \ ?Ng h * (?Nf h + kF)" . qed finally show "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" . qed qed subsection {* Product Rule *} lemma (in bounded_bilinear) FDERIV_lemma: "a' ** b' - a ** b - (a ** B + A ** b) = a ** (b' - b - B) + (a' - a - A) ** b' + A ** (b' - b)" by (simp add: diff_left diff_right) lemma (in bounded_bilinear) FDERIV: fixes x :: "'d::real_normed_vector" assumes f: "FDERIV f x :> F" assumes g: "FDERIV g x :> G" shows "FDERIV (\x. f x ** g x) x :> (\h. f x ** G h + F h ** g x)" proof (rule FDERIV_I) show "bounded_linear (\h. f x ** G h + F h ** g x)" apply (rule bounded_linear_add) apply (rule bounded_linear_compose [OF bounded_linear_right]) apply (rule FDERIV_bounded_linear [OF g]) apply (rule bounded_linear_compose [OF bounded_linear_left]) apply (rule FDERIV_bounded_linear [OF f]) done next from bounded_linear.bounded [OF FDERIV_bounded_linear [OF f]] obtain KF where norm_F: "\x. norm (F x) \ norm x * KF" by fast from pos_bounded obtain K where K: "0 < K" and norm_prod: "\a b. norm (a ** b) \ norm a * norm b * K" by fast let ?Rf = "\h. f (x + h) - f x - F h" let ?Rg = "\h. g (x + h) - g x - G h" let ?fun1 = "\h. norm (f x ** ?Rg h + ?Rf h ** g (x + h) + F h ** (g (x + h) - g x)) / norm h" let ?fun2 = "\h. norm (f x) * (norm (?Rg h) / norm h) * K + norm (?Rf h) / norm h * norm (g (x + h)) * K + KF * norm (g (x + h) - g x) * K" have "?fun1 -- 0 --> 0" proof (rule real_LIM_sandwich_zero) from f g isCont_iff [THEN iffD1, OF FDERIV_isCont [OF g]] have "?fun2 -- 0 --> norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K" by (intro LIM_add LIM_mult LIM_const LIM_norm LIM_zero FDERIV_D) thus "?fun2 -- 0 --> 0" by simp next fix h::'d assume "h \ 0" thus "0 \ ?fun1 h" by (simp add: divide_nonneg_pos) next fix h::'d assume "h \ 0" have "?fun1 h \ (norm (f x) * norm (?Rg h) * K + norm (?Rf h) * norm (g (x + h)) * K + norm h * KF * norm (g (x + h) - g x) * K) / norm h" by (intro divide_right_mono mult_mono' order_trans [OF norm_triangle_ineq add_mono] order_trans [OF norm_prod mult_right_mono] mult_nonneg_nonneg order_refl norm_ge_zero norm_F K [THEN order_less_imp_le] ) also have "\ = ?fun2 h" by (simp add: add_divide_distrib) finally show "?fun1 h \ ?fun2 h" . qed thus "(\h. norm (f (x + h) ** g (x + h) - f x ** g x - (f x ** G h + F h ** g x)) / norm h) -- 0 --> 0" by (simp only: FDERIV_lemma) qed lemmas FDERIV_mult = mult.FDERIV lemmas FDERIV_scaleR = scaleR.FDERIV subsection {* Powers *} lemma FDERIV_power_Suc: fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" shows "FDERIV (\x. x ^ Suc n) x :> (\h. (1 + of_nat n) * x ^ n * h)" apply (induct n) apply (simp add: power_Suc FDERIV_ident) apply (drule FDERIV_mult [OF FDERIV_ident]) apply (simp only: of_nat_Suc left_distrib mult_1_left) apply (simp only: power_Suc right_distrib add_ac mult_ac) done lemma FDERIV_power: fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}" shows "FDERIV (\x. x ^ n) x :> (\h. of_nat n * x ^ (n - 1) * h)" apply (cases n) apply (simp add: FDERIV_const) apply (simp add: FDERIV_power_Suc del: power_Suc) done subsection {* Inverse *} lemma inverse_diff_inverse: "\(a::'a::division_ring) \ 0; b \ 0\ \ inverse a - inverse b = - (inverse a * (a - b) * inverse b)" by (simp add: right_diff_distrib left_diff_distrib mult_assoc) lemmas bounded_linear_mult_const = mult.bounded_linear_left [THEN bounded_linear_compose] lemmas bounded_linear_const_mult = mult.bounded_linear_right [THEN bounded_linear_compose] lemma FDERIV_inverse: fixes x :: "'a::real_normed_div_algebra" assumes x: "x \ 0" shows "FDERIV inverse x :> (\h. - (inverse x * h * inverse x))" (is "FDERIV ?inv _ :> _") proof (rule FDERIV_I) show "bounded_linear (\h. - (?inv x * h * ?inv x))" apply (rule bounded_linear_minus) apply (rule bounded_linear_mult_const) apply (rule bounded_linear_const_mult) apply (rule bounded_linear_ident) done next show "(\h. norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h) -- 0 --> 0" proof (rule LIM_equal2) show "0 < norm x" using x by simp next fix h::'a assume 1: "h \ 0" assume "norm (h - 0) < norm x" hence "h \ -x" by clarsimp hence 2: "x + h \ 0" apply (rule contrapos_nn) apply (rule sym) apply (erule equals_zero_I) done show "norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h = norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" apply (subst inverse_diff_inverse [OF 2 x]) apply (subst minus_diff_minus) apply (subst norm_minus_cancel) apply (simp add: left_diff_distrib) done next show "(\h. norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h) -- 0 --> 0" proof (rule real_LIM_sandwich_zero) show "(\h. norm (?inv (x + h) - ?inv x) * norm (?inv x)) -- 0 --> 0" apply (rule LIM_mult_left_zero) apply (rule LIM_norm_zero) apply (rule LIM_zero) apply (rule LIM_offset_zero) apply (rule LIM_inverse) apply (rule LIM_ident) apply (rule x) done next fix h::'a assume h: "h \ 0" show "0 \ norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" apply (rule divide_nonneg_pos) apply (rule norm_ge_zero) apply (simp add: h) done next fix h::'a assume h: "h \ 0" have "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h \ norm (?inv (x + h) - ?inv x) * norm h * norm (?inv x) / norm h" apply (rule divide_right_mono [OF _ norm_ge_zero]) apply (rule order_trans [OF norm_mult_ineq]) apply (rule mult_right_mono [OF _ norm_ge_zero]) apply (rule norm_mult_ineq) done also have "\ = norm (?inv (x + h) - ?inv x) * norm (?inv x)" by simp finally show "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h \ norm (?inv (x + h) - ?inv x) * norm (?inv x)" . qed qed qed subsection {* Alternate definition *} lemma field_fderiv_def: fixes x :: "'a::real_normed_field" shows "FDERIV f x :> (\h. h * D) = (\h. (f (x + h) - f x) / h) -- 0 --> D" apply (unfold fderiv_def) apply (simp add: mult.bounded_linear_left) apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric]) apply (subst diff_divide_distrib) apply (subst times_divide_eq_left [symmetric]) apply (simp cong: LIM_cong) apply (simp add: LIM_norm_zero_iff LIM_zero_iff) done end diff --git a/src/HOL/Library/Inner_Product.thy b/src/HOL/Library/Inner_Product.thy --- a/src/HOL/Library/Inner_Product.thy +++ b/src/HOL/Library/Inner_Product.thy @@ -1,296 +1,296 @@ (* Title: Inner_Product.thy Author: Brian Huffman *) header {* Inner Product Spaces and the Gradient Derivative *} theory Inner_Product imports Complex_Main FrechetDeriv begin subsection {* Real inner product spaces *} class real_inner = real_vector + sgn_div_norm + fixes inner :: "'a \ 'a \ real" assumes inner_commute: "inner x y = inner y x" and inner_left_distrib: "inner (x + y) z = inner x z + inner y z" and inner_scaleR_left: "inner (scaleR r x) y = r * (inner x y)" and inner_ge_zero [simp]: "0 \ inner x x" and inner_eq_zero_iff [simp]: "inner x x = 0 \ x = 0" and norm_eq_sqrt_inner: "norm x = sqrt (inner x x)" begin lemma inner_zero_left [simp]: "inner 0 x = 0" using inner_left_distrib [of 0 0 x] by simp lemma inner_minus_left [simp]: "inner (- x) y = - inner x y" using inner_left_distrib [of x "- x" y] by simp lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z" by (simp add: diff_minus inner_left_distrib) text {* Transfer distributivity rules to right argument. *} lemma inner_right_distrib: "inner x (y + z) = inner x y + inner x z" using inner_left_distrib [of y z x] by (simp only: inner_commute) lemma inner_scaleR_right: "inner x (scaleR r y) = r * (inner x y)" using inner_scaleR_left [of r y x] by (simp only: inner_commute) lemma inner_zero_right [simp]: "inner x 0 = 0" using inner_zero_left [of x] by (simp only: inner_commute) lemma inner_minus_right [simp]: "inner x (- y) = - inner x y" using inner_minus_left [of y x] by (simp only: inner_commute) lemma inner_diff_right: "inner x (y - z) = inner x y - inner x z" using inner_diff_left [of y z x] by (simp only: inner_commute) lemmas inner_distrib = inner_left_distrib inner_right_distrib lemmas inner_diff = inner_diff_left inner_diff_right lemmas inner_scaleR = inner_scaleR_left inner_scaleR_right lemma inner_gt_zero_iff [simp]: "0 < inner x x \ x \ 0" by (simp add: order_less_le) lemma power2_norm_eq_inner: "(norm x)\ = inner x x" by (simp add: norm_eq_sqrt_inner) lemma Cauchy_Schwarz_ineq: "(inner x y)\ \ inner x x * inner y y" proof (cases) assume "y = 0" thus ?thesis by simp next assume y: "y \ 0" let ?r = "inner x y / inner y y" have "0 \ inner (x - scaleR ?r y) (x - scaleR ?r y)" by (rule inner_ge_zero) also have "\ = inner x x - inner y x * ?r" by (simp add: inner_diff inner_scaleR) also have "\ = inner x x - (inner x y)\ / inner y y" by (simp add: power2_eq_square inner_commute) finally have "0 \ inner x x - (inner x y)\ / inner y y" . hence "(inner x y)\ / inner y y \ inner x x" by (simp add: le_diff_eq) thus "(inner x y)\ \ inner x x * inner y y" by (simp add: pos_divide_le_eq y) qed lemma Cauchy_Schwarz_ineq2: "\inner x y\ \ norm x * norm y" proof (rule power2_le_imp_le) have "(inner x y)\ \ inner x x * inner y y" using Cauchy_Schwarz_ineq . thus "\inner x y\\ \ (norm x * norm y)\" by (simp add: power_mult_distrib power2_norm_eq_inner) show "0 \ norm x * norm y" unfolding norm_eq_sqrt_inner by (intro mult_nonneg_nonneg real_sqrt_ge_zero inner_ge_zero) qed subclass real_normed_vector proof fix a :: real and x y :: 'a show "0 \ norm x" unfolding norm_eq_sqrt_inner by simp show "norm x = 0 \ x = 0" unfolding norm_eq_sqrt_inner by simp show "norm (x + y) \ norm x + norm y" proof (rule power2_le_imp_le) have "inner x y \ norm x * norm y" by (rule order_trans [OF abs_ge_self Cauchy_Schwarz_ineq2]) thus "(norm (x + y))\ \ (norm x + norm y)\" unfolding power2_sum power2_norm_eq_inner by (simp add: inner_distrib inner_commute) show "0 \ norm x + norm y" unfolding norm_eq_sqrt_inner by (simp add: add_nonneg_nonneg) qed have "sqrt (a\ * inner x x) = \a\ * sqrt (inner x x)" by (simp add: real_sqrt_mult_distrib) then show "norm (a *\<^sub>R x) = \a\ * norm x" unfolding norm_eq_sqrt_inner by (simp add: inner_scaleR power2_eq_square mult_assoc) qed end -interpretation inner!: +interpretation inner: bounded_bilinear "inner::'a::real_inner \ 'a \ real" proof fix x y z :: 'a and r :: real show "inner (x + y) z = inner x z + inner y z" by (rule inner_left_distrib) show "inner x (y + z) = inner x y + inner x z" by (rule inner_right_distrib) show "inner (scaleR r x) y = scaleR r (inner x y)" unfolding real_scaleR_def by (rule inner_scaleR_left) show "inner x (scaleR r y) = scaleR r (inner x y)" unfolding real_scaleR_def by (rule inner_scaleR_right) show "\K. \x y::'a. norm (inner x y) \ norm x * norm y * K" proof show "\x y::'a. norm (inner x y) \ norm x * norm y * 1" by (simp add: Cauchy_Schwarz_ineq2) qed qed -interpretation inner_left!: +interpretation inner_left: bounded_linear "\x::'a::real_inner. inner x y" by (rule inner.bounded_linear_left) -interpretation inner_right!: +interpretation inner_right: bounded_linear "\y::'a::real_inner. inner x y" by (rule inner.bounded_linear_right) subsection {* Class instances *} instantiation real :: real_inner begin definition inner_real_def [simp]: "inner = op *" instance proof fix x y z r :: real show "inner x y = inner y x" unfolding inner_real_def by (rule mult_commute) show "inner (x + y) z = inner x z + inner y z" unfolding inner_real_def by (rule left_distrib) show "inner (scaleR r x) y = r * inner x y" unfolding inner_real_def real_scaleR_def by (rule mult_assoc) show "0 \ inner x x" unfolding inner_real_def by simp show "inner x x = 0 \ x = 0" unfolding inner_real_def by simp show "norm x = sqrt (inner x x)" unfolding inner_real_def by simp qed end instantiation complex :: real_inner begin definition inner_complex_def: "inner x y = Re x * Re y + Im x * Im y" instance proof fix x y z :: complex and r :: real show "inner x y = inner y x" unfolding inner_complex_def by (simp add: mult_commute) show "inner (x + y) z = inner x z + inner y z" unfolding inner_complex_def by (simp add: left_distrib) show "inner (scaleR r x) y = r * inner x y" unfolding inner_complex_def by (simp add: right_distrib) show "0 \ inner x x" unfolding inner_complex_def by (simp add: add_nonneg_nonneg) show "inner x x = 0 \ x = 0" unfolding inner_complex_def by (simp add: add_nonneg_eq_0_iff complex_Re_Im_cancel_iff) show "norm x = sqrt (inner x x)" unfolding inner_complex_def complex_norm_def by (simp add: power2_eq_square) qed end subsection {* Gradient derivative *} definition gderiv :: "['a::real_inner \ real, 'a, 'a] \ bool" ("(GDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "GDERIV f x :> D \ FDERIV f x :> (\h. inner h D)" lemma deriv_fderiv: "DERIV f x :> D \ FDERIV f x :> (\h. h * D)" by (simp only: deriv_def field_fderiv_def) lemma gderiv_deriv [simp]: "GDERIV f x :> D \ DERIV f x :> D" by (simp only: gderiv_def deriv_fderiv inner_real_def) lemma GDERIV_DERIV_compose: "\GDERIV f x :> df; DERIV g (f x) :> dg\ \ GDERIV (\x. g (f x)) x :> scaleR dg df" unfolding gderiv_def deriv_fderiv apply (drule (1) FDERIV_compose) apply (simp add: inner_scaleR_right mult_ac) done lemma FDERIV_subst: "\FDERIV f x :> df; df = d\ \ FDERIV f x :> d" by simp lemma GDERIV_subst: "\GDERIV f x :> df; df = d\ \ GDERIV f x :> d" by simp lemma GDERIV_const: "GDERIV (\x. k) x :> 0" unfolding gderiv_def inner_right.zero by (rule FDERIV_const) lemma GDERIV_add: "\GDERIV f x :> df; GDERIV g x :> dg\ \ GDERIV (\x. f x + g x) x :> df + dg" unfolding gderiv_def inner_right.add by (rule FDERIV_add) lemma GDERIV_minus: "GDERIV f x :> df \ GDERIV (\x. - f x) x :> - df" unfolding gderiv_def inner_right.minus by (rule FDERIV_minus) lemma GDERIV_diff: "\GDERIV f x :> df; GDERIV g x :> dg\ \ GDERIV (\x. f x - g x) x :> df - dg" unfolding gderiv_def inner_right.diff by (rule FDERIV_diff) lemma GDERIV_scaleR: "\DERIV f x :> df; GDERIV g x :> dg\ \ GDERIV (\x. scaleR (f x) (g x)) x :> (scaleR (f x) dg + scaleR df (g x))" unfolding gderiv_def deriv_fderiv inner_right.add inner_right.scaleR apply (rule FDERIV_subst) apply (erule (1) scaleR.FDERIV) apply (simp add: mult_ac) done lemma GDERIV_mult: "\GDERIV f x :> df; GDERIV g x :> dg\ \ GDERIV (\x. f x * g x) x :> scaleR (f x) dg + scaleR (g x) df" unfolding gderiv_def apply (rule FDERIV_subst) apply (erule (1) FDERIV_mult) apply (simp add: inner_distrib inner_scaleR mult_ac) done lemma GDERIV_inverse: "\GDERIV f x :> df; f x \ 0\ \ GDERIV (\x. inverse (f x)) x :> - (inverse (f x))\ *\<^sub>R df" apply (erule GDERIV_DERIV_compose) apply (erule DERIV_inverse [folded numeral_2_eq_2]) done lemma GDERIV_norm: assumes "x \ 0" shows "GDERIV (\x. norm x) x :> sgn x" proof - have 1: "FDERIV (\x. inner x x) x :> (\h. inner x h + inner h x)" by (intro inner.FDERIV FDERIV_ident) have 2: "(\h. inner x h + inner h x) = (\h. inner h (scaleR 2 x))" by (simp add: expand_fun_eq inner_scaleR inner_commute) have "0 < inner x x" using `x \ 0` by simp then have 3: "DERIV sqrt (inner x x) :> (inverse (sqrt (inner x x)) / 2)" by (rule DERIV_real_sqrt) have 4: "(inverse (sqrt (inner x x)) / 2) *\<^sub>R 2 *\<^sub>R x = sgn x" by (simp add: sgn_div_norm norm_eq_sqrt_inner) show ?thesis unfolding norm_eq_sqrt_inner apply (rule GDERIV_subst [OF _ 4]) apply (rule GDERIV_DERIV_compose [where g=sqrt and df="scaleR 2 x"]) apply (subst gderiv_def) apply (rule FDERIV_subst [OF _ 2]) apply (rule 1) apply (rule 3) done qed lemmas FDERIV_norm = GDERIV_norm [unfolded gderiv_def] end diff --git a/src/HOL/Library/Multiset.thy b/src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy +++ b/src/HOL/Library/Multiset.thy @@ -1,1632 +1,1632 @@ (* Title: HOL/Library/Multiset.thy Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker *) header {* Multisets *} theory Multiset imports List Main begin subsection {* The type of multisets *} typedef 'a multiset = "{f::'a => nat. finite {x . f x > 0}}" proof show "(\x. 0::nat) \ ?multiset" by simp qed lemmas multiset_typedef [simp] = Abs_multiset_inverse Rep_multiset_inverse Rep_multiset and [simp] = Rep_multiset_inject [symmetric] definition Mempty :: "'a multiset" ("{#}") where [code del]: "{#} = Abs_multiset (\a. 0)" definition single :: "'a => 'a multiset" where [code del]: "single a = Abs_multiset (\b. if b = a then 1 else 0)" definition count :: "'a multiset => 'a => nat" where "count = Rep_multiset" definition MCollect :: "'a multiset => ('a => bool) => 'a multiset" where "MCollect M P = Abs_multiset (\x. if P x then Rep_multiset M x else 0)" abbreviation Melem :: "'a => 'a multiset => bool" ("(_/ :# _)" [50, 51] 50) where "a :# M == 0 < count M a" notation (xsymbols) Melem (infix "\#" 50) syntax "_MCollect" :: "pttrn => 'a multiset => bool => 'a multiset" ("(1{# _ :# _./ _#})") translations "{#x :# M. P#}" == "CONST MCollect M (\x. P)" definition set_of :: "'a multiset => 'a set" where "set_of M = {x. x :# M}" instantiation multiset :: (type) "{plus, minus, zero, size}" begin definition union_def [code del]: "M + N = Abs_multiset (\a. Rep_multiset M a + Rep_multiset N a)" definition diff_def [code del]: "M - N = Abs_multiset (\a. Rep_multiset M a - Rep_multiset N a)" definition Zero_multiset_def [simp]: "0 = {#}" definition size_def: "size M = setsum (count M) (set_of M)" instance .. end definition multiset_inter :: "'a multiset \ 'a multiset \ 'a multiset" (infixl "#\" 70) where "multiset_inter A B = A - (A - B)" text {* Multiset Enumeration *} syntax "_multiset" :: "args => 'a multiset" ("{#(_)#}") translations "{#x, xs#}" == "{#x#} + {#xs#}" "{#x#}" == "CONST single x" text {* \medskip Preservation of the representing set @{term multiset}. *} lemma const0_in_multiset: "(\a. 0) \ multiset" by (simp add: multiset_def) lemma only1_in_multiset: "(\b. if b = a then 1 else 0) \ multiset" by (simp add: multiset_def) lemma union_preserves_multiset: "M \ multiset ==> N \ multiset ==> (\a. M a + N a) \ multiset" by (simp add: multiset_def) lemma diff_preserves_multiset: "M \ multiset ==> (\a. M a - N a) \ multiset" apply (simp add: multiset_def) apply (rule finite_subset) apply auto done lemma MCollect_preserves_multiset: "M \ multiset ==> (\x. if P x then M x else 0) \ multiset" apply (simp add: multiset_def) apply (rule finite_subset, auto) done lemmas in_multiset = const0_in_multiset only1_in_multiset union_preserves_multiset diff_preserves_multiset MCollect_preserves_multiset subsection {* Algebraic properties *} subsubsection {* Union *} lemma union_empty [simp]: "M + {#} = M \ {#} + M = M" by (simp add: union_def Mempty_def in_multiset) lemma union_commute: "M + N = N + (M::'a multiset)" by (simp add: union_def add_ac in_multiset) lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" by (simp add: union_def add_ac in_multiset) lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" proof - have "M + (N + K) = (N + K) + M" by (rule union_commute) also have "\ = N + (K + M)" by (rule union_assoc) also have "K + M = M + K" by (rule union_commute) finally show ?thesis . qed lemmas union_ac = union_assoc union_commute union_lcomm instance multiset :: (type) comm_monoid_add proof fix a b c :: "'a multiset" show "(a + b) + c = a + (b + c)" by (rule union_assoc) show "a + b = b + a" by (rule union_commute) show "0 + a = a" by simp qed subsubsection {* Difference *} lemma diff_empty [simp]: "M - {#} = M \ {#} - M = {#}" by (simp add: Mempty_def diff_def in_multiset) lemma diff_union_inverse2 [simp]: "M + {#a#} - {#a#} = M" by (simp add: union_def diff_def in_multiset) lemma diff_cancel: "A - A = {#}" by (simp add: diff_def Mempty_def) subsubsection {* Count of elements *} lemma count_empty [simp]: "count {#} a = 0" by (simp add: count_def Mempty_def in_multiset) lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)" by (simp add: count_def single_def in_multiset) lemma count_union [simp]: "count (M + N) a = count M a + count N a" by (simp add: count_def union_def in_multiset) lemma count_diff [simp]: "count (M - N) a = count M a - count N a" by (simp add: count_def diff_def in_multiset) lemma count_MCollect [simp]: "count {# x:#M. P x #} a = (if P a then count M a else 0)" by (simp add: count_def MCollect_def in_multiset) subsubsection {* Set of elements *} lemma set_of_empty [simp]: "set_of {#} = {}" by (simp add: set_of_def) lemma set_of_single [simp]: "set_of {#b#} = {b}" by (simp add: set_of_def) lemma set_of_union [simp]: "set_of (M + N) = set_of M \ set_of N" by (auto simp add: set_of_def) lemma set_of_eq_empty_iff [simp]: "(set_of M = {}) = (M = {#})" by (auto simp: set_of_def Mempty_def in_multiset count_def expand_fun_eq [where f="Rep_multiset M"]) lemma mem_set_of_iff [simp]: "(x \ set_of M) = (x :# M)" by (auto simp add: set_of_def) lemma set_of_MCollect [simp]: "set_of {# x:#M. P x #} = set_of M \ {x. P x}" by (auto simp add: set_of_def) subsubsection {* Size *} lemma size_empty [simp]: "size {#} = 0" by (simp add: size_def) lemma size_single [simp]: "size {#b#} = 1" by (simp add: size_def) lemma finite_set_of [iff]: "finite (set_of M)" using Rep_multiset [of M] by (simp add: multiset_def set_of_def count_def) lemma setsum_count_Int: "finite A ==> setsum (count N) (A \ set_of N) = setsum (count N) A" apply (induct rule: finite_induct) apply simp apply (simp add: Int_insert_left set_of_def) done lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" apply (unfold size_def) apply (subgoal_tac "count (M + N) = (\a. count M a + count N a)") prefer 2 apply (rule ext, simp) apply (simp (no_asm_simp) add: setsum_Un_nat setsum_addf setsum_count_Int) apply (subst Int_commute) apply (simp (no_asm_simp) add: setsum_count_Int) done lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" apply (unfold size_def Mempty_def count_def, auto simp: in_multiset) apply (simp add: set_of_def count_def in_multiset expand_fun_eq) done lemma nonempty_has_size: "(S \ {#}) = (0 < size S)" by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) lemma size_eq_Suc_imp_elem: "size M = Suc n ==> \a. a :# M" apply (unfold size_def) apply (drule setsum_SucD) apply auto done subsubsection {* Equality of multisets *} lemma multiset_eq_conv_count_eq: "(M = N) = (\a. count M a = count N a)" by (simp add: count_def expand_fun_eq) lemma single_not_empty [simp]: "{#a#} \ {#} \ {#} \ {#a#}" by (simp add: single_def Mempty_def in_multiset expand_fun_eq) lemma single_eq_single [simp]: "({#a#} = {#b#}) = (a = b)" by (auto simp add: single_def in_multiset expand_fun_eq) lemma union_eq_empty [iff]: "(M + N = {#}) = (M = {#} \ N = {#})" by (auto simp add: union_def Mempty_def in_multiset expand_fun_eq) lemma empty_eq_union [iff]: "({#} = M + N) = (M = {#} \ N = {#})" by (auto simp add: union_def Mempty_def in_multiset expand_fun_eq) lemma union_right_cancel [simp]: "(M + K = N + K) = (M = (N::'a multiset))" by (simp add: union_def in_multiset expand_fun_eq) lemma union_left_cancel [simp]: "(K + M = K + N) = (M = (N::'a multiset))" by (simp add: union_def in_multiset expand_fun_eq) lemma union_is_single: "(M + N = {#a#}) = (M = {#a#} \ N={#} \ M = {#} \ N = {#a#})" apply (simp add: Mempty_def single_def union_def in_multiset add_is_1 expand_fun_eq) apply blast done lemma single_is_union: "({#a#} = M + N) \ ({#a#} = M \ N = {#} \ M = {#} \ {#a#} = N)" apply (unfold Mempty_def single_def union_def) apply (simp add: add_is_1 one_is_add in_multiset expand_fun_eq) apply (blast dest: sym) done lemma add_eq_conv_diff: "(M + {#a#} = N + {#b#}) = (M = N \ a = b \ M = N - {#a#} + {#b#} \ N = M - {#b#} + {#a#})" using [[simproc del: neq]] apply (unfold single_def union_def diff_def) apply (simp (no_asm) add: in_multiset expand_fun_eq) apply (rule conjI, force, safe, simp_all) apply (simp add: eq_sym_conv) done declare Rep_multiset_inject [symmetric, simp del] instance multiset :: (type) cancel_ab_semigroup_add proof fix a b c :: "'a multiset" show "a + b = a + c \ b = c" by simp qed lemma insert_DiffM: "x \# M \ {#x#} + (M - {#x#}) = M" by (clarsimp simp: multiset_eq_conv_count_eq) lemma insert_DiffM2[simp]: "x \# M \ M - {#x#} + {#x#} = M" by (clarsimp simp: multiset_eq_conv_count_eq) lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \ X = Y" by (induct A arbitrary: X Y) auto lemma multi_self_add_other_not_self[simp]: "(A = A + {#x#}) = False" by (metis single_not_empty union_empty union_left_cancel) lemma insert_noteq_member: assumes BC: "B + {#b#} = C + {#c#}" and bnotc: "b \ c" shows "c \# B" proof - have "c \# C + {#c#}" by simp have nc: "\ c \# {#b#}" using bnotc by simp then have "c \# B + {#b#}" using BC by simp then show "c \# B" using nc by simp qed lemma add_eq_conv_ex: "(M + {#a#} = N + {#b#}) = (M = N \ a = b \ (\K. M = K + {#b#} \ N = K + {#a#}))" by (auto simp add: add_eq_conv_diff) lemma empty_multiset_count: "(\x. count A x = 0) = (A = {#})" by (metis count_empty multiset_eq_conv_count_eq) subsubsection {* Intersection *} lemma multiset_inter_count: "count (A #\ B) x = min (count A x) (count B x)" by (simp add: multiset_inter_def min_def) lemma multiset_inter_commute: "A #\ B = B #\ A" by (simp add: multiset_eq_conv_count_eq multiset_inter_count min_max.inf_commute) lemma multiset_inter_assoc: "A #\ (B #\ C) = A #\ B #\ C" by (simp add: multiset_eq_conv_count_eq multiset_inter_count min_max.inf_assoc) lemma multiset_inter_left_commute: "A #\ (B #\ C) = B #\ (A #\ C)" by (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def) lemmas multiset_inter_ac = multiset_inter_commute multiset_inter_assoc multiset_inter_left_commute lemma multiset_inter_single: "a \ b \ {#a#} #\ {#b#} = {#}" by (simp add: multiset_eq_conv_count_eq multiset_inter_count) lemma multiset_union_diff_commute: "B #\ C = {#} \ A + B - C = A - C + B" apply (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def split: split_if_asm) apply clarsimp apply (erule_tac x = a in allE) apply auto done subsubsection {* Comprehension (filter) *} lemma MCollect_empty [simp]: "MCollect {#} P = {#}" by (simp add: MCollect_def Mempty_def Abs_multiset_inject in_multiset expand_fun_eq) lemma MCollect_single [simp]: "MCollect {#x#} P = (if P x then {#x#} else {#})" by (simp add: MCollect_def Mempty_def single_def Abs_multiset_inject in_multiset expand_fun_eq) lemma MCollect_union [simp]: "MCollect (M+N) f = MCollect M f + MCollect N f" by (simp add: MCollect_def union_def Abs_multiset_inject in_multiset expand_fun_eq) subsection {* Induction and case splits *} lemma setsum_decr: "finite F ==> (0::nat) < f a ==> setsum (f (a := f a - 1)) F = (if a\F then setsum f F - 1 else setsum f F)" apply (induct rule: finite_induct) apply auto apply (drule_tac a = a in mk_disjoint_insert, auto) done lemma rep_multiset_induct_aux: assumes 1: "P (\a. (0::nat))" and 2: "!!f b. f \ multiset ==> P f ==> P (f (b := f b + 1))" shows "\f. f \ multiset --> setsum f {x. f x \ 0} = n --> P f" apply (unfold multiset_def) apply (induct_tac n, simp, clarify) apply (subgoal_tac "f = (\a.0)") apply simp apply (rule 1) apply (rule ext, force, clarify) apply (frule setsum_SucD, clarify) apply (rename_tac a) apply (subgoal_tac "finite {x. (f (a := f a - 1)) x > 0}") prefer 2 apply (rule finite_subset) prefer 2 apply assumption apply simp apply blast apply (subgoal_tac "f = (f (a := f a - 1))(a := (f (a := f a - 1)) a + 1)") prefer 2 apply (rule ext) apply (simp (no_asm_simp)) apply (erule ssubst, rule 2 [unfolded multiset_def], blast) apply (erule allE, erule impE, erule_tac [2] mp, blast) apply (simp (no_asm_simp) add: setsum_decr del: fun_upd_apply One_nat_def) apply (subgoal_tac "{x. x \ a --> f x \ 0} = {x. f x \ 0}") prefer 2 apply blast apply (subgoal_tac "{x. x \ a \ f x \ 0} = {x. f x \ 0} - {a}") prefer 2 apply blast apply (simp add: le_imp_diff_is_add setsum_diff1_nat cong: conj_cong) done theorem rep_multiset_induct: "f \ multiset ==> P (\a. 0) ==> (!!f b. f \ multiset ==> P f ==> P (f (b := f b + 1))) ==> P f" using rep_multiset_induct_aux by blast theorem multiset_induct [case_names empty add, induct type: multiset]: assumes empty: "P {#}" and add: "!!M x. P M ==> P (M + {#x#})" shows "P M" proof - note defns = union_def single_def Mempty_def show ?thesis apply (rule Rep_multiset_inverse [THEN subst]) apply (rule Rep_multiset [THEN rep_multiset_induct]) apply (rule empty [unfolded defns]) apply (subgoal_tac "f(b := f b + 1) = (\a. f a + (if a=b then 1 else 0))") prefer 2 apply (simp add: expand_fun_eq) apply (erule ssubst) apply (erule Abs_multiset_inverse [THEN subst]) apply (drule add [unfolded defns, simplified]) apply(simp add:in_multiset) done qed lemma multi_nonempty_split: "M \ {#} \ \A a. M = A + {#a#}" by (induct M) auto lemma multiset_cases [cases type, case_names empty add]: assumes em: "M = {#} \ P" assumes add: "\N x. M = N + {#x#} \ P" shows "P" proof (cases "M = {#}") assume "M = {#}" then show ?thesis using em by simp next assume "M \ {#}" then obtain M' m where "M = M' + {#m#}" by (blast dest: multi_nonempty_split) then show ?thesis using add by simp qed lemma multi_member_split: "x \# M \ \A. M = A + {#x#}" apply (cases M) apply simp apply (rule_tac x="M - {#x#}" in exI, simp) done lemma multiset_partition: "M = {# x:#M. P x #} + {# x:#M. \ P x #}" apply (subst multiset_eq_conv_count_eq) apply auto done declare multiset_typedef [simp del] lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" by (cases "B = {#}") (auto dest: multi_member_split) subsection {* Orderings *} subsubsection {* Well-foundedness *} definition mult1 :: "('a \ 'a) set => ('a multiset \ 'a multiset) set" where [code del]: "mult1 r = {(N, M). \a M0 K. M = M0 + {#a#} \ N = M0 + K \ (\b. b :# K --> (b, a) \ r)}" definition mult :: "('a \ 'a) set => ('a multiset \ 'a multiset) set" where "mult r = (mult1 r)\<^sup>+" lemma not_less_empty [iff]: "(M, {#}) \ mult1 r" by (simp add: mult1_def) lemma less_add: "(N, M0 + {#a#}) \ mult1 r ==> (\M. (M, M0) \ mult1 r \ N = M + {#a#}) \ (\K. (\b. b :# K --> (b, a) \ r) \ N = M0 + K)" (is "_ \ ?case1 (mult1 r) \ ?case2") proof (unfold mult1_def) let ?r = "\K a. \b. b :# K --> (b, a) \ r" let ?R = "\N M. \a M0 K. M = M0 + {#a#} \ N = M0 + K \ ?r K a" let ?case1 = "?case1 {(N, M). ?R N M}" assume "(N, M0 + {#a#}) \ {(N, M). ?R N M}" then have "\a' M0' K. M0 + {#a#} = M0' + {#a'#} \ N = M0' + K \ ?r K a'" by simp then show "?case1 \ ?case2" proof (elim exE conjE) fix a' M0' K assume N: "N = M0' + K" and r: "?r K a'" assume "M0 + {#a#} = M0' + {#a'#}" then have "M0 = M0' \ a = a' \ (\K'. M0 = K' + {#a'#} \ M0' = K' + {#a#})" by (simp only: add_eq_conv_ex) then show ?thesis proof (elim disjE conjE exE) assume "M0 = M0'" "a = a'" with N r have "?r K a \ N = M0 + K" by simp then have ?case2 .. then show ?thesis .. next fix K' assume "M0' = K' + {#a#}" with N have n: "N = K' + K + {#a#}" by (simp add: union_ac) assume "M0 = K' + {#a'#}" with r have "?R (K' + K) M0" by blast with n have ?case1 by simp then show ?thesis .. qed qed qed lemma all_accessible: "wf r ==> \M. M \ acc (mult1 r)" proof let ?R = "mult1 r" let ?W = "acc ?R" { fix M M0 a assume M0: "M0 \ ?W" and wf_hyp: "!!b. (b, a) \ r ==> (\M \ ?W. M + {#b#} \ ?W)" and acc_hyp: "\M. (M, M0) \ ?R --> M + {#a#} \ ?W" have "M0 + {#a#} \ ?W" proof (rule accI [of "M0 + {#a#}"]) fix N assume "(N, M0 + {#a#}) \ ?R" then have "((\M. (M, M0) \ ?R \ N = M + {#a#}) \ (\K. (\b. b :# K --> (b, a) \ r) \ N = M0 + K))" by (rule less_add) then show "N \ ?W" proof (elim exE disjE conjE) fix M assume "(M, M0) \ ?R" and N: "N = M + {#a#}" from acc_hyp have "(M, M0) \ ?R --> M + {#a#} \ ?W" .. from this and `(M, M0) \ ?R` have "M + {#a#} \ ?W" .. then show "N \ ?W" by (simp only: N) next fix K assume N: "N = M0 + K" assume "\b. b :# K --> (b, a) \ r" then have "M0 + K \ ?W" proof (induct K) case empty from M0 show "M0 + {#} \ ?W" by simp next case (add K x) from add.prems have "(x, a) \ r" by simp with wf_hyp have "\M \ ?W. M + {#x#} \ ?W" by blast moreover from add have "M0 + K \ ?W" by simp ultimately have "(M0 + K) + {#x#} \ ?W" .. then show "M0 + (K + {#x#}) \ ?W" by (simp only: union_assoc) qed then show "N \ ?W" by (simp only: N) qed qed } note tedious_reasoning = this assume wf: "wf r" fix M show "M \ ?W" proof (induct M) show "{#} \ ?W" proof (rule accI) fix b assume "(b, {#}) \ ?R" with not_less_empty show "b \ ?W" by contradiction qed fix M a assume "M \ ?W" from wf have "\M \ ?W. M + {#a#} \ ?W" proof induct fix a assume r: "!!b. (b, a) \ r ==> (\M \ ?W. M + {#b#} \ ?W)" show "\M \ ?W. M + {#a#} \ ?W" proof fix M assume "M \ ?W" then show "M + {#a#} \ ?W" by (rule acc_induct) (rule tedious_reasoning [OF _ r]) qed qed from this and `M \ ?W` show "M + {#a#} \ ?W" .. qed qed theorem wf_mult1: "wf r ==> wf (mult1 r)" by (rule acc_wfI) (rule all_accessible) theorem wf_mult: "wf r ==> wf (mult r)" unfolding mult_def by (rule wf_trancl) (rule wf_mult1) subsubsection {* Closure-free presentation *} (*Badly needed: a linear arithmetic procedure for multisets*) lemma diff_union_single_conv: "a :# J ==> I + J - {#a#} = I + (J - {#a#})" by (simp add: multiset_eq_conv_count_eq) text {* One direction. *} lemma mult_implies_one_step: "trans r ==> (M, N) \ mult r ==> \I J K. N = I + J \ M = I + K \ J \ {#} \ (\k \ set_of K. \j \ set_of J. (k, j) \ r)" apply (unfold mult_def mult1_def set_of_def) apply (erule converse_trancl_induct, clarify) apply (rule_tac x = M0 in exI, simp, clarify) apply (case_tac "a :# K") apply (rule_tac x = I in exI) apply (simp (no_asm)) apply (rule_tac x = "(K - {#a#}) + Ka" in exI) apply (simp (no_asm_simp) add: union_assoc [symmetric]) apply (drule_tac f = "\M. M - {#a#}" in arg_cong) apply (simp add: diff_union_single_conv) apply (simp (no_asm_use) add: trans_def) apply blast apply (subgoal_tac "a :# I") apply (rule_tac x = "I - {#a#}" in exI) apply (rule_tac x = "J + {#a#}" in exI) apply (rule_tac x = "K + Ka" in exI) apply (rule conjI) apply (simp add: multiset_eq_conv_count_eq split: nat_diff_split) apply (rule conjI) apply (drule_tac f = "\M. M - {#a#}" in arg_cong, simp) apply (simp add: multiset_eq_conv_count_eq split: nat_diff_split) apply (simp (no_asm_use) add: trans_def) apply blast apply (subgoal_tac "a :# (M0 + {#a#})") apply simp apply (simp (no_asm)) done lemma elem_imp_eq_diff_union: "a :# M ==> M = M - {#a#} + {#a#}" by (simp add: multiset_eq_conv_count_eq) lemma size_eq_Suc_imp_eq_union: "size M = Suc n ==> \a N. M = N + {#a#}" apply (erule size_eq_Suc_imp_elem [THEN exE]) apply (drule elem_imp_eq_diff_union, auto) done lemma one_step_implies_mult_aux: "trans r ==> \I J K. (size J = n \ J \ {#} \ (\k \ set_of K. \j \ set_of J. (k, j) \ r)) --> (I + K, I + J) \ mult r" apply (induct_tac n, auto) apply (frule size_eq_Suc_imp_eq_union, clarify) apply (rename_tac "J'", simp) apply (erule notE, auto) apply (case_tac "J' = {#}") apply (simp add: mult_def) apply (rule r_into_trancl) apply (simp add: mult1_def set_of_def, blast) txt {* Now we know @{term "J' \ {#}"}. *} apply (cut_tac M = K and P = "\x. (x, a) \ r" in multiset_partition) apply (erule_tac P = "\k \ set_of K. ?P k" in rev_mp) apply (erule ssubst) apply (simp add: Ball_def, auto) apply (subgoal_tac "((I + {# x :# K. (x, a) \ r #}) + {# x :# K. (x, a) \ r #}, (I + {# x :# K. (x, a) \ r #}) + J') \ mult r") prefer 2 apply force apply (simp (no_asm_use) add: union_assoc [symmetric] mult_def) apply (erule trancl_trans) apply (rule r_into_trancl) apply (simp add: mult1_def set_of_def) apply (rule_tac x = a in exI) apply (rule_tac x = "I + J'" in exI) apply (simp add: union_ac) done lemma one_step_implies_mult: "trans r ==> J \ {#} ==> \k \ set_of K. \j \ set_of J. (k, j) \ r ==> (I + K, I + J) \ mult r" using one_step_implies_mult_aux by blast subsubsection {* Partial-order properties *} instantiation multiset :: (order) order begin definition less_multiset_def [code del]: "M' < M \ (M', M) \ mult {(x', x). x' < x}" definition le_multiset_def [code del]: "M' <= M \ M' = M \ M' < (M::'a multiset)" lemma trans_base_order: "trans {(x', x). x' < (x::'a::order)}" unfolding trans_def by (blast intro: order_less_trans) text {* \medskip Irreflexivity. *} lemma mult_irrefl_aux: "finite A ==> (\x \ A. \y \ A. x < (y::'a::order)) \ A = {}" by (induct rule: finite_induct) (auto intro: order_less_trans) lemma mult_less_not_refl: "\ M < (M::'a::order multiset)" apply (unfold less_multiset_def, auto) apply (drule trans_base_order [THEN mult_implies_one_step], auto) apply (drule finite_set_of [THEN mult_irrefl_aux [rule_format (no_asm)]]) apply (simp add: set_of_eq_empty_iff) done lemma mult_less_irrefl [elim!]: "M < (M::'a::order multiset) ==> R" using insert mult_less_not_refl by fast text {* Transitivity. *} theorem mult_less_trans: "K < M ==> M < N ==> K < (N::'a::order multiset)" unfolding less_multiset_def mult_def by (blast intro: trancl_trans) text {* Asymmetry. *} theorem mult_less_not_sym: "M < N ==> \ N < (M::'a::order multiset)" apply auto apply (rule mult_less_not_refl [THEN notE]) apply (erule mult_less_trans, assumption) done theorem mult_less_asym: "M < N ==> (\ P ==> N < (M::'a::order multiset)) ==> P" using mult_less_not_sym by blast theorem mult_le_refl [iff]: "M <= (M::'a::order multiset)" unfolding le_multiset_def by auto text {* Anti-symmetry. *} theorem mult_le_antisym: "M <= N ==> N <= M ==> M = (N::'a::order multiset)" unfolding le_multiset_def by (blast dest: mult_less_not_sym) text {* Transitivity. *} theorem mult_le_trans: "K <= M ==> M <= N ==> K <= (N::'a::order multiset)" unfolding le_multiset_def by (blast intro: mult_less_trans) theorem mult_less_le: "(M < N) = (M <= N \ M \ (N::'a::order multiset))" unfolding le_multiset_def by auto instance proof qed (auto simp add: mult_less_le dest: mult_le_antisym elim: mult_le_trans) end subsubsection {* Monotonicity of multiset union *} lemma mult1_union: "(B, D) \ mult1 r ==> trans r ==> (C + B, C + D) \ mult1 r" apply (unfold mult1_def) apply auto apply (rule_tac x = a in exI) apply (rule_tac x = "C + M0" in exI) apply (simp add: union_assoc) done lemma union_less_mono2: "B < D ==> C + B < C + (D::'a::order multiset)" apply (unfold less_multiset_def mult_def) apply (erule trancl_induct) apply (blast intro: mult1_union transI order_less_trans r_into_trancl) apply (blast intro: mult1_union transI order_less_trans r_into_trancl trancl_trans) done lemma union_less_mono1: "B < D ==> B + C < D + (C::'a::order multiset)" apply (subst union_commute [of B C]) apply (subst union_commute [of D C]) apply (erule union_less_mono2) done lemma union_less_mono: "A < C ==> B < D ==> A + B < C + (D::'a::order multiset)" by (blast intro!: union_less_mono1 union_less_mono2 mult_less_trans) lemma union_le_mono: "A <= C ==> B <= D ==> A + B <= C + (D::'a::order multiset)" unfolding le_multiset_def by (blast intro: union_less_mono union_less_mono1 union_less_mono2) lemma empty_leI [iff]: "{#} <= (M::'a::order multiset)" apply (unfold le_multiset_def less_multiset_def) apply (case_tac "M = {#}") prefer 2 apply (subgoal_tac "({#} + {#}, {#} + M) \ mult (Collect (split op <))") prefer 2 apply (rule one_step_implies_mult) apply (simp only: trans_def) apply auto done lemma union_upper1: "A <= A + (B::'a::order multiset)" proof - have "A + {#} <= A + B" by (blast intro: union_le_mono) then show ?thesis by simp qed lemma union_upper2: "B <= A + (B::'a::order multiset)" by (subst union_commute) (rule union_upper1) instance multiset :: (order) pordered_ab_semigroup_add apply intro_classes apply (erule union_le_mono[OF mult_le_refl]) done subsection {* Link with lists *} primrec multiset_of :: "'a list \ 'a multiset" where "multiset_of [] = {#}" | "multiset_of (a # x) = multiset_of x + {# a #}" lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])" by (induct x) auto lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])" by (induct x) auto lemma set_of_multiset_of[simp]: "set_of(multiset_of x) = set x" by (induct x) auto lemma mem_set_multiset_eq: "x \ set xs = (x :# multiset_of xs)" by (induct xs) auto lemma multiset_of_append [simp]: "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" by (induct xs arbitrary: ys) (auto simp: union_ac) lemma surj_multiset_of: "surj multiset_of" apply (unfold surj_def) apply (rule allI) apply (rule_tac M = y in multiset_induct) apply auto apply (rule_tac x = "x # xa" in exI) apply auto done lemma set_count_greater_0: "set x = {a. count (multiset_of x) a > 0}" by (induct x) auto lemma distinct_count_atmost_1: "distinct x = (! a. count (multiset_of x) a = (if a \ set x then 1 else 0))" apply (induct x, simp, rule iffI, simp_all) apply (rule conjI) apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of) apply (erule_tac x = a in allE, simp, clarify) apply (erule_tac x = aa in allE, simp) done lemma multiset_of_eq_setD: "multiset_of xs = multiset_of ys \ set xs = set ys" by (rule) (auto simp add:multiset_eq_conv_count_eq set_count_greater_0) lemma set_eq_iff_multiset_of_eq_distinct: "distinct x \ distinct y \ (set x = set y) = (multiset_of x = multiset_of y)" by (auto simp: multiset_eq_conv_count_eq distinct_count_atmost_1) lemma set_eq_iff_multiset_of_remdups_eq: "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))" apply (rule iffI) apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1]) apply (drule distinct_remdups [THEN distinct_remdups [THEN set_eq_iff_multiset_of_eq_distinct [THEN iffD2]]]) apply simp done lemma multiset_of_compl_union [simp]: "multiset_of [x\xs. P x] + multiset_of [x\xs. \P x] = multiset_of xs" by (induct xs) (auto simp: union_ac) lemma count_filter: "count (multiset_of xs) x = length [y \ xs. y = x]" by (induct xs) auto lemma nth_mem_multiset_of: "i < length ls \ (ls ! i) :# multiset_of ls" apply (induct ls arbitrary: i) apply simp apply (case_tac i) apply auto done lemma multiset_of_remove1: "multiset_of (remove1 a xs) = multiset_of xs - {#a#}" by (induct xs) (auto simp add: multiset_eq_conv_count_eq) lemma multiset_of_eq_length: assumes "multiset_of xs = multiset_of ys" shows "length xs = length ys" using assms proof (induct arbitrary: ys rule: length_induct) case (1 xs ys) show ?case proof (cases xs) case Nil with "1.prems" show ?thesis by simp next case (Cons x xs') note xCons = Cons show ?thesis proof (cases ys) case Nil with "1.prems" Cons show ?thesis by simp next case (Cons y ys') have x_in_ys: "x = y \ x \ set ys'" proof (cases "x = y") case True then show ?thesis .. next case False from "1.prems" [symmetric] xCons Cons have "x :# multiset_of ys' + {#y#}" by simp with False show ?thesis by (simp add: mem_set_multiset_eq) qed from "1.hyps" have IH: "length xs' < length xs \ (\x. multiset_of xs' = multiset_of x \ length xs' = length x)" by blast from "1.prems" x_in_ys Cons xCons have "multiset_of xs' = multiset_of (remove1 x (y#ys'))" apply - apply (simp add: multiset_of_remove1, simp only: add_eq_conv_diff) apply fastsimp done with IH xCons have IH': "length xs' = length (remove1 x (y#ys'))" by fastsimp from x_in_ys have "x \ y \ length ys' > 0" by auto with Cons xCons x_in_ys IH' show ?thesis by (auto simp add: length_remove1) qed qed qed text {* This lemma shows which properties suffice to show that a function @{text "f"} with @{text "f xs = ys"} behaves like sort. *} lemma properties_for_sort: "multiset_of ys = multiset_of xs \ sorted ys \ sort xs = ys" proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) then have "x \ set ys" by (auto simp add: mem_set_multiset_eq intro!: ccontr) with Cons.prems Cons.hyps [of "remove1 x ys"] show ?case by (simp add: sorted_remove1 multiset_of_remove1 insort_remove1) qed subsection {* Pointwise ordering induced by count *} definition mset_le :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where [code del]: "A \# B \ (\a. count A a \ count B a)" definition mset_less :: "'a multiset \ 'a multiset \ bool" (infix "<#" 50) where [code del]: "A <# B \ A \# B \ A \ B" notation mset_le (infix "\#" 50) notation mset_less (infix "\#" 50) lemma mset_le_refl[simp]: "A \# A" unfolding mset_le_def by auto lemma mset_le_trans: "A \# B \ B \# C \ A \# C" unfolding mset_le_def by (fast intro: order_trans) lemma mset_le_antisym: "A \# B \ B \# A \ A = B" apply (unfold mset_le_def) apply (rule multiset_eq_conv_count_eq [THEN iffD2]) apply (blast intro: order_antisym) done lemma mset_le_exists_conv: "(A \# B) = (\C. B = A + C)" apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI) apply (auto intro: multiset_eq_conv_count_eq [THEN iffD2]) done lemma mset_le_mono_add_right_cancel[simp]: "(A + C \# B + C) = (A \# B)" unfolding mset_le_def by auto lemma mset_le_mono_add_left_cancel[simp]: "(C + A \# C + B) = (A \# B)" unfolding mset_le_def by auto lemma mset_le_mono_add: "\ A \# B; C \# D \ \ A + C \# B + D" apply (unfold mset_le_def) apply auto apply (erule_tac x = a in allE)+ apply auto done lemma mset_le_add_left[simp]: "A \# A + B" unfolding mset_le_def by auto lemma mset_le_add_right[simp]: "B \# A + B" unfolding mset_le_def by auto lemma mset_le_single: "a :# B \ {#a#} \# B" by (simp add: mset_le_def) lemma multiset_diff_union_assoc: "C \# B \ A + B - C = A + (B - C)" by (simp add: multiset_eq_conv_count_eq mset_le_def) lemma mset_le_multiset_union_diff_commute: assumes "B \# A" shows "A - B + C = A + C - B" proof - from mset_le_exists_conv [of "B" "A"] assms have "\D. A = B + D" .. from this obtain D where "A = B + D" .. then show ?thesis apply simp apply (subst union_commute) apply (subst multiset_diff_union_assoc) apply simp apply (simp add: diff_cancel) apply (subst union_assoc) apply (subst union_commute[of "B" _]) apply (subst multiset_diff_union_assoc) apply simp apply (simp add: diff_cancel) done qed lemma multiset_of_remdups_le: "multiset_of (remdups xs) \# multiset_of xs" apply (induct xs) apply auto apply (rule mset_le_trans) apply auto done lemma multiset_of_update: "i < length ls \ multiset_of (ls[i := v]) = multiset_of ls - {#ls ! i#} + {#v#}" proof (induct ls arbitrary: i) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases i) case 0 then show ?thesis by simp next case (Suc i') with Cons show ?thesis apply simp apply (subst union_assoc) apply (subst union_commute [where M = "{#v#}" and N = "{#x#}"]) apply (subst union_assoc [symmetric]) apply simp apply (rule mset_le_multiset_union_diff_commute) apply (simp add: mset_le_single nth_mem_multiset_of) done qed qed lemma multiset_of_swap: "i < length ls \ j < length ls \ multiset_of (ls[j := ls ! i, i := ls ! j]) = multiset_of ls" apply (case_tac "i = j") apply simp apply (simp add: multiset_of_update) apply (subst elem_imp_eq_diff_union[symmetric]) apply (simp add: nth_mem_multiset_of) apply simp done -interpretation mset_order!: order "op \#" "op <#" +interpretation mset_order: order "op \#" "op <#" proof qed (auto intro: order.intro mset_le_refl mset_le_antisym mset_le_trans simp: mset_less_def) -interpretation mset_order_cancel_semigroup!: +interpretation mset_order_cancel_semigroup: pordered_cancel_ab_semigroup_add "op +" "op \#" "op <#" proof qed (erule mset_le_mono_add [OF mset_le_refl]) -interpretation mset_order_semigroup_cancel!: +interpretation mset_order_semigroup_cancel: pordered_ab_semigroup_add_imp_le "op +" "op \#" "op <#" proof qed simp lemma mset_lessD: "A \# B \ x \# A \ x \# B" apply (clarsimp simp: mset_le_def mset_less_def) apply (erule_tac x=x in allE) apply auto done lemma mset_leD: "A \# B \ x \# A \ x \# B" apply (clarsimp simp: mset_le_def mset_less_def) apply (erule_tac x = x in allE) apply auto done lemma mset_less_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" apply (rule conjI) apply (simp add: mset_lessD) apply (clarsimp simp: mset_le_def mset_less_def) apply safe apply (erule_tac x = a in allE) apply (auto split: split_if_asm) done lemma mset_le_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" apply (rule conjI) apply (simp add: mset_leD) apply (force simp: mset_le_def mset_less_def split: split_if_asm) done lemma mset_less_of_empty[simp]: "A \# {#} = False" by (induct A) (auto simp: mset_le_def mset_less_def) lemma multi_psub_of_add_self[simp]: "A \# A + {#x#}" by (auto simp: mset_le_def mset_less_def) lemma multi_psub_self[simp]: "A \# A = False" by (auto simp: mset_le_def mset_less_def) lemma mset_less_add_bothsides: "T + {#x#} \# S + {#x#} \ T \# S" by (auto simp: mset_le_def mset_less_def) lemma mset_less_empty_nonempty: "({#} \# S) = (S \ {#})" by (auto simp: mset_le_def mset_less_def) lemma mset_less_size: "A \# B \ size A < size B" proof (induct A arbitrary: B) case (empty M) then have "M \ {#}" by (simp add: mset_less_empty_nonempty) then obtain M' x where "M = M' + {#x#}" by (blast dest: multi_nonempty_split) then show ?case by simp next case (add S x T) have IH: "\B. S \# B \ size S < size B" by fact have SxsubT: "S + {#x#} \# T" by fact then have "x \# T" and "S \# T" by (auto dest: mset_less_insertD) then obtain T' where T: "T = T' + {#x#}" by (blast dest: multi_member_split) then have "S \# T'" using SxsubT by (blast intro: mset_less_add_bothsides) then have "size S < size T'" using IH by simp then show ?case using T by simp qed lemmas mset_less_trans = mset_order.less_trans lemma mset_less_diff_self: "c \# B \ B - {#c#} \# B" by (auto simp: mset_le_def mset_less_def multi_drop_mem_not_eq) subsection {* Strong induction and subset induction for multisets *} text {* Well-foundedness of proper subset operator: *} text {* proper multiset subset *} definition mset_less_rel :: "('a multiset * 'a multiset) set" where "mset_less_rel = {(A,B). A \# B}" lemma multiset_add_sub_el_shuffle: assumes "c \# B" and "b \ c" shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}" proof - from `c \# B` obtain A where B: "B = A + {#c#}" by (blast dest: multi_member_split) have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" by (simp add: union_ac) then show ?thesis using B by simp qed lemma wf_mset_less_rel: "wf mset_less_rel" apply (unfold mset_less_rel_def) apply (rule wf_measure [THEN wf_subset, where f1=size]) apply (clarsimp simp: measure_def inv_image_def mset_less_size) done text {* The induction rules: *} lemma full_multiset_induct [case_names less]: assumes ih: "\B. \A. A \# B \ P A \ P B" shows "P B" apply (rule wf_mset_less_rel [THEN wf_induct]) apply (rule ih, auto simp: mset_less_rel_def) done lemma multi_subset_induct [consumes 2, case_names empty add]: assumes "F \# A" and empty: "P {#}" and insert: "\a F. a \# A \ P F \ P (F + {#a#})" shows "P F" proof - from `F \# A` show ?thesis proof (induct F) show "P {#}" by fact next fix x F assume P: "F \# A \ P F" and i: "F + {#x#} \# A" show "P (F + {#x#})" proof (rule insert) from i show "x \# A" by (auto dest: mset_le_insertD) from i have "F \# A" by (auto dest: mset_le_insertD) with P show "P F" . qed qed qed text{* A consequence: Extensionality. *} lemma multi_count_eq: "(\x. count A x = count B x) = (A = B)" apply (rule iffI) prefer 2 apply clarsimp apply (induct A arbitrary: B rule: full_multiset_induct) apply (rename_tac C) apply (case_tac B rule: multiset_cases) apply (simp add: empty_multiset_count) apply simp apply (case_tac "x \# C") apply (force dest: multi_member_split) apply (erule_tac x = x in allE) apply simp done lemmas multi_count_ext = multi_count_eq [THEN iffD1, rule_format] subsection {* The fold combinator *} text {* The intended behaviour is @{text "fold_mset f z {#x\<^isub>1, ..., x\<^isub>n#} = f x\<^isub>1 (\ (f x\<^isub>n z)\)"} if @{text f} is associative-commutative. *} text {* The graph of @{text "fold_mset"}, @{text "z"}: the start element, @{text "f"}: folding function, @{text "A"}: the multiset, @{text "y"}: the result. *} inductive fold_msetG :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b \ bool" for f :: "'a \ 'b \ 'b" and z :: 'b where emptyI [intro]: "fold_msetG f z {#} z" | insertI [intro]: "fold_msetG f z A y \ fold_msetG f z (A + {#x#}) (f x y)" inductive_cases empty_fold_msetGE [elim!]: "fold_msetG f z {#} x" inductive_cases insert_fold_msetGE: "fold_msetG f z (A + {#}) y" definition fold_mset :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b" where "fold_mset f z A = (THE x. fold_msetG f z A x)" lemma Diff1_fold_msetG: "fold_msetG f z (A - {#x#}) y \ x \# A \ fold_msetG f z A (f x y)" apply (frule_tac x = x in fold_msetG.insertI) apply auto done lemma fold_msetG_nonempty: "\x. fold_msetG f z A x" apply (induct A) apply blast apply clarsimp apply (drule_tac x = x in fold_msetG.insertI) apply auto done lemma fold_mset_empty[simp]: "fold_mset f z {#} = z" unfolding fold_mset_def by blast locale left_commutative = fixes f :: "'a => 'b => 'b" assumes left_commute: "f x (f y z) = f y (f x z)" begin lemma fold_msetG_determ: "fold_msetG f z A x \ fold_msetG f z A y \ y = x" proof (induct arbitrary: x y z rule: full_multiset_induct) case (less M x\<^isub>1 x\<^isub>2 Z) have IH: "\A. A \# M \ (\x x' x''. fold_msetG f x'' A x \ fold_msetG f x'' A x' \ x' = x)" by fact have Mfoldx\<^isub>1: "fold_msetG f Z M x\<^isub>1" and Mfoldx\<^isub>2: "fold_msetG f Z M x\<^isub>2" by fact+ show ?case proof (rule fold_msetG.cases [OF Mfoldx\<^isub>1]) assume "M = {#}" and "x\<^isub>1 = Z" then show ?case using Mfoldx\<^isub>2 by auto next fix B b u assume "M = B + {#b#}" and "x\<^isub>1 = f b u" and Bu: "fold_msetG f Z B u" then have MBb: "M = B + {#b#}" and x\<^isub>1: "x\<^isub>1 = f b u" by auto show ?case proof (rule fold_msetG.cases [OF Mfoldx\<^isub>2]) assume "M = {#}" "x\<^isub>2 = Z" then show ?case using Mfoldx\<^isub>1 by auto next fix C c v assume "M = C + {#c#}" and "x\<^isub>2 = f c v" and Cv: "fold_msetG f Z C v" then have MCc: "M = C + {#c#}" and x\<^isub>2: "x\<^isub>2 = f c v" by auto then have CsubM: "C \# M" by simp from MBb have BsubM: "B \# M" by simp show ?case proof cases assume "b=c" then moreover have "B = C" using MBb MCc by auto ultimately show ?thesis using Bu Cv x\<^isub>1 x\<^isub>2 CsubM IH by auto next assume diff: "b \ c" let ?D = "B - {#c#}" have cinB: "c \# B" and binC: "b \# C" using MBb MCc diff by (auto intro: insert_noteq_member dest: sym) have "B - {#c#} \# B" using cinB by (rule mset_less_diff_self) then have DsubM: "?D \# M" using BsubM by (blast intro: mset_less_trans) from MBb MCc have "B + {#b#} = C + {#c#}" by blast then have [simp]: "B + {#b#} - {#c#} = C" using MBb MCc binC cinB by auto have B: "B = ?D + {#c#}" and C: "C = ?D + {#b#}" using MBb MCc diff binC cinB by (auto simp: multiset_add_sub_el_shuffle) then obtain d where Dfoldd: "fold_msetG f Z ?D d" using fold_msetG_nonempty by iprover then have "fold_msetG f Z B (f c d)" using cinB by (rule Diff1_fold_msetG) then have "f c d = u" using IH BsubM Bu by blast moreover have "fold_msetG f Z C (f b d)" using binC cinB diff Dfoldd by (auto simp: multiset_add_sub_el_shuffle dest: fold_msetG.insertI [where x=b]) then have "f b d = v" using IH CsubM Cv by blast ultimately show ?thesis using x\<^isub>1 x\<^isub>2 by (auto simp: left_commute) qed qed qed qed lemma fold_mset_insert_aux: "(fold_msetG f z (A + {#x#}) v) = (\y. fold_msetG f z A y \ v = f x y)" apply (rule iffI) prefer 2 apply blast apply (rule_tac A=A and f=f in fold_msetG_nonempty [THEN exE, standard]) apply (blast intro: fold_msetG_determ) done lemma fold_mset_equality: "fold_msetG f z A y \ fold_mset f z A = y" unfolding fold_mset_def by (blast intro: fold_msetG_determ) lemma fold_mset_insert: "fold_mset f z (A + {#x#}) = f x (fold_mset f z A)" apply (simp add: fold_mset_def fold_mset_insert_aux union_commute) apply (rule the_equality) apply (auto cong add: conj_cong simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty) done lemma fold_mset_insert_idem: "fold_mset f z (A + {#a#}) = f a (fold_mset f z A)" apply (simp add: fold_mset_def fold_mset_insert_aux) apply (rule the_equality) apply (auto cong add: conj_cong simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty) done lemma fold_mset_commute: "f x (fold_mset f z A) = fold_mset f (f x z) A" by (induct A) (auto simp: fold_mset_insert left_commute [of x]) lemma fold_mset_single [simp]: "fold_mset f z {#x#} = f x z" using fold_mset_insert [of z "{#}"] by simp lemma fold_mset_union [simp]: "fold_mset f z (A+B) = fold_mset f (fold_mset f z A) B" proof (induct A) case empty then show ?case by simp next case (add A x) have "A + {#x#} + B = (A+B) + {#x#}" by(simp add:union_ac) then have "fold_mset f z (A + {#x#} + B) = f x (fold_mset f z (A + B))" by (simp add: fold_mset_insert) also have "\ = fold_mset f (fold_mset f z (A + {#x#})) B" by (simp add: fold_mset_commute[of x,symmetric] add fold_mset_insert) finally show ?case . qed lemma fold_mset_fusion: assumes "left_commutative g" shows "(\x y. h (g x y) = f x (h y)) \ h (fold_mset g w A) = fold_mset f (h w) A" (is "PROP ?P") proof - interpret left_commutative g by fact show "PROP ?P" by (induct A) auto qed lemma fold_mset_rec: assumes "a \# A" shows "fold_mset f z A = f a (fold_mset f z (A - {#a#}))" proof - from assms obtain A' where "A = A' + {#a#}" by (blast dest: multi_member_split) then show ?thesis by simp qed end text {* A note on code generation: When defining some function containing a subterm @{term"fold_mset F"}, code generation is not automatic. When interpreting locale @{text left_commutative} with @{text F}, the would be code thms for @{const fold_mset} become thms like @{term"fold_mset F z {#} = z"} where @{text F} is not a pattern but contains defined symbols, i.e.\ is not a code thm. Hence a separate constant with its own code thms needs to be introduced for @{text F}. See the image operator below. *} subsection {* Image *} definition [code del]: "image_mset f = fold_mset (op + o single o f) {#}" -interpretation image_left_comm!: left_commutative "op + o single o f" +interpretation image_left_comm: left_commutative "op + o single o f" proof qed (simp add:union_ac) lemma image_mset_empty [simp]: "image_mset f {#} = {#}" by (simp add: image_mset_def) lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}" by (simp add: image_mset_def) lemma image_mset_insert: "image_mset f (M + {#a#}) = image_mset f M + {#f a#}" by (simp add: image_mset_def add_ac) lemma image_mset_union [simp]: "image_mset f (M+N) = image_mset f M + image_mset f N" apply (induct N) apply simp apply (simp add: union_assoc [symmetric] image_mset_insert) done lemma size_image_mset [simp]: "size (image_mset f M) = size M" by (induct M) simp_all lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \ M = {#}" by (cases M) auto syntax comprehension1_mset :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ :# _#})") translations "{#e. x:#M#}" == "CONST image_mset (%x. e) M" syntax comprehension2_mset :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ :# _./ _#})") translations "{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}" text {* This allows to write not just filters like @{term "{#x:#M. x# XS \ x \# {# y #} + XS" and multi_member_this: "x \# {# x #} + XS" and multi_member_last: "x \# {# x #}" by auto definition "ms_strict = mult pair_less" definition [code del]: "ms_weak = ms_strict \ Id" lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)" unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def by (auto intro: wf_mult1 wf_trancl simp: mult_def) lemma smsI: "(set_of A, set_of B) \ max_strict \ (Z + A, Z + B) \ ms_strict" unfolding ms_strict_def by (rule one_step_implies_mult) (auto simp add: max_strict_def pair_less_def elim!:max_ext.cases) lemma wmsI: "(set_of A, set_of B) \ max_strict \ A = {#} \ B = {#} \ (Z + A, Z + B) \ ms_weak" unfolding ms_weak_def ms_strict_def by (auto simp add: pair_less_def max_strict_def elim!:max_ext.cases intro: one_step_implies_mult) inductive pw_leq where pw_leq_empty: "pw_leq {#} {#}" | pw_leq_step: "\(x,y) \ pair_leq; pw_leq X Y \ \ pw_leq ({#x#} + X) ({#y#} + Y)" lemma pw_leq_lstep: "(x, y) \ pair_leq \ pw_leq {#x#} {#y#}" by (drule pw_leq_step) (rule pw_leq_empty, simp) lemma pw_leq_split: assumes "pw_leq X Y" shows "\A B Z. X = A + Z \ Y = B + Z \ ((set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#}))" using assms proof (induct) case pw_leq_empty thus ?case by auto next case (pw_leq_step x y X Y) then obtain A B Z where [simp]: "X = A + Z" "Y = B + Z" and 1[simp]: "(set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#})" by auto from pw_leq_step have "x = y \ (x, y) \ pair_less" unfolding pair_leq_def by auto thus ?case proof assume [simp]: "x = y" have "{#x#} + X = A + ({#y#}+Z) \ {#y#} + Y = B + ({#y#}+Z) \ ((set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#}))" by (auto simp: add_ac) thus ?case by (intro exI) next assume A: "(x, y) \ pair_less" let ?A' = "{#x#} + A" and ?B' = "{#y#} + B" have "{#x#} + X = ?A' + Z" "{#y#} + Y = ?B' + Z" by (auto simp add: add_ac) moreover have "(set_of ?A', set_of ?B') \ max_strict" using 1 A unfolding max_strict_def by (auto elim!: max_ext.cases) ultimately show ?thesis by blast qed qed lemma assumes pwleq: "pw_leq Z Z'" shows ms_strictI: "(set_of A, set_of B) \ max_strict \ (Z + A, Z' + B) \ ms_strict" and ms_weakI1: "(set_of A, set_of B) \ max_strict \ (Z + A, Z' + B) \ ms_weak" and ms_weakI2: "(Z + {#}, Z' + {#}) \ ms_weak" proof - from pw_leq_split[OF pwleq] obtain A' B' Z'' where [simp]: "Z = A' + Z''" "Z' = B' + Z''" and mx_or_empty: "(set_of A', set_of B') \ max_strict \ (A' = {#} \ B' = {#})" by blast { assume max: "(set_of A, set_of B) \ max_strict" from mx_or_empty have "(Z'' + (A + A'), Z'' + (B + B')) \ ms_strict" proof assume max': "(set_of A', set_of B') \ max_strict" with max have "(set_of (A + A'), set_of (B + B')) \ max_strict" by (auto simp: max_strict_def intro: max_ext_additive) thus ?thesis by (rule smsI) next assume [simp]: "A' = {#} \ B' = {#}" show ?thesis by (rule smsI) (auto intro: max) qed thus "(Z + A, Z' + B) \ ms_strict" by (simp add:add_ac) thus "(Z + A, Z' + B) \ ms_weak" by (simp add: ms_weak_def) } from mx_or_empty have "(Z'' + A', Z'' + B') \ ms_weak" by (rule wmsI) thus "(Z + {#}, Z' + {#}) \ ms_weak" by (simp add:add_ac) qed lemma empty_idemp: "{#} + x = x" "x + {#} = x" and nonempty_plus: "{# x #} + rs \ {#}" and nonempty_single: "{# x #} \ {#}" by auto setup {* let fun msetT T = Type ("Multiset.multiset", [T]); fun mk_mset T [] = Const (@{const_name Mempty}, msetT T) | mk_mset T [x] = Const (@{const_name single}, T --> msetT T) $ x | mk_mset T (x :: xs) = Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $ mk_mset T [x] $ mk_mset T xs fun mset_member_tac m i = (if m <= 0 then rtac @{thm multi_member_this} i ORELSE rtac @{thm multi_member_last} i else rtac @{thm multi_member_skip} i THEN mset_member_tac (m - 1) i) val mset_nonempty_tac = rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single} val regroup_munion_conv = FundefLib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus} (map (fn t => t RS eq_reflection) (@{thms union_ac} @ @{thms empty_idemp})) fun unfold_pwleq_tac i = (rtac @{thm pw_leq_step} i THEN (fn st => unfold_pwleq_tac (i + 1) st)) ORELSE (rtac @{thm pw_leq_lstep} i) ORELSE (rtac @{thm pw_leq_empty} i) val set_of_simps = [@{thm set_of_empty}, @{thm set_of_single}, @{thm set_of_union}, @{thm Un_insert_left}, @{thm Un_empty_left}] in ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset { msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_of_simps, smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1}, reduction_pair= @{thm ms_reduction_pair} }) end *} end diff --git a/src/HOL/Library/Numeral_Type.thy b/src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy +++ b/src/HOL/Library/Numeral_Type.thy @@ -1,461 +1,461 @@ (* Title: HOL/Library/Numeral_Type.thy Author: Brian Huffman *) header {* Numeral Syntax for Types *} theory Numeral_Type imports Main begin subsection {* Preliminary lemmas *} (* These should be moved elsewhere *) lemma (in type_definition) univ: "UNIV = Abs ` A" proof show "Abs ` A \ UNIV" by (rule subset_UNIV) show "UNIV \ Abs ` A" proof fix x :: 'b have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) moreover have "Rep x \ A" by (rule Rep) ultimately show "x \ Abs ` A" by (rule image_eqI) qed qed lemma (in type_definition) card: "card (UNIV :: 'b set) = card A" by (simp add: univ card_image inj_on_def Abs_inject) subsection {* Cardinalities of types *} syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))") translations "CARD(t)" => "CONST card (CONST UNIV \ t set)" typed_print_translation {* let fun card_univ_tr' show_sorts _ [Const (@{const_syntax UNIV}, Type(_,[T,_]))] = Syntax.const "_type_card" $ Syntax.term_of_typ show_sorts T; in [(@{const_syntax card}, card_univ_tr')] end *} lemma card_unit [simp]: "CARD(unit) = 1" unfolding UNIV_unit by simp lemma card_bool [simp]: "CARD(bool) = 2" unfolding UNIV_bool by simp lemma card_prod [simp]: "CARD('a \ 'b) = CARD('a::finite) * CARD('b::finite)" unfolding UNIV_Times_UNIV [symmetric] by (simp only: card_cartesian_product) lemma card_sum [simp]: "CARD('a + 'b) = CARD('a::finite) + CARD('b::finite)" unfolding UNIV_Plus_UNIV [symmetric] by (simp only: finite card_Plus) lemma card_option [simp]: "CARD('a option) = Suc CARD('a::finite)" unfolding insert_None_conv_UNIV [symmetric] apply (subgoal_tac "(None::'a option) \ range Some") apply (simp add: card_image) apply fast done lemma card_set [simp]: "CARD('a set) = 2 ^ CARD('a::finite)" unfolding Pow_UNIV [symmetric] by (simp only: card_Pow finite numeral_2_eq_2) lemma card_nat [simp]: "CARD(nat) = 0" by (simp add: infinite_UNIV_nat card_eq_0_iff) subsection {* Classes with at least 1 and 2 *} text {* Class finite already captures "at least 1" *} lemma zero_less_card_finite [simp]: "0 < CARD('a::finite)" unfolding neq0_conv [symmetric] by simp lemma one_le_card_finite [simp]: "Suc 0 \ CARD('a::finite)" by (simp add: less_Suc_eq_le [symmetric]) text {* Class for cardinality "at least 2" *} class card2 = finite + assumes two_le_card: "2 \ CARD('a)" lemma one_less_card: "Suc 0 < CARD('a::card2)" using two_le_card [where 'a='a] by simp lemma one_less_int_card: "1 < int CARD('a::card2)" using one_less_card [where 'a='a] by simp subsection {* Numeral Types *} typedef (open) num0 = "UNIV :: nat set" .. typedef (open) num1 = "UNIV :: unit set" .. typedef (open) 'a bit0 = "{0 ..< 2 * int CARD('a::finite)}" proof show "0 \ {0 ..< 2 * int CARD('a)}" by simp qed typedef (open) 'a bit1 = "{0 ..< 1 + 2 * int CARD('a::finite)}" proof show "0 \ {0 ..< 1 + 2 * int CARD('a)}" by simp qed lemma card_num0 [simp]: "CARD (num0) = 0" unfolding type_definition.card [OF type_definition_num0] by simp lemma card_num1 [simp]: "CARD(num1) = 1" unfolding type_definition.card [OF type_definition_num1] by (simp only: card_unit) lemma card_bit0 [simp]: "CARD('a bit0) = 2 * CARD('a::finite)" unfolding type_definition.card [OF type_definition_bit0] by simp lemma card_bit1 [simp]: "CARD('a bit1) = Suc (2 * CARD('a::finite))" unfolding type_definition.card [OF type_definition_bit1] by simp instance num1 :: finite proof show "finite (UNIV::num1 set)" unfolding type_definition.univ [OF type_definition_num1] using finite by (rule finite_imageI) qed instance bit0 :: (finite) card2 proof show "finite (UNIV::'a bit0 set)" unfolding type_definition.univ [OF type_definition_bit0] by simp show "2 \ CARD('a bit0)" by simp qed instance bit1 :: (finite) card2 proof show "finite (UNIV::'a bit1 set)" unfolding type_definition.univ [OF type_definition_bit1] by simp show "2 \ CARD('a bit1)" by simp qed subsection {* Locale for modular arithmetic subtypes *} locale mod_type = fixes n :: int and Rep :: "'a::{zero,one,plus,times,uminus,minus,power} \ int" and Abs :: "int \ 'a::{zero,one,plus,times,uminus,minus,power}" assumes type: "type_definition Rep Abs {0.. n" by (rule Rep_less_n [THEN order_less_imp_le]) lemma Rep_inject_sym: "x = y \ Rep x = Rep y" by (rule type_definition.Rep_inject [OF type, symmetric]) lemma Rep_inverse: "Abs (Rep x) = x" by (rule type_definition.Rep_inverse [OF type]) lemma Abs_inverse: "m \ {0.. Rep (Abs m) = m" by (rule type_definition.Abs_inverse [OF type]) lemma Rep_Abs_mod: "Rep (Abs (m mod n)) = m mod n" by (simp add: Abs_inverse IntDiv.pos_mod_conj [OF size0]) lemma Rep_Abs_0: "Rep (Abs 0) = 0" by (simp add: Abs_inverse size0) lemma Rep_0: "Rep 0 = 0" by (simp add: zero_def Rep_Abs_0) lemma Rep_Abs_1: "Rep (Abs 1) = 1" by (simp add: Abs_inverse size1) lemma Rep_1: "Rep 1 = 1" by (simp add: one_def Rep_Abs_1) lemma Rep_mod: "Rep x mod n = Rep x" apply (rule_tac x=x in type_definition.Abs_cases [OF type]) apply (simp add: type_definition.Abs_inverse [OF type]) apply (simp add: mod_pos_pos_trivial) done lemmas Rep_simps = Rep_inject_sym Rep_inverse Rep_Abs_mod Rep_mod Rep_Abs_0 Rep_Abs_1 lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)" apply (intro_classes, unfold definitions) apply (simp_all add: Rep_simps zmod_simps ring_simps) done lemma recpower: "OFCLASS('a, recpower_class)" apply (intro_classes, unfold definitions) apply (simp_all add: Rep_simps zmod_simps add_ac mult_assoc mod_pos_pos_trivial size1) done end locale mod_ring = mod_type + constrains n :: int and Rep :: "'a::{number_ring,power} \ int" and Abs :: "int \ 'a::{number_ring,power}" begin lemma of_nat_eq: "of_nat k = Abs (int k mod n)" apply (induct k) apply (simp add: zero_def) apply (simp add: Rep_simps add_def one_def zmod_simps add_ac) done lemma of_int_eq: "of_int z = Abs (z mod n)" apply (cases z rule: int_diff_cases) apply (simp add: Rep_simps of_nat_eq diff_def zmod_simps) done lemma Rep_number_of: "Rep (number_of w) = number_of w mod n" by (simp add: number_of_eq of_int_eq Rep_Abs_mod) lemma iszero_number_of: "iszero (number_of w::'a) \ number_of w mod n = 0" by (simp add: Rep_simps number_of_eq of_int_eq iszero_def zero_def) lemma cases: assumes 1: "\z. \(x::'a) = of_int z; 0 \ z; z < n\ \ P" shows "P" apply (cases x rule: type_definition.Abs_cases [OF type]) apply (rule_tac z="y" in 1) apply (simp_all add: of_int_eq mod_pos_pos_trivial) done lemma induct: "(\z. \0 \ z; z < n\ \ P (of_int z)) \ P (x::'a)" by (cases x rule: cases) simp end subsection {* Number ring instances *} text {* Unfortunately a number ring instance is not possible for @{typ num1}, since 0 and 1 are not distinct. *} instantiation num1 :: "{comm_ring,comm_monoid_mult,number,recpower}" begin lemma num1_eq_iff: "(x::num1) = (y::num1) \ True" by (induct x, induct y) simp instance proof qed (simp_all add: num1_eq_iff) end instantiation bit0 and bit1 :: (finite) "{zero,one,plus,times,uminus,minus,power}" begin definition Abs_bit0' :: "int \ 'a bit0" where "Abs_bit0' x = Abs_bit0 (x mod int CARD('a bit0))" definition Abs_bit1' :: "int \ 'a bit1" where "Abs_bit1' x = Abs_bit1 (x mod int CARD('a bit1))" definition "0 = Abs_bit0 0" definition "1 = Abs_bit0 1" definition "x + y = Abs_bit0' (Rep_bit0 x + Rep_bit0 y)" definition "x * y = Abs_bit0' (Rep_bit0 x * Rep_bit0 y)" definition "x - y = Abs_bit0' (Rep_bit0 x - Rep_bit0 y)" definition "- x = Abs_bit0' (- Rep_bit0 x)" definition "x ^ k = Abs_bit0' (Rep_bit0 x ^ k)" definition "0 = Abs_bit1 0" definition "1 = Abs_bit1 1" definition "x + y = Abs_bit1' (Rep_bit1 x + Rep_bit1 y)" definition "x * y = Abs_bit1' (Rep_bit1 x * Rep_bit1 y)" definition "x - y = Abs_bit1' (Rep_bit1 x - Rep_bit1 y)" definition "- x = Abs_bit1' (- Rep_bit1 x)" definition "x ^ k = Abs_bit1' (Rep_bit1 x ^ k)" instance .. end -interpretation bit0!: +interpretation bit0: mod_type "int CARD('a::finite bit0)" "Rep_bit0 :: 'a::finite bit0 \ int" "Abs_bit0 :: int \ 'a::finite bit0" apply (rule mod_type.intro) apply (simp add: int_mult type_definition_bit0) apply (rule one_less_int_card) apply (rule zero_bit0_def) apply (rule one_bit0_def) apply (rule plus_bit0_def [unfolded Abs_bit0'_def]) apply (rule times_bit0_def [unfolded Abs_bit0'_def]) apply (rule minus_bit0_def [unfolded Abs_bit0'_def]) apply (rule uminus_bit0_def [unfolded Abs_bit0'_def]) apply (rule power_bit0_def [unfolded Abs_bit0'_def]) done -interpretation bit1!: +interpretation bit1: mod_type "int CARD('a::finite bit1)" "Rep_bit1 :: 'a::finite bit1 \ int" "Abs_bit1 :: int \ 'a::finite bit1" apply (rule mod_type.intro) apply (simp add: int_mult type_definition_bit1) apply (rule one_less_int_card) apply (rule zero_bit1_def) apply (rule one_bit1_def) apply (rule plus_bit1_def [unfolded Abs_bit1'_def]) apply (rule times_bit1_def [unfolded Abs_bit1'_def]) apply (rule minus_bit1_def [unfolded Abs_bit1'_def]) apply (rule uminus_bit1_def [unfolded Abs_bit1'_def]) apply (rule power_bit1_def [unfolded Abs_bit1'_def]) done instance bit0 :: (finite) "{comm_ring_1,recpower}" by (rule bit0.comm_ring_1 bit0.recpower)+ instance bit1 :: (finite) "{comm_ring_1,recpower}" by (rule bit1.comm_ring_1 bit1.recpower)+ instantiation bit0 and bit1 :: (finite) number_ring begin definition "(number_of w :: _ bit0) = of_int w" definition "(number_of w :: _ bit1) = of_int w" instance proof qed (rule number_of_bit0_def number_of_bit1_def)+ end -interpretation bit0!: +interpretation bit0: mod_ring "int CARD('a::finite bit0)" "Rep_bit0 :: 'a::finite bit0 \ int" "Abs_bit0 :: int \ 'a::finite bit0" .. -interpretation bit1!: +interpretation bit1: mod_ring "int CARD('a::finite bit1)" "Rep_bit1 :: 'a::finite bit1 \ int" "Abs_bit1 :: int \ 'a::finite bit1" .. text {* Set up cases, induction, and arithmetic *} lemmas bit0_cases [case_names of_int, cases type: bit0] = bit0.cases lemmas bit1_cases [case_names of_int, cases type: bit1] = bit1.cases lemmas bit0_induct [case_names of_int, induct type: bit0] = bit0.induct lemmas bit1_induct [case_names of_int, induct type: bit1] = bit1.induct lemmas bit0_iszero_number_of [simp] = bit0.iszero_number_of lemmas bit1_iszero_number_of [simp] = bit1.iszero_number_of declare power_Suc [where ?'a="'a::finite bit0", standard, simp] declare power_Suc [where ?'a="'a::finite bit1", standard, simp] subsection {* Syntax *} syntax "_NumeralType" :: "num_const => type" ("_") "_NumeralType0" :: type ("0") "_NumeralType1" :: type ("1") translations "_NumeralType1" == (type) "num1" "_NumeralType0" == (type) "num0" parse_translation {* let val num1_const = Syntax.const "Numeral_Type.num1"; val num0_const = Syntax.const "Numeral_Type.num0"; val B0_const = Syntax.const "Numeral_Type.bit0"; val B1_const = Syntax.const "Numeral_Type.bit1"; fun mk_bintype n = let fun mk_bit n = if n = 0 then B0_const else B1_const; fun bin_of n = if n = 1 then num1_const else if n = 0 then num0_const else if n = ~1 then raise TERM ("negative type numeral", []) else let val (q, r) = Integer.div_mod n 2; in mk_bit r $ bin_of q end; in bin_of n end; fun numeral_tr (*"_NumeralType"*) [Const (str, _)] = mk_bintype (valOf (Int.fromString str)) | numeral_tr (*"_NumeralType"*) ts = raise TERM ("numeral_tr", ts); in [("_NumeralType", numeral_tr)] end; *} print_translation {* let fun int_of [] = 0 | int_of (b :: bs) = b + 2 * int_of bs; fun bin_of (Const ("num0", _)) = [] | bin_of (Const ("num1", _)) = [1] | bin_of (Const ("bit0", _) $ bs) = 0 :: bin_of bs | bin_of (Const ("bit1", _) $ bs) = 1 :: bin_of bs | bin_of t = raise TERM("bin_of", [t]); fun bit_tr' b [t] = let val rev_digs = b :: bin_of t handle TERM _ => raise Match val i = int_of rev_digs; val num = string_of_int (abs i); in Syntax.const "_NumeralType" $ Syntax.free num end | bit_tr' b _ = raise Match; in [("bit0", bit_tr' 0), ("bit1", bit_tr' 1)] end; *} subsection {* Examples *} lemma "CARD(0) = 0" by simp lemma "CARD(17) = 17" by simp lemma "8 * 11 ^ 3 - 6 = (2::5)" by simp end diff --git a/src/HOL/Library/Product_Vector.thy b/src/HOL/Library/Product_Vector.thy --- a/src/HOL/Library/Product_Vector.thy +++ b/src/HOL/Library/Product_Vector.thy @@ -1,273 +1,273 @@ (* Title: HOL/Library/Product_Vector.thy Author: Brian Huffman *) header {* Cartesian Products as Vector Spaces *} theory Product_Vector imports Inner_Product Product_plus begin subsection {* Product is a real vector space *} instantiation "*" :: (real_vector, real_vector) real_vector begin definition scaleR_prod_def: "scaleR r A = (scaleR r (fst A), scaleR r (snd A))" lemma fst_scaleR [simp]: "fst (scaleR r A) = scaleR r (fst A)" unfolding scaleR_prod_def by simp lemma snd_scaleR [simp]: "snd (scaleR r A) = scaleR r (snd A)" unfolding scaleR_prod_def by simp lemma scaleR_Pair [simp]: "scaleR r (a, b) = (scaleR r a, scaleR r b)" unfolding scaleR_prod_def by simp instance proof fix a b :: real and x y :: "'a \ 'b" show "scaleR a (x + y) = scaleR a x + scaleR a y" by (simp add: expand_prod_eq scaleR_right_distrib) show "scaleR (a + b) x = scaleR a x + scaleR b x" by (simp add: expand_prod_eq scaleR_left_distrib) show "scaleR a (scaleR b x) = scaleR (a * b) x" by (simp add: expand_prod_eq) show "scaleR 1 x = x" by (simp add: expand_prod_eq) qed end subsection {* Product is a normed vector space *} instantiation "*" :: (real_normed_vector, real_normed_vector) real_normed_vector begin definition norm_prod_def: "norm x = sqrt ((norm (fst x))\ + (norm (snd x))\)" definition sgn_prod_def: "sgn (x::'a \ 'b) = scaleR (inverse (norm x)) x" lemma norm_Pair: "norm (a, b) = sqrt ((norm a)\ + (norm b)\)" unfolding norm_prod_def by simp instance proof fix r :: real and x y :: "'a \ 'b" show "0 \ norm x" unfolding norm_prod_def by simp show "norm x = 0 \ x = 0" unfolding norm_prod_def by (simp add: expand_prod_eq) show "norm (x + y) \ norm x + norm y" unfolding norm_prod_def apply (rule order_trans [OF _ real_sqrt_sum_squares_triangle_ineq]) apply (simp add: add_mono power_mono norm_triangle_ineq) done show "norm (scaleR r x) = \r\ * norm x" unfolding norm_prod_def apply (simp add: norm_scaleR power_mult_distrib) apply (simp add: right_distrib [symmetric]) apply (simp add: real_sqrt_mult_distrib) done show "sgn x = scaleR (inverse (norm x)) x" by (rule sgn_prod_def) qed end subsection {* Product is an inner product space *} instantiation "*" :: (real_inner, real_inner) real_inner begin definition inner_prod_def: "inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)" lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d" unfolding inner_prod_def by simp instance proof fix r :: real fix x y z :: "'a::real_inner * 'b::real_inner" show "inner x y = inner y x" unfolding inner_prod_def by (simp add: inner_commute) show "inner (x + y) z = inner x z + inner y z" unfolding inner_prod_def by (simp add: inner_left_distrib) show "inner (scaleR r x) y = r * inner x y" unfolding inner_prod_def by (simp add: inner_scaleR_left right_distrib) show "0 \ inner x x" unfolding inner_prod_def by (intro add_nonneg_nonneg inner_ge_zero) show "inner x x = 0 \ x = 0" unfolding inner_prod_def expand_prod_eq by (simp add: add_nonneg_eq_0_iff) show "norm x = sqrt (inner x x)" unfolding norm_prod_def inner_prod_def by (simp add: power2_norm_eq_inner) qed end subsection {* Pair operations are linear and continuous *} -interpretation fst!: bounded_linear fst +interpretation fst: bounded_linear fst apply (unfold_locales) apply (rule fst_add) apply (rule fst_scaleR) apply (rule_tac x="1" in exI, simp add: norm_Pair) done -interpretation snd!: bounded_linear snd +interpretation snd: bounded_linear snd apply (unfold_locales) apply (rule snd_add) apply (rule snd_scaleR) apply (rule_tac x="1" in exI, simp add: norm_Pair) done text {* TODO: move to NthRoot *} lemma sqrt_add_le_add_sqrt: assumes x: "0 \ x" and y: "0 \ y" shows "sqrt (x + y) \ sqrt x + sqrt y" apply (rule power2_le_imp_le) apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y) apply (simp add: mult_nonneg_nonneg x y) apply (simp add: add_nonneg_nonneg x y) done lemma bounded_linear_Pair: assumes f: "bounded_linear f" assumes g: "bounded_linear g" shows "bounded_linear (\x. (f x, g x))" proof interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact fix x y and r :: real show "(f (x + y), g (x + y)) = (f x, g x) + (f y, g y)" by (simp add: f.add g.add) show "(f (r *\<^sub>R x), g (r *\<^sub>R x)) = r *\<^sub>R (f x, g x)" by (simp add: f.scaleR g.scaleR) obtain Kf where "0 < Kf" and norm_f: "\x. norm (f x) \ norm x * Kf" using f.pos_bounded by fast obtain Kg where "0 < Kg" and norm_g: "\x. norm (g x) \ norm x * Kg" using g.pos_bounded by fast have "\x. norm (f x, g x) \ norm x * (Kf + Kg)" apply (rule allI) apply (simp add: norm_Pair) apply (rule order_trans [OF sqrt_add_le_add_sqrt], simp, simp) apply (simp add: right_distrib) apply (rule add_mono [OF norm_f norm_g]) done then show "\K. \x. norm (f x, g x) \ norm x * K" .. qed text {* TODO: The next three proofs are nearly identical to each other. Is there a good way to factor out the common parts? *} lemma LIMSEQ_Pair: assumes "X ----> a" and "Y ----> b" shows "(\n. (X n, Y n)) ----> (a, b)" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" then have "0 < r / sqrt 2" (is "0 < ?s") by (simp add: divide_pos_pos) obtain M where M: "\n\M. norm (X n - a) < ?s" using LIMSEQ_D [OF `X ----> a` `0 < ?s`] .. obtain N where N: "\n\N. norm (Y n - b) < ?s" using LIMSEQ_D [OF `Y ----> b` `0 < ?s`] .. have "\n\max M N. norm ((X n, Y n) - (a, b)) < r" using M N by (simp add: real_sqrt_sum_squares_less norm_Pair) then show "\n0. \n\n0. norm ((X n, Y n) - (a, b)) < r" .. qed lemma Cauchy_Pair: assumes "Cauchy X" and "Cauchy Y" shows "Cauchy (\n. (X n, Y n))" proof (rule CauchyI) fix r :: real assume "0 < r" then have "0 < r / sqrt 2" (is "0 < ?s") by (simp add: divide_pos_pos) obtain M where M: "\m\M. \n\M. norm (X m - X n) < ?s" using CauchyD [OF `Cauchy X` `0 < ?s`] .. obtain N where N: "\m\N. \n\N. norm (Y m - Y n) < ?s" using CauchyD [OF `Cauchy Y` `0 < ?s`] .. have "\m\max M N. \n\max M N. norm ((X m, Y m) - (X n, Y n)) < r" using M N by (simp add: real_sqrt_sum_squares_less norm_Pair) then show "\n0. \m\n0. \n\n0. norm ((X m, Y m) - (X n, Y n)) < r" .. qed lemma LIM_Pair: assumes "f -- x --> a" and "g -- x --> b" shows "(\x. (f x, g x)) -- x --> (a, b)" proof (rule LIM_I) fix r :: real assume "0 < r" then have "0 < r / sqrt 2" (is "0 < ?e") by (simp add: divide_pos_pos) obtain s where s: "0 < s" "\z. z \ x \ norm (z - x) < s \ norm (f z - a) < ?e" using LIM_D [OF `f -- x --> a` `0 < ?e`] by fast obtain t where t: "0 < t" "\z. z \ x \ norm (z - x) < t \ norm (g z - b) < ?e" using LIM_D [OF `g -- x --> b` `0 < ?e`] by fast have "0 < min s t \ (\z. z \ x \ norm (z - x) < min s t \ norm ((f z, g z) - (a, b)) < r)" using s t by (simp add: real_sqrt_sum_squares_less norm_Pair) then show "\s>0. \z. z \ x \ norm (z - x) < s \ norm ((f z, g z) - (a, b)) < r" .. qed lemma isCont_Pair [simp]: "\isCont f x; isCont g x\ \ isCont (\x. (f x, g x)) x" unfolding isCont_def by (rule LIM_Pair) subsection {* Product is a complete vector space *} instance "*" :: (banach, banach) banach proof fix X :: "nat \ 'a \ 'b" assume "Cauchy X" have 1: "(\n. fst (X n)) ----> lim (\n. fst (X n))" using fst.Cauchy [OF `Cauchy X`] by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) have 2: "(\n. snd (X n)) ----> lim (\n. snd (X n))" using snd.Cauchy [OF `Cauchy X`] by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) have "X ----> (lim (\n. fst (X n)), lim (\n. snd (X n)))" using LIMSEQ_Pair [OF 1 2] by simp then show "convergent X" by (rule convergentI) qed subsection {* Frechet derivatives involving pairs *} lemma FDERIV_Pair: assumes f: "FDERIV f x :> f'" and g: "FDERIV g x :> g'" shows "FDERIV (\x. (f x, g x)) x :> (\h. (f' h, g' h))" apply (rule FDERIV_I) apply (rule bounded_linear_Pair) apply (rule FDERIV_bounded_linear [OF f]) apply (rule FDERIV_bounded_linear [OF g]) apply (simp add: norm_Pair) apply (rule real_LIM_sandwich_zero) apply (rule LIM_add_zero) apply (rule FDERIV_D [OF f]) apply (rule FDERIV_D [OF g]) apply (rename_tac h) apply (simp add: divide_nonneg_pos) apply (rename_tac h) apply (subst add_divide_distrib [symmetric]) apply (rule divide_right_mono [OF _ norm_ge_zero]) apply (rule order_trans [OF sqrt_add_le_add_sqrt]) apply simp apply simp apply simp done end diff --git a/src/HOL/Library/SetsAndFunctions.thy b/src/HOL/Library/SetsAndFunctions.thy --- a/src/HOL/Library/SetsAndFunctions.thy +++ b/src/HOL/Library/SetsAndFunctions.thy @@ -1,373 +1,373 @@ (* Title: HOL/Library/SetsAndFunctions.thy Author: Jeremy Avigad and Kevin Donnelly *) header {* Operations on sets and functions *} theory SetsAndFunctions imports Plain begin text {* This library lifts operations like addition and muliplication to sets and functions of appropriate types. It was designed to support asymptotic calculations. See the comments at the top of theory @{text BigO}. *} subsection {* Basic definitions *} definition set_plus :: "('a::plus) set => 'a set => 'a set" (infixl "\" 65) where "A \ B == {c. EX a:A. EX b:B. c = a + b}" instantiation "fun" :: (type, plus) plus begin definition func_plus: "f + g == (%x. f x + g x)" instance .. end definition set_times :: "('a::times) set => 'a set => 'a set" (infixl "\" 70) where "A \ B == {c. EX a:A. EX b:B. c = a * b}" instantiation "fun" :: (type, times) times begin definition func_times: "f * g == (%x. f x * g x)" instance .. end instantiation "fun" :: (type, zero) zero begin definition func_zero: "0::(('a::type) => ('b::zero)) == %x. 0" instance .. end instantiation "fun" :: (type, one) one begin definition func_one: "1::(('a::type) => ('b::one)) == %x. 1" instance .. end definition elt_set_plus :: "'a::plus => 'a set => 'a set" (infixl "+o" 70) where "a +o B = {c. EX b:B. c = a + b}" definition elt_set_times :: "'a::times => 'a set => 'a set" (infixl "*o" 80) where "a *o B = {c. EX b:B. c = a * b}" abbreviation (input) elt_set_eq :: "'a => 'a set => bool" (infix "=o" 50) where "x =o A == x : A" instance "fun" :: (type,semigroup_add)semigroup_add by default (auto simp add: func_plus add_assoc) instance "fun" :: (type,comm_monoid_add)comm_monoid_add by default (auto simp add: func_zero func_plus add_ac) instance "fun" :: (type,ab_group_add)ab_group_add apply default apply (simp add: fun_Compl_def func_plus func_zero) apply (simp add: fun_Compl_def func_plus fun_diff_def diff_minus) done instance "fun" :: (type,semigroup_mult)semigroup_mult apply default apply (auto simp add: func_times mult_assoc) done instance "fun" :: (type,comm_monoid_mult)comm_monoid_mult apply default apply (auto simp add: func_one func_times mult_ac) done instance "fun" :: (type,comm_ring_1)comm_ring_1 apply default apply (auto simp add: func_plus func_times fun_Compl_def fun_diff_def func_one func_zero algebra_simps) apply (drule fun_cong) apply simp done -interpretation set_semigroup_add!: semigroup_add "op \ :: ('a::semigroup_add) set => 'a set => 'a set" +interpretation set_semigroup_add: semigroup_add "op \ :: ('a::semigroup_add) set => 'a set => 'a set" apply default apply (unfold set_plus_def) apply (force simp add: add_assoc) done -interpretation set_semigroup_mult!: semigroup_mult "op \ :: ('a::semigroup_mult) set => 'a set => 'a set" +interpretation set_semigroup_mult: semigroup_mult "op \ :: ('a::semigroup_mult) set => 'a set => 'a set" apply default apply (unfold set_times_def) apply (force simp add: mult_assoc) done -interpretation set_comm_monoid_add!: comm_monoid_add "{0}" "op \ :: ('a::comm_monoid_add) set => 'a set => 'a set" +interpretation set_comm_monoid_add: comm_monoid_add "{0}" "op \ :: ('a::comm_monoid_add) set => 'a set => 'a set" apply default apply (unfold set_plus_def) apply (force simp add: add_ac) apply force done -interpretation set_comm_monoid_mult!: comm_monoid_mult "{1}" "op \ :: ('a::comm_monoid_mult) set => 'a set => 'a set" +interpretation set_comm_monoid_mult: comm_monoid_mult "{1}" "op \ :: ('a::comm_monoid_mult) set => 'a set => 'a set" apply default apply (unfold set_times_def) apply (force simp add: mult_ac) apply force done subsection {* Basic properties *} lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C \ D" by (auto simp add: set_plus_def) lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C" by (auto simp add: elt_set_plus_def) lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) \ (b +o D) = (a + b) +o (C \ D)" apply (auto simp add: elt_set_plus_def set_plus_def add_ac) apply (rule_tac x = "ba + bb" in exI) apply (auto simp add: add_ac) apply (rule_tac x = "aa + a" in exI) apply (auto simp add: add_ac) done lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = (a + b) +o C" by (auto simp add: elt_set_plus_def add_assoc) lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) \ C = a +o (B \ C)" apply (auto simp add: elt_set_plus_def set_plus_def) apply (blast intro: add_ac) apply (rule_tac x = "a + aa" in exI) apply (rule conjI) apply (rule_tac x = "aa" in bexI) apply auto apply (rule_tac x = "ba" in bexI) apply (auto simp add: add_ac) done theorem set_plus_rearrange4: "C \ ((a::'a::comm_monoid_add) +o D) = a +o (C \ D)" apply (auto intro!: subsetI simp add: elt_set_plus_def set_plus_def add_ac) apply (rule_tac x = "aa + ba" in exI) apply (auto simp add: add_ac) done theorems set_plus_rearranges = set_plus_rearrange set_plus_rearrange2 set_plus_rearrange3 set_plus_rearrange4 lemma set_plus_mono [intro!]: "C <= D ==> a +o C <= a +o D" by (auto simp add: elt_set_plus_def) lemma set_plus_mono2 [intro]: "(C::('a::plus) set) <= D ==> E <= F ==> C \ E <= D \ F" by (auto simp add: set_plus_def) lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C \ D" by (auto simp add: elt_set_plus_def set_plus_def) lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==> a +o D <= D \ C" by (auto simp add: elt_set_plus_def set_plus_def add_ac) lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C \ D" apply (subgoal_tac "a +o B <= a +o D") apply (erule order_trans) apply (erule set_plus_mono3) apply (erule set_plus_mono) done lemma set_plus_mono_b: "C <= D ==> x : a +o C ==> x : a +o D" apply (frule set_plus_mono) apply auto done lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C \ E ==> x : D \ F" apply (frule set_plus_mono2) prefer 2 apply force apply assumption done lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C \ D" apply (frule set_plus_mono3) apply auto done lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==> x : a +o D ==> x : D \ C" apply (frule set_plus_mono4) apply auto done lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C" by (auto simp add: elt_set_plus_def) lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A \ B" apply (auto intro!: subsetI simp add: set_plus_def) apply (rule_tac x = 0 in bexI) apply (rule_tac x = x in bexI) apply (auto simp add: add_ac) done lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C" by (auto simp add: elt_set_plus_def add_ac diff_minus) lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C ==> a : b +o C" apply (auto simp add: elt_set_plus_def add_ac diff_minus) apply (subgoal_tac "a = (a + - b) + b") apply (rule bexI, assumption, assumption) apply (auto simp add: add_ac) done lemma set_minus_plus: "((a::'a::ab_group_add) - b : C) = (a : b +o C)" by (rule iffI, rule set_minus_imp_plus, assumption, rule set_plus_imp_minus, assumption) lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C \ D" by (auto simp add: set_times_def) lemma set_times_intro2 [intro!]: "b : C ==> a * b : a *o C" by (auto simp add: elt_set_times_def) lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) \ (b *o D) = (a * b) *o (C \ D)" apply (auto simp add: elt_set_times_def set_times_def) apply (rule_tac x = "ba * bb" in exI) apply (auto simp add: mult_ac) apply (rule_tac x = "aa * a" in exI) apply (auto simp add: mult_ac) done lemma set_times_rearrange2: "(a::'a::semigroup_mult) *o (b *o C) = (a * b) *o C" by (auto simp add: elt_set_times_def mult_assoc) lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) \ C = a *o (B \ C)" apply (auto simp add: elt_set_times_def set_times_def) apply (blast intro: mult_ac) apply (rule_tac x = "a * aa" in exI) apply (rule conjI) apply (rule_tac x = "aa" in bexI) apply auto apply (rule_tac x = "ba" in bexI) apply (auto simp add: mult_ac) done theorem set_times_rearrange4: "C \ ((a::'a::comm_monoid_mult) *o D) = a *o (C \ D)" apply (auto intro!: subsetI simp add: elt_set_times_def set_times_def mult_ac) apply (rule_tac x = "aa * ba" in exI) apply (auto simp add: mult_ac) done theorems set_times_rearranges = set_times_rearrange set_times_rearrange2 set_times_rearrange3 set_times_rearrange4 lemma set_times_mono [intro]: "C <= D ==> a *o C <= a *o D" by (auto simp add: elt_set_times_def) lemma set_times_mono2 [intro]: "(C::('a::times) set) <= D ==> E <= F ==> C \ E <= D \ F" by (auto simp add: set_times_def) lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C \ D" by (auto simp add: elt_set_times_def set_times_def) lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==> a *o D <= D \ C" by (auto simp add: elt_set_times_def set_times_def mult_ac) lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C \ D" apply (subgoal_tac "a *o B <= a *o D") apply (erule order_trans) apply (erule set_times_mono3) apply (erule set_times_mono) done lemma set_times_mono_b: "C <= D ==> x : a *o C ==> x : a *o D" apply (frule set_times_mono) apply auto done lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C \ E ==> x : D \ F" apply (frule set_times_mono2) prefer 2 apply force apply assumption done lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C \ D" apply (frule set_times_mono3) apply auto done lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==> x : a *o D ==> x : D \ C" apply (frule set_times_mono4) apply auto done lemma set_one_times [simp]: "(1::'a::comm_monoid_mult) *o C = C" by (auto simp add: elt_set_times_def) lemma set_times_plus_distrib: "(a::'a::semiring) *o (b +o C)= (a * b) +o (a *o C)" by (auto simp add: elt_set_plus_def elt_set_times_def ring_distribs) lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B \ C) = (a *o B) \ (a *o C)" apply (auto simp add: set_plus_def elt_set_times_def ring_distribs) apply blast apply (rule_tac x = "b + bb" in exI) apply (auto simp add: ring_distribs) done lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) \ D <= a *o D \ C \ D" apply (auto intro!: subsetI simp add: elt_set_plus_def elt_set_times_def set_times_def set_plus_def ring_distribs) apply auto done theorems set_times_plus_distribs = set_times_plus_distrib set_times_plus_distrib2 lemma set_neg_intro: "(a::'a::ring_1) : (- 1) *o C ==> - a : C" by (auto simp add: elt_set_times_def) lemma set_neg_intro2: "(a::'a::ring_1) : C ==> - a : (- 1) *o C" by (auto simp add: elt_set_times_def) end diff --git a/src/HOL/NSA/StarDef.thy b/src/HOL/NSA/StarDef.thy --- a/src/HOL/NSA/StarDef.thy +++ b/src/HOL/NSA/StarDef.thy @@ -1,1033 +1,1033 @@ (* Title : HOL/Hyperreal/StarDef.thy ID : $Id$ Author : Jacques D. Fleuriot and Brian Huffman *) header {* Construction of Star Types Using Ultrafilters *} theory StarDef imports Filter uses ("transfer.ML") begin subsection {* A Free Ultrafilter over the Naturals *} definition FreeUltrafilterNat :: "nat set set" ("\") where "\ = (SOME U. freeultrafilter U)" lemma freeultrafilter_FreeUltrafilterNat: "freeultrafilter \" apply (unfold FreeUltrafilterNat_def) apply (rule someI_ex [where P=freeultrafilter]) apply (rule freeultrafilter_Ex) apply (rule nat_infinite) done -interpretation FreeUltrafilterNat!: freeultrafilter FreeUltrafilterNat +interpretation FreeUltrafilterNat: freeultrafilter FreeUltrafilterNat by (rule freeultrafilter_FreeUltrafilterNat) text {* This rule takes the place of the old ultra tactic *} lemma ultra: "\{n. P n} \ \; {n. P n \ Q n} \ \\ \ {n. Q n} \ \" by (simp add: Collect_imp_eq FreeUltrafilterNat.Un_iff FreeUltrafilterNat.Compl_iff) subsection {* Definition of @{text star} type constructor *} definition starrel :: "((nat \ 'a) \ (nat \ 'a)) set" where "starrel = {(X,Y). {n. X n = Y n} \ \}" typedef 'a star = "(UNIV :: (nat \ 'a) set) // starrel" by (auto intro: quotientI) definition star_n :: "(nat \ 'a) \ 'a star" where "star_n X = Abs_star (starrel `` {X})" theorem star_cases [case_names star_n, cases type: star]: "(\X. x = star_n X \ P) \ P" by (cases x, unfold star_n_def star_def, erule quotientE, fast) lemma all_star_eq: "(\x. P x) = (\X. P (star_n X))" by (auto, rule_tac x=x in star_cases, simp) lemma ex_star_eq: "(\x. P x) = (\X. P (star_n X))" by (auto, rule_tac x=x in star_cases, auto) text {* Proving that @{term starrel} is an equivalence relation *} lemma starrel_iff [iff]: "((X,Y) \ starrel) = ({n. X n = Y n} \ \)" by (simp add: starrel_def) lemma equiv_starrel: "equiv UNIV starrel" proof (rule equiv.intro) show "refl starrel" by (simp add: refl_on_def) show "sym starrel" by (simp add: sym_def eq_commute) show "trans starrel" by (auto intro: transI elim!: ultra) qed lemmas equiv_starrel_iff = eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I] lemma starrel_in_star: "starrel``{x} \ star" by (simp add: star_def quotientI) lemma star_n_eq_iff: "(star_n X = star_n Y) = ({n. X n = Y n} \ \)" by (simp add: star_n_def Abs_star_inject starrel_in_star equiv_starrel_iff) subsection {* Transfer principle *} text {* This introduction rule starts each transfer proof. *} lemma transfer_start: "P \ {n. Q} \ \ \ Trueprop P \ Trueprop Q" by (subgoal_tac "P \ Q", simp, simp add: atomize_eq) text {*Initialize transfer tactic.*} use "transfer.ML" setup Transfer.setup text {* Transfer introduction rules. *} lemma transfer_ex [transfer_intro]: "\\X. p (star_n X) \ {n. P n (X n)} \ \\ \ \x::'a star. p x \ {n. \x. P n x} \ \" by (simp only: ex_star_eq FreeUltrafilterNat.Collect_ex) lemma transfer_all [transfer_intro]: "\\X. p (star_n X) \ {n. P n (X n)} \ \\ \ \x::'a star. p x \ {n. \x. P n x} \ \" by (simp only: all_star_eq FreeUltrafilterNat.Collect_all) lemma transfer_not [transfer_intro]: "\p \ {n. P n} \ \\ \ \ p \ {n. \ P n} \ \" by (simp only: FreeUltrafilterNat.Collect_not) lemma transfer_conj [transfer_intro]: "\p \ {n. P n} \ \; q \ {n. Q n} \ \\ \ p \ q \ {n. P n \ Q n} \ \" by (simp only: FreeUltrafilterNat.Collect_conj) lemma transfer_disj [transfer_intro]: "\p \ {n. P n} \ \; q \ {n. Q n} \ \\ \ p \ q \ {n. P n \ Q n} \ \" by (simp only: FreeUltrafilterNat.Collect_disj) lemma transfer_imp [transfer_intro]: "\p \ {n. P n} \ \; q \ {n. Q n} \ \\ \ p \ q \ {n. P n \ Q n} \ \" by (simp only: imp_conv_disj transfer_disj transfer_not) lemma transfer_iff [transfer_intro]: "\p \ {n. P n} \ \; q \ {n. Q n} \ \\ \ p = q \ {n. P n = Q n} \ \" by (simp only: iff_conv_conj_imp transfer_conj transfer_imp) lemma transfer_if_bool [transfer_intro]: "\p \ {n. P n} \ \; x \ {n. X n} \ \; y \ {n. Y n} \ \\ \ (if p then x else y) \ {n. if P n then X n else Y n} \ \" by (simp only: if_bool_eq_conj transfer_conj transfer_imp transfer_not) lemma transfer_eq [transfer_intro]: "\x \ star_n X; y \ star_n Y\ \ x = y \ {n. X n = Y n} \ \" by (simp only: star_n_eq_iff) lemma transfer_if [transfer_intro]: "\p \ {n. P n} \ \; x \ star_n X; y \ star_n Y\ \ (if p then x else y) \ star_n (\n. if P n then X n else Y n)" apply (rule eq_reflection) apply (auto simp add: star_n_eq_iff transfer_not elim!: ultra) done lemma transfer_fun_eq [transfer_intro]: "\\X. f (star_n X) = g (star_n X) \ {n. F n (X n) = G n (X n)} \ \\ \ f = g \ {n. F n = G n} \ \" by (simp only: expand_fun_eq transfer_all) lemma transfer_star_n [transfer_intro]: "star_n X \ star_n (\n. X n)" by (rule reflexive) lemma transfer_bool [transfer_intro]: "p \ {n. p} \ \" by (simp add: atomize_eq) subsection {* Standard elements *} definition star_of :: "'a \ 'a star" where "star_of x == star_n (\n. x)" definition Standard :: "'a star set" where "Standard = range star_of" text {* Transfer tactic should remove occurrences of @{term star_of} *} setup {* Transfer.add_const "StarDef.star_of" *} declare star_of_def [transfer_intro] lemma star_of_inject: "(star_of x = star_of y) = (x = y)" by (transfer, rule refl) lemma Standard_star_of [simp]: "star_of x \ Standard" by (simp add: Standard_def) subsection {* Internal functions *} definition Ifun :: "('a \ 'b) star \ 'a star \ 'b star" ("_ \ _" [300,301] 300) where "Ifun f \ \x. Abs_star (\F\Rep_star f. \X\Rep_star x. starrel``{\n. F n (X n)})" lemma Ifun_congruent2: "congruent2 starrel starrel (\F X. starrel``{\n. F n (X n)})" by (auto simp add: congruent2_def equiv_starrel_iff elim!: ultra) lemma Ifun_star_n: "star_n F \ star_n X = star_n (\n. F n (X n))" by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2]) text {* Transfer tactic should remove occurrences of @{term Ifun} *} setup {* Transfer.add_const "StarDef.Ifun" *} lemma transfer_Ifun [transfer_intro]: "\f \ star_n F; x \ star_n X\ \ f \ x \ star_n (\n. F n (X n))" by (simp only: Ifun_star_n) lemma Ifun_star_of [simp]: "star_of f \ star_of x = star_of (f x)" by (transfer, rule refl) lemma Standard_Ifun [simp]: "\f \ Standard; x \ Standard\ \ f \ x \ Standard" by (auto simp add: Standard_def) text {* Nonstandard extensions of functions *} definition starfun :: "('a \ 'b) \ ('a star \ 'b star)" ("*f* _" [80] 80) where "starfun f == \x. star_of f \ x" definition starfun2 :: "('a \ 'b \ 'c) \ ('a star \ 'b star \ 'c star)" ("*f2* _" [80] 80) where "starfun2 f == \x y. star_of f \ x \ y" declare starfun_def [transfer_unfold] declare starfun2_def [transfer_unfold] lemma starfun_star_n: "( *f* f) (star_n X) = star_n (\n. f (X n))" by (simp only: starfun_def star_of_def Ifun_star_n) lemma starfun2_star_n: "( *f2* f) (star_n X) (star_n Y) = star_n (\n. f (X n) (Y n))" by (simp only: starfun2_def star_of_def Ifun_star_n) lemma starfun_star_of [simp]: "( *f* f) (star_of x) = star_of (f x)" by (transfer, rule refl) lemma starfun2_star_of [simp]: "( *f2* f) (star_of x) = *f* f x" by (transfer, rule refl) lemma Standard_starfun [simp]: "x \ Standard \ starfun f x \ Standard" by (simp add: starfun_def) lemma Standard_starfun2 [simp]: "\x \ Standard; y \ Standard\ \ starfun2 f x y \ Standard" by (simp add: starfun2_def) lemma Standard_starfun_iff: assumes inj: "\x y. f x = f y \ x = y" shows "(starfun f x \ Standard) = (x \ Standard)" proof assume "x \ Standard" thus "starfun f x \ Standard" by simp next have inj': "\x y. starfun f x = starfun f y \ x = y" using inj by transfer assume "starfun f x \ Standard" then obtain b where b: "starfun f x = star_of b" unfolding Standard_def .. hence "\x. starfun f x = star_of b" .. hence "\a. f a = b" by transfer then obtain a where "f a = b" .. hence "starfun f (star_of a) = star_of b" by transfer with b have "starfun f x = starfun f (star_of a)" by simp hence "x = star_of a" by (rule inj') thus "x \ Standard" unfolding Standard_def by auto qed lemma Standard_starfun2_iff: assumes inj: "\a b a' b'. f a b = f a' b' \ a = a' \ b = b'" shows "(starfun2 f x y \ Standard) = (x \ Standard \ y \ Standard)" proof assume "x \ Standard \ y \ Standard" thus "starfun2 f x y \ Standard" by simp next have inj': "\x y z w. starfun2 f x y = starfun2 f z w \ x = z \ y = w" using inj by transfer assume "starfun2 f x y \ Standard" then obtain c where c: "starfun2 f x y = star_of c" unfolding Standard_def .. hence "\x y. starfun2 f x y = star_of c" by auto hence "\a b. f a b = c" by transfer then obtain a b where "f a b = c" by auto hence "starfun2 f (star_of a) (star_of b) = star_of c" by transfer with c have "starfun2 f x y = starfun2 f (star_of a) (star_of b)" by simp hence "x = star_of a \ y = star_of b" by (rule inj') thus "x \ Standard \ y \ Standard" unfolding Standard_def by auto qed subsection {* Internal predicates *} definition unstar :: "bool star \ bool" where [code del]: "unstar b \ b = star_of True" lemma unstar_star_n: "unstar (star_n P) = ({n. P n} \ \)" by (simp add: unstar_def star_of_def star_n_eq_iff) lemma unstar_star_of [simp]: "unstar (star_of p) = p" by (simp add: unstar_def star_of_inject) text {* Transfer tactic should remove occurrences of @{term unstar} *} setup {* Transfer.add_const "StarDef.unstar" *} lemma transfer_unstar [transfer_intro]: "p \ star_n P \ unstar p \ {n. P n} \ \" by (simp only: unstar_star_n) definition starP :: "('a \ bool) \ 'a star \ bool" ("*p* _" [80] 80) where "*p* P = (\x. unstar (star_of P \ x))" definition starP2 :: "('a \ 'b \ bool) \ 'a star \ 'b star \ bool" ("*p2* _" [80] 80) where "*p2* P = (\x y. unstar (star_of P \ x \ y))" declare starP_def [transfer_unfold] declare starP2_def [transfer_unfold] lemma starP_star_n: "( *p* P) (star_n X) = ({n. P (X n)} \ \)" by (simp only: starP_def star_of_def Ifun_star_n unstar_star_n) lemma starP2_star_n: "( *p2* P) (star_n X) (star_n Y) = ({n. P (X n) (Y n)} \ \)" by (simp only: starP2_def star_of_def Ifun_star_n unstar_star_n) lemma starP_star_of [simp]: "( *p* P) (star_of x) = P x" by (transfer, rule refl) lemma starP2_star_of [simp]: "( *p2* P) (star_of x) = *p* P x" by (transfer, rule refl) subsection {* Internal sets *} definition Iset :: "'a set star \ 'a star set" where "Iset A = {x. ( *p2* op \) x A}" lemma Iset_star_n: "(star_n X \ Iset (star_n A)) = ({n. X n \ A n} \ \)" by (simp add: Iset_def starP2_star_n) text {* Transfer tactic should remove occurrences of @{term Iset} *} setup {* Transfer.add_const "StarDef.Iset" *} lemma transfer_mem [transfer_intro]: "\x \ star_n X; a \ Iset (star_n A)\ \ x \ a \ {n. X n \ A n} \ \" by (simp only: Iset_star_n) lemma transfer_Collect [transfer_intro]: "\\X. p (star_n X) \ {n. P n (X n)} \ \\ \ Collect p \ Iset (star_n (\n. Collect (P n)))" by (simp add: atomize_eq expand_set_eq all_star_eq Iset_star_n) lemma transfer_set_eq [transfer_intro]: "\a \ Iset (star_n A); b \ Iset (star_n B)\ \ a = b \ {n. A n = B n} \ \" by (simp only: expand_set_eq transfer_all transfer_iff transfer_mem) lemma transfer_ball [transfer_intro]: "\a \ Iset (star_n A); \X. p (star_n X) \ {n. P n (X n)} \ \\ \ \x\a. p x \ {n. \x\A n. P n x} \ \" by (simp only: Ball_def transfer_all transfer_imp transfer_mem) lemma transfer_bex [transfer_intro]: "\a \ Iset (star_n A); \X. p (star_n X) \ {n. P n (X n)} \ \\ \ \x\a. p x \ {n. \x\A n. P n x} \ \" by (simp only: Bex_def transfer_ex transfer_conj transfer_mem) lemma transfer_Iset [transfer_intro]: "\a \ star_n A\ \ Iset a \ Iset (star_n (\n. A n))" by simp text {* Nonstandard extensions of sets. *} definition starset :: "'a set \ 'a star set" ("*s* _" [80] 80) where "starset A = Iset (star_of A)" declare starset_def [transfer_unfold] lemma starset_mem: "(star_of x \ *s* A) = (x \ A)" by (transfer, rule refl) lemma starset_UNIV: "*s* (UNIV::'a set) = (UNIV::'a star set)" by (transfer UNIV_def, rule refl) lemma starset_empty: "*s* {} = {}" by (transfer empty_def, rule refl) lemma starset_insert: "*s* (insert x A) = insert (star_of x) ( *s* A)" by (transfer insert_def Un_def, rule refl) lemma starset_Un: "*s* (A \ B) = *s* A \ *s* B" by (transfer Un_def, rule refl) lemma starset_Int: "*s* (A \ B) = *s* A \ *s* B" by (transfer Int_def, rule refl) lemma starset_Compl: "*s* -A = -( *s* A)" by (transfer Compl_eq, rule refl) lemma starset_diff: "*s* (A - B) = *s* A - *s* B" by (transfer set_diff_eq, rule refl) lemma starset_image: "*s* (f ` A) = ( *f* f) ` ( *s* A)" by (transfer image_def, rule refl) lemma starset_vimage: "*s* (f -` A) = ( *f* f) -` ( *s* A)" by (transfer vimage_def, rule refl) lemma starset_subset: "( *s* A \ *s* B) = (A \ B)" by (transfer subset_eq, rule refl) lemma starset_eq: "( *s* A = *s* B) = (A = B)" by (transfer, rule refl) lemmas starset_simps [simp] = starset_mem starset_UNIV starset_empty starset_insert starset_Un starset_Int starset_Compl starset_diff starset_image starset_vimage starset_subset starset_eq subsection {* Syntactic classes *} instantiation star :: (zero) zero begin definition star_zero_def [code del]: "0 \ star_of 0" instance .. end instantiation star :: (one) one begin definition star_one_def [code del]: "1 \ star_of 1" instance .. end instantiation star :: (plus) plus begin definition star_add_def [code del]: "(op +) \ *f2* (op +)" instance .. end instantiation star :: (times) times begin definition star_mult_def [code del]: "(op *) \ *f2* (op *)" instance .. end instantiation star :: (uminus) uminus begin definition star_minus_def [code del]: "uminus \ *f* uminus" instance .. end instantiation star :: (minus) minus begin definition star_diff_def [code del]: "(op -) \ *f2* (op -)" instance .. end instantiation star :: (abs) abs begin definition star_abs_def: "abs \ *f* abs" instance .. end instantiation star :: (sgn) sgn begin definition star_sgn_def: "sgn \ *f* sgn" instance .. end instantiation star :: (inverse) inverse begin definition star_divide_def: "(op /) \ *f2* (op /)" definition star_inverse_def: "inverse \ *f* inverse" instance .. end instantiation star :: (number) number begin definition star_number_def: "number_of b \ star_of (number_of b)" instance .. end instance star :: (Ring_and_Field.dvd) Ring_and_Field.dvd .. instantiation star :: (Divides.div) Divides.div begin definition star_div_def: "(op div) \ *f2* (op div)" definition star_mod_def: "(op mod) \ *f2* (op mod)" instance .. end instantiation star :: (power) power begin definition star_power_def: "(op ^) \ \x n. ( *f* (\x. x ^ n)) x" instance .. end instantiation star :: (ord) ord begin definition star_le_def: "(op \) \ *p2* (op \)" definition star_less_def: "(op <) \ *p2* (op <)" instance .. end lemmas star_class_defs [transfer_unfold] = star_zero_def star_one_def star_number_def star_add_def star_diff_def star_minus_def star_mult_def star_divide_def star_inverse_def star_le_def star_less_def star_abs_def star_sgn_def star_div_def star_mod_def star_power_def text {* Class operations preserve standard elements *} lemma Standard_zero: "0 \ Standard" by (simp add: star_zero_def) lemma Standard_one: "1 \ Standard" by (simp add: star_one_def) lemma Standard_number_of: "number_of b \ Standard" by (simp add: star_number_def) lemma Standard_add: "\x \ Standard; y \ Standard\ \ x + y \ Standard" by (simp add: star_add_def) lemma Standard_diff: "\x \ Standard; y \ Standard\ \ x - y \ Standard" by (simp add: star_diff_def) lemma Standard_minus: "x \ Standard \ - x \ Standard" by (simp add: star_minus_def) lemma Standard_mult: "\x \ Standard; y \ Standard\ \ x * y \ Standard" by (simp add: star_mult_def) lemma Standard_divide: "\x \ Standard; y \ Standard\ \ x / y \ Standard" by (simp add: star_divide_def) lemma Standard_inverse: "x \ Standard \ inverse x \ Standard" by (simp add: star_inverse_def) lemma Standard_abs: "x \ Standard \ abs x \ Standard" by (simp add: star_abs_def) lemma Standard_div: "\x \ Standard; y \ Standard\ \ x div y \ Standard" by (simp add: star_div_def) lemma Standard_mod: "\x \ Standard; y \ Standard\ \ x mod y \ Standard" by (simp add: star_mod_def) lemma Standard_power: "x \ Standard \ x ^ n \ Standard" by (simp add: star_power_def) lemmas Standard_simps [simp] = Standard_zero Standard_one Standard_number_of Standard_add Standard_diff Standard_minus Standard_mult Standard_divide Standard_inverse Standard_abs Standard_div Standard_mod Standard_power text {* @{term star_of} preserves class operations *} lemma star_of_add: "star_of (x + y) = star_of x + star_of y" by transfer (rule refl) lemma star_of_diff: "star_of (x - y) = star_of x - star_of y" by transfer (rule refl) lemma star_of_minus: "star_of (-x) = - star_of x" by transfer (rule refl) lemma star_of_mult: "star_of (x * y) = star_of x * star_of y" by transfer (rule refl) lemma star_of_divide: "star_of (x / y) = star_of x / star_of y" by transfer (rule refl) lemma star_of_inverse: "star_of (inverse x) = inverse (star_of x)" by transfer (rule refl) lemma star_of_div: "star_of (x div y) = star_of x div star_of y" by transfer (rule refl) lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y" by transfer (rule refl) lemma star_of_power: "star_of (x ^ n) = star_of x ^ n" by transfer (rule refl) lemma star_of_abs: "star_of (abs x) = abs (star_of x)" by transfer (rule refl) text {* @{term star_of} preserves numerals *} lemma star_of_zero: "star_of 0 = 0" by transfer (rule refl) lemma star_of_one: "star_of 1 = 1" by transfer (rule refl) lemma star_of_number_of: "star_of (number_of x) = number_of x" by transfer (rule refl) text {* @{term star_of} preserves orderings *} lemma star_of_less: "(star_of x < star_of y) = (x < y)" by transfer (rule refl) lemma star_of_le: "(star_of x \ star_of y) = (x \ y)" by transfer (rule refl) lemma star_of_eq: "(star_of x = star_of y) = (x = y)" by transfer (rule refl) text{*As above, for 0*} lemmas star_of_0_less = star_of_less [of 0, simplified star_of_zero] lemmas star_of_0_le = star_of_le [of 0, simplified star_of_zero] lemmas star_of_0_eq = star_of_eq [of 0, simplified star_of_zero] lemmas star_of_less_0 = star_of_less [of _ 0, simplified star_of_zero] lemmas star_of_le_0 = star_of_le [of _ 0, simplified star_of_zero] lemmas star_of_eq_0 = star_of_eq [of _ 0, simplified star_of_zero] text{*As above, for 1*} lemmas star_of_1_less = star_of_less [of 1, simplified star_of_one] lemmas star_of_1_le = star_of_le [of 1, simplified star_of_one] lemmas star_of_1_eq = star_of_eq [of 1, simplified star_of_one] lemmas star_of_less_1 = star_of_less [of _ 1, simplified star_of_one] lemmas star_of_le_1 = star_of_le [of _ 1, simplified star_of_one] lemmas star_of_eq_1 = star_of_eq [of _ 1, simplified star_of_one] text{*As above, for numerals*} lemmas star_of_number_less = star_of_less [of "number_of w", standard, simplified star_of_number_of] lemmas star_of_number_le = star_of_le [of "number_of w", standard, simplified star_of_number_of] lemmas star_of_number_eq = star_of_eq [of "number_of w", standard, simplified star_of_number_of] lemmas star_of_less_number = star_of_less [of _ "number_of w", standard, simplified star_of_number_of] lemmas star_of_le_number = star_of_le [of _ "number_of w", standard, simplified star_of_number_of] lemmas star_of_eq_number = star_of_eq [of _ "number_of w", standard, simplified star_of_number_of] lemmas star_of_simps [simp] = star_of_add star_of_diff star_of_minus star_of_mult star_of_divide star_of_inverse star_of_div star_of_mod star_of_power star_of_abs star_of_zero star_of_one star_of_number_of star_of_less star_of_le star_of_eq star_of_0_less star_of_0_le star_of_0_eq star_of_less_0 star_of_le_0 star_of_eq_0 star_of_1_less star_of_1_le star_of_1_eq star_of_less_1 star_of_le_1 star_of_eq_1 star_of_number_less star_of_number_le star_of_number_eq star_of_less_number star_of_le_number star_of_eq_number subsection {* Ordering and lattice classes *} instance star :: (order) order apply (intro_classes) apply (transfer, rule less_le_not_le) apply (transfer, rule order_refl) apply (transfer, erule (1) order_trans) apply (transfer, erule (1) order_antisym) done instantiation star :: (lower_semilattice) lower_semilattice begin definition star_inf_def [transfer_unfold]: "inf \ *f2* inf" instance by default (transfer star_inf_def, auto)+ end instantiation star :: (upper_semilattice) upper_semilattice begin definition star_sup_def [transfer_unfold]: "sup \ *f2* sup" instance by default (transfer star_sup_def, auto)+ end instance star :: (lattice) lattice .. instance star :: (distrib_lattice) distrib_lattice by default (transfer, auto simp add: sup_inf_distrib1) lemma Standard_inf [simp]: "\x \ Standard; y \ Standard\ \ inf x y \ Standard" by (simp add: star_inf_def) lemma Standard_sup [simp]: "\x \ Standard; y \ Standard\ \ sup x y \ Standard" by (simp add: star_sup_def) lemma star_of_inf [simp]: "star_of (inf x y) = inf (star_of x) (star_of y)" by transfer (rule refl) lemma star_of_sup [simp]: "star_of (sup x y) = sup (star_of x) (star_of y)" by transfer (rule refl) instance star :: (linorder) linorder by (intro_classes, transfer, rule linorder_linear) lemma star_max_def [transfer_unfold]: "max = *f2* max" apply (rule ext, rule ext) apply (unfold max_def, transfer, fold max_def) apply (rule refl) done lemma star_min_def [transfer_unfold]: "min = *f2* min" apply (rule ext, rule ext) apply (unfold min_def, transfer, fold min_def) apply (rule refl) done lemma Standard_max [simp]: "\x \ Standard; y \ Standard\ \ max x y \ Standard" by (simp add: star_max_def) lemma Standard_min [simp]: "\x \ Standard; y \ Standard\ \ min x y \ Standard" by (simp add: star_min_def) lemma star_of_max [simp]: "star_of (max x y) = max (star_of x) (star_of y)" by transfer (rule refl) lemma star_of_min [simp]: "star_of (min x y) = min (star_of x) (star_of y)" by transfer (rule refl) subsection {* Ordered group classes *} instance star :: (semigroup_add) semigroup_add by (intro_classes, transfer, rule add_assoc) instance star :: (ab_semigroup_add) ab_semigroup_add by (intro_classes, transfer, rule add_commute) instance star :: (semigroup_mult) semigroup_mult by (intro_classes, transfer, rule mult_assoc) instance star :: (ab_semigroup_mult) ab_semigroup_mult by (intro_classes, transfer, rule mult_commute) instance star :: (comm_monoid_add) comm_monoid_add by (intro_classes, transfer, rule comm_monoid_add_class.add_0) instance star :: (monoid_mult) monoid_mult apply (intro_classes) apply (transfer, rule mult_1_left) apply (transfer, rule mult_1_right) done instance star :: (comm_monoid_mult) comm_monoid_mult by (intro_classes, transfer, rule mult_1) instance star :: (cancel_semigroup_add) cancel_semigroup_add apply (intro_classes) apply (transfer, erule add_left_imp_eq) apply (transfer, erule add_right_imp_eq) done instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add by (intro_classes, transfer, rule add_imp_eq) instance star :: (cancel_comm_monoid_add) cancel_comm_monoid_add .. instance star :: (ab_group_add) ab_group_add apply (intro_classes) apply (transfer, rule left_minus) apply (transfer, rule diff_minus) done instance star :: (pordered_ab_semigroup_add) pordered_ab_semigroup_add by (intro_classes, transfer, rule add_left_mono) instance star :: (pordered_cancel_ab_semigroup_add) pordered_cancel_ab_semigroup_add .. instance star :: (pordered_ab_semigroup_add_imp_le) pordered_ab_semigroup_add_imp_le by (intro_classes, transfer, rule add_le_imp_le_left) instance star :: (pordered_comm_monoid_add) pordered_comm_monoid_add .. instance star :: (pordered_ab_group_add) pordered_ab_group_add .. instance star :: (pordered_ab_group_add_abs) pordered_ab_group_add_abs by intro_classes (transfer, simp add: abs_ge_self abs_leI abs_triangle_ineq)+ instance star :: (ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add .. instance star :: (lordered_ab_group_add_meet) lordered_ab_group_add_meet .. instance star :: (lordered_ab_group_add_meet) lordered_ab_group_add_meet .. instance star :: (lordered_ab_group_add) lordered_ab_group_add .. instance star :: (lordered_ab_group_add_abs) lordered_ab_group_add_abs by (intro_classes, transfer, rule abs_lattice) subsection {* Ring and field classes *} instance star :: (semiring) semiring apply (intro_classes) apply (transfer, rule left_distrib) apply (transfer, rule right_distrib) done instance star :: (semiring_0) semiring_0 by intro_classes (transfer, simp)+ instance star :: (semiring_0_cancel) semiring_0_cancel .. instance star :: (comm_semiring) comm_semiring by (intro_classes, transfer, rule left_distrib) instance star :: (comm_semiring_0) comm_semiring_0 .. instance star :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. instance star :: (zero_neq_one) zero_neq_one by (intro_classes, transfer, rule zero_neq_one) instance star :: (semiring_1) semiring_1 .. instance star :: (comm_semiring_1) comm_semiring_1 .. instance star :: (no_zero_divisors) no_zero_divisors by (intro_classes, transfer, rule no_zero_divisors) instance star :: (semiring_1_cancel) semiring_1_cancel .. instance star :: (comm_semiring_1_cancel) comm_semiring_1_cancel .. instance star :: (ring) ring .. instance star :: (comm_ring) comm_ring .. instance star :: (ring_1) ring_1 .. instance star :: (comm_ring_1) comm_ring_1 .. instance star :: (ring_no_zero_divisors) ring_no_zero_divisors .. instance star :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors .. instance star :: (idom) idom .. instance star :: (division_ring) division_ring apply (intro_classes) apply (transfer, erule left_inverse) apply (transfer, erule right_inverse) done instance star :: (field) field apply (intro_classes) apply (transfer, erule left_inverse) apply (transfer, rule divide_inverse) done instance star :: (division_by_zero) division_by_zero by (intro_classes, transfer, rule inverse_zero) instance star :: (pordered_semiring) pordered_semiring apply (intro_classes) apply (transfer, erule (1) mult_left_mono) apply (transfer, erule (1) mult_right_mono) done instance star :: (pordered_cancel_semiring) pordered_cancel_semiring .. instance star :: (ordered_semiring_strict) ordered_semiring_strict apply (intro_classes) apply (transfer, erule (1) mult_strict_left_mono) apply (transfer, erule (1) mult_strict_right_mono) done instance star :: (pordered_comm_semiring) pordered_comm_semiring by (intro_classes, transfer, rule mult_mono1_class.mult_mono1) instance star :: (pordered_cancel_comm_semiring) pordered_cancel_comm_semiring .. instance star :: (ordered_comm_semiring_strict) ordered_comm_semiring_strict by (intro_classes, transfer, rule ordered_comm_semiring_strict_class.mult_strict_left_mono_comm) instance star :: (pordered_ring) pordered_ring .. instance star :: (pordered_ring_abs) pordered_ring_abs by intro_classes (transfer, rule abs_eq_mult) instance star :: (lordered_ring) lordered_ring .. instance star :: (abs_if) abs_if by (intro_classes, transfer, rule abs_if) instance star :: (sgn_if) sgn_if by (intro_classes, transfer, rule sgn_if) instance star :: (ordered_ring_strict) ordered_ring_strict .. instance star :: (pordered_comm_ring) pordered_comm_ring .. instance star :: (ordered_semidom) ordered_semidom by (intro_classes, transfer, rule zero_less_one) instance star :: (ordered_idom) ordered_idom .. instance star :: (ordered_field) ordered_field .. subsection {* Power classes *} text {* Proving the class axiom @{thm [source] power_Suc} for type @{typ "'a star"} is a little tricky, because it quantifies over values of type @{typ nat}. The transfer principle does not handle quantification over non-star types in general, but we can work around this by fixing an arbitrary @{typ nat} value, and then applying the transfer principle. *} instance star :: (recpower) recpower proof show "\a::'a star. a ^ 0 = 1" by transfer (rule power_0) next fix n show "\a::'a star. a ^ Suc n = a * a ^ n" by transfer (rule power_Suc) qed subsection {* Number classes *} lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)" by (induct n, simp_all) lemma Standard_of_nat [simp]: "of_nat n \ Standard" by (simp add: star_of_nat_def) lemma star_of_of_nat [simp]: "star_of (of_nat n) = of_nat n" by transfer (rule refl) lemma star_of_int_def [transfer_unfold]: "of_int z = star_of (of_int z)" by (rule_tac z=z in int_diff_cases, simp) lemma Standard_of_int [simp]: "of_int z \ Standard" by (simp add: star_of_int_def) lemma star_of_of_int [simp]: "star_of (of_int z) = of_int z" by transfer (rule refl) instance star :: (semiring_char_0) semiring_char_0 by intro_classes (simp only: star_of_nat_def star_of_eq of_nat_eq_iff) instance star :: (ring_char_0) ring_char_0 .. instance star :: (number_ring) number_ring by (intro_classes, simp only: star_number_def star_of_int_def number_of_eq) subsection {* Finite class *} lemma starset_finite: "finite A \ *s* A = star_of ` A" by (erule finite_induct, simp_all) instance star :: (finite) finite apply (intro_classes) apply (subst starset_UNIV [symmetric]) apply (subst starset_finite [OF finite]) apply (rule finite_imageI [OF finite]) done end diff --git a/src/HOL/RealVector.thy b/src/HOL/RealVector.thy --- a/src/HOL/RealVector.thy +++ b/src/HOL/RealVector.thy @@ -1,834 +1,834 @@ (* Title: HOL/RealVector.thy Author: Brian Huffman *) header {* Vector Spaces and Algebras over the Reals *} theory RealVector imports RealPow begin subsection {* Locale for additive functions *} locale additive = fixes f :: "'a::ab_group_add \ 'b::ab_group_add" assumes add: "f (x + y) = f x + f y" begin lemma zero: "f 0 = 0" proof - have "f 0 = f (0 + 0)" by simp also have "\ = f 0 + f 0" by (rule add) finally show "f 0 = 0" by simp qed lemma minus: "f (- x) = - f x" proof - have "f (- x) + f x = f (- x + x)" by (rule add [symmetric]) also have "\ = - f x + f x" by (simp add: zero) finally show "f (- x) = - f x" by (rule add_right_imp_eq) qed lemma diff: "f (x - y) = f x - f y" by (simp add: diff_def add minus) lemma setsum: "f (setsum g A) = (\x\A. f (g x))" apply (cases "finite A") apply (induct set: finite) apply (simp add: zero) apply (simp add: add) apply (simp add: zero) done end subsection {* Vector spaces *} locale vector_space = fixes scale :: "'a::field \ 'b::ab_group_add \ 'b" assumes scale_right_distrib [algebra_simps]: "scale a (x + y) = scale a x + scale a y" and scale_left_distrib [algebra_simps]: "scale (a + b) x = scale a x + scale b x" and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x" and scale_one [simp]: "scale 1 x = x" begin lemma scale_left_commute: "scale a (scale b x) = scale b (scale a x)" by (simp add: mult_commute) lemma scale_zero_left [simp]: "scale 0 x = 0" and scale_minus_left [simp]: "scale (- a) x = - (scale a x)" and scale_left_diff_distrib [algebra_simps]: "scale (a - b) x = scale a x - scale b x" proof - interpret s: additive "\a. scale a x" proof qed (rule scale_left_distrib) show "scale 0 x = 0" by (rule s.zero) show "scale (- a) x = - (scale a x)" by (rule s.minus) show "scale (a - b) x = scale a x - scale b x" by (rule s.diff) qed lemma scale_zero_right [simp]: "scale a 0 = 0" and scale_minus_right [simp]: "scale a (- x) = - (scale a x)" and scale_right_diff_distrib [algebra_simps]: "scale a (x - y) = scale a x - scale a y" proof - interpret s: additive "\x. scale a x" proof qed (rule scale_right_distrib) show "scale a 0 = 0" by (rule s.zero) show "scale a (- x) = - (scale a x)" by (rule s.minus) show "scale a (x - y) = scale a x - scale a y" by (rule s.diff) qed lemma scale_eq_0_iff [simp]: "scale a x = 0 \ a = 0 \ x = 0" proof cases assume "a = 0" thus ?thesis by simp next assume anz [simp]: "a \ 0" { assume "scale a x = 0" hence "scale (inverse a) (scale a x) = 0" by simp hence "x = 0" by simp } thus ?thesis by force qed lemma scale_left_imp_eq: "\a \ 0; scale a x = scale a y\ \ x = y" proof - assume nonzero: "a \ 0" assume "scale a x = scale a y" hence "scale a (x - y) = 0" by (simp add: scale_right_diff_distrib) hence "x - y = 0" by (simp add: nonzero) thus "x = y" by (simp only: right_minus_eq) qed lemma scale_right_imp_eq: "\x \ 0; scale a x = scale b x\ \ a = b" proof - assume nonzero: "x \ 0" assume "scale a x = scale b x" hence "scale (a - b) x = 0" by (simp add: scale_left_diff_distrib) hence "a - b = 0" by (simp add: nonzero) thus "a = b" by (simp only: right_minus_eq) qed lemma scale_cancel_left: "scale a x = scale a y \ x = y \ a = 0" by (auto intro: scale_left_imp_eq) lemma scale_cancel_right: "scale a x = scale b x \ a = b \ x = 0" by (auto intro: scale_right_imp_eq) end subsection {* Real vector spaces *} class scaleR = fixes scaleR :: "real \ 'a \ 'a" (infixr "*\<^sub>R" 75) begin abbreviation divideR :: "'a \ real \ 'a" (infixl "'/\<^sub>R" 70) where "x /\<^sub>R r == scaleR (inverse r) x" end class real_vector = scaleR + ab_group_add + assumes scaleR_right_distrib: "scaleR a (x + y) = scaleR a x + scaleR a y" and scaleR_left_distrib: "scaleR (a + b) x = scaleR a x + scaleR b x" and scaleR_scaleR: "scaleR a (scaleR b x) = scaleR (a * b) x" and scaleR_one: "scaleR 1 x = x" -interpretation real_vector!: +interpretation real_vector: vector_space "scaleR :: real \ 'a \ 'a::real_vector" apply unfold_locales apply (rule scaleR_right_distrib) apply (rule scaleR_left_distrib) apply (rule scaleR_scaleR) apply (rule scaleR_one) done text {* Recover original theorem names *} lemmas scaleR_left_commute = real_vector.scale_left_commute lemmas scaleR_zero_left = real_vector.scale_zero_left lemmas scaleR_minus_left = real_vector.scale_minus_left lemmas scaleR_left_diff_distrib = real_vector.scale_left_diff_distrib lemmas scaleR_zero_right = real_vector.scale_zero_right lemmas scaleR_minus_right = real_vector.scale_minus_right lemmas scaleR_right_diff_distrib = real_vector.scale_right_diff_distrib lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq lemmas scaleR_cancel_left = real_vector.scale_cancel_left lemmas scaleR_cancel_right = real_vector.scale_cancel_right class real_algebra = real_vector + ring + assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)" and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)" class real_algebra_1 = real_algebra + ring_1 class real_div_algebra = real_algebra_1 + division_ring class real_field = real_div_algebra + field instantiation real :: real_field begin definition real_scaleR_def [simp]: "scaleR a x = a * x" instance proof qed (simp_all add: algebra_simps) end -interpretation scaleR_left!: additive "(\a. scaleR a x::'a::real_vector)" +interpretation scaleR_left: additive "(\a. scaleR a x::'a::real_vector)" proof qed (rule scaleR_left_distrib) -interpretation scaleR_right!: additive "(\x. scaleR a x::'a::real_vector)" +interpretation scaleR_right: additive "(\x. scaleR a x::'a::real_vector)" proof qed (rule scaleR_right_distrib) lemma nonzero_inverse_scaleR_distrib: fixes x :: "'a::real_div_algebra" shows "\a \ 0; x \ 0\ \ inverse (scaleR a x) = scaleR (inverse a) (inverse x)" by (rule inverse_unique, simp) lemma inverse_scaleR_distrib: fixes x :: "'a::{real_div_algebra,division_by_zero}" shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)" apply (case_tac "a = 0", simp) apply (case_tac "x = 0", simp) apply (erule (1) nonzero_inverse_scaleR_distrib) done subsection {* Embedding of the Reals into any @{text real_algebra_1}: @{term of_real} *} definition of_real :: "real \ 'a::real_algebra_1" where "of_real r = scaleR r 1" lemma scaleR_conv_of_real: "scaleR r x = of_real r * x" by (simp add: of_real_def) lemma of_real_0 [simp]: "of_real 0 = 0" by (simp add: of_real_def) lemma of_real_1 [simp]: "of_real 1 = 1" by (simp add: of_real_def) lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y" by (simp add: of_real_def scaleR_left_distrib) lemma of_real_minus [simp]: "of_real (- x) = - of_real x" by (simp add: of_real_def) lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y" by (simp add: of_real_def scaleR_left_diff_distrib) lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y" by (simp add: of_real_def mult_commute) lemma nonzero_of_real_inverse: "x \ 0 \ of_real (inverse x) = inverse (of_real x :: 'a::real_div_algebra)" by (simp add: of_real_def nonzero_inverse_scaleR_distrib) lemma of_real_inverse [simp]: "of_real (inverse x) = inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})" by (simp add: of_real_def inverse_scaleR_distrib) lemma nonzero_of_real_divide: "y \ 0 \ of_real (x / y) = (of_real x / of_real y :: 'a::real_field)" by (simp add: divide_inverse nonzero_of_real_inverse) lemma of_real_divide [simp]: "of_real (x / y) = (of_real x / of_real y :: 'a::{real_field,division_by_zero})" by (simp add: divide_inverse) lemma of_real_power [simp]: "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1,recpower}) ^ n" by (induct n) simp_all lemma of_real_eq_iff [simp]: "(of_real x = of_real y) = (x = y)" by (simp add: of_real_def scaleR_cancel_right) lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified] lemma of_real_eq_id [simp]: "of_real = (id :: real \ real)" proof fix r show "of_real r = id r" by (simp add: of_real_def) qed text{*Collapse nested embeddings*} lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n" by (induct n) auto lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z" by (cases z rule: int_diff_cases, simp) lemma of_real_number_of_eq: "of_real (number_of w) = (number_of w :: 'a::{number_ring,real_algebra_1})" by (simp add: number_of_eq) text{*Every real algebra has characteristic zero*} instance real_algebra_1 < ring_char_0 proof fix m n :: nat have "(of_real (of_nat m) = (of_real (of_nat n)::'a)) = (m = n)" by (simp only: of_real_eq_iff of_nat_eq_iff) thus "(of_nat m = (of_nat n::'a)) = (m = n)" by (simp only: of_real_of_nat_eq) qed instance real_field < field_char_0 .. subsection {* The Set of Real Numbers *} definition Reals :: "'a::real_algebra_1 set" where [code del]: "Reals = range of_real" notation (xsymbols) Reals ("\") lemma Reals_of_real [simp]: "of_real r \ Reals" by (simp add: Reals_def) lemma Reals_of_int [simp]: "of_int z \ Reals" by (subst of_real_of_int_eq [symmetric], rule Reals_of_real) lemma Reals_of_nat [simp]: "of_nat n \ Reals" by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real) lemma Reals_number_of [simp]: "(number_of w::'a::{number_ring,real_algebra_1}) \ Reals" by (subst of_real_number_of_eq [symmetric], rule Reals_of_real) lemma Reals_0 [simp]: "0 \ Reals" apply (unfold Reals_def) apply (rule range_eqI) apply (rule of_real_0 [symmetric]) done lemma Reals_1 [simp]: "1 \ Reals" apply (unfold Reals_def) apply (rule range_eqI) apply (rule of_real_1 [symmetric]) done lemma Reals_add [simp]: "\a \ Reals; b \ Reals\ \ a + b \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (rule of_real_add [symmetric]) done lemma Reals_minus [simp]: "a \ Reals \ - a \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (rule of_real_minus [symmetric]) done lemma Reals_diff [simp]: "\a \ Reals; b \ Reals\ \ a - b \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (rule of_real_diff [symmetric]) done lemma Reals_mult [simp]: "\a \ Reals; b \ Reals\ \ a * b \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (rule of_real_mult [symmetric]) done lemma nonzero_Reals_inverse: fixes a :: "'a::real_div_algebra" shows "\a \ Reals; a \ 0\ \ inverse a \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (erule nonzero_of_real_inverse [symmetric]) done lemma Reals_inverse [simp]: fixes a :: "'a::{real_div_algebra,division_by_zero}" shows "a \ Reals \ inverse a \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (rule of_real_inverse [symmetric]) done lemma nonzero_Reals_divide: fixes a b :: "'a::real_field" shows "\a \ Reals; b \ Reals; b \ 0\ \ a / b \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (erule nonzero_of_real_divide [symmetric]) done lemma Reals_divide [simp]: fixes a b :: "'a::{real_field,division_by_zero}" shows "\a \ Reals; b \ Reals\ \ a / b \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (rule of_real_divide [symmetric]) done lemma Reals_power [simp]: fixes a :: "'a::{real_algebra_1,recpower}" shows "a \ Reals \ a ^ n \ Reals" apply (auto simp add: Reals_def) apply (rule range_eqI) apply (rule of_real_power [symmetric]) done lemma Reals_cases [cases set: Reals]: assumes "q \ \" obtains (of_real) r where "q = of_real r" unfolding Reals_def proof - from `q \ \` have "q \ range of_real" unfolding Reals_def . then obtain r where "q = of_real r" .. then show thesis .. qed lemma Reals_induct [case_names of_real, induct set: Reals]: "q \ \ \ (\r. P (of_real r)) \ P q" by (rule Reals_cases) auto subsection {* Real normed vector spaces *} class norm = fixes norm :: "'a \ real" class sgn_div_norm = scaleR + norm + sgn + assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x" class real_normed_vector = real_vector + sgn_div_norm + assumes norm_ge_zero [simp]: "0 \ norm x" and norm_eq_zero [simp]: "norm x = 0 \ x = 0" and norm_triangle_ineq: "norm (x + y) \ norm x + norm y" and norm_scaleR: "norm (scaleR a x) = \a\ * norm x" class real_normed_algebra = real_algebra + real_normed_vector + assumes norm_mult_ineq: "norm (x * y) \ norm x * norm y" class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra + assumes norm_one [simp]: "norm 1 = 1" class real_normed_div_algebra = real_div_algebra + real_normed_vector + assumes norm_mult: "norm (x * y) = norm x * norm y" class real_normed_field = real_field + real_normed_div_algebra instance real_normed_div_algebra < real_normed_algebra_1 proof fix x y :: 'a show "norm (x * y) \ norm x * norm y" by (simp add: norm_mult) next have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)" by (rule norm_mult) thus "norm (1::'a) = 1" by simp qed instantiation real :: real_normed_field begin definition real_norm_def [simp]: "norm r = \r\" instance apply (intro_classes, unfold real_norm_def real_scaleR_def) apply (simp add: real_sgn_def) apply (rule abs_ge_zero) apply (rule abs_eq_0) apply (rule abs_triangle_ineq) apply (rule abs_mult) apply (rule abs_mult) done end lemma norm_zero [simp]: "norm (0::'a::real_normed_vector) = 0" by simp lemma zero_less_norm_iff [simp]: fixes x :: "'a::real_normed_vector" shows "(0 < norm x) = (x \ 0)" by (simp add: order_less_le) lemma norm_not_less_zero [simp]: fixes x :: "'a::real_normed_vector" shows "\ norm x < 0" by (simp add: linorder_not_less) lemma norm_le_zero_iff [simp]: fixes x :: "'a::real_normed_vector" shows "(norm x \ 0) = (x = 0)" by (simp add: order_le_less) lemma norm_minus_cancel [simp]: fixes x :: "'a::real_normed_vector" shows "norm (- x) = norm x" proof - have "norm (- x) = norm (scaleR (- 1) x)" by (simp only: scaleR_minus_left scaleR_one) also have "\ = \- 1\ * norm x" by (rule norm_scaleR) finally show ?thesis by simp qed lemma norm_minus_commute: fixes a b :: "'a::real_normed_vector" shows "norm (a - b) = norm (b - a)" proof - have "norm (- (b - a)) = norm (b - a)" by (rule norm_minus_cancel) thus ?thesis by simp qed lemma norm_triangle_ineq2: fixes a b :: "'a::real_normed_vector" shows "norm a - norm b \ norm (a - b)" proof - have "norm (a - b + b) \ norm (a - b) + norm b" by (rule norm_triangle_ineq) thus ?thesis by simp qed lemma norm_triangle_ineq3: fixes a b :: "'a::real_normed_vector" shows "\norm a - norm b\ \ norm (a - b)" apply (subst abs_le_iff) apply auto apply (rule norm_triangle_ineq2) apply (subst norm_minus_commute) apply (rule norm_triangle_ineq2) done lemma norm_triangle_ineq4: fixes a b :: "'a::real_normed_vector" shows "norm (a - b) \ norm a + norm b" proof - have "norm (a + - b) \ norm a + norm (- b)" by (rule norm_triangle_ineq) thus ?thesis by (simp only: diff_minus norm_minus_cancel) qed lemma norm_diff_ineq: fixes a b :: "'a::real_normed_vector" shows "norm a - norm b \ norm (a + b)" proof - have "norm a - norm (- b) \ norm (a - - b)" by (rule norm_triangle_ineq2) thus ?thesis by simp qed lemma norm_diff_triangle_ineq: fixes a b c d :: "'a::real_normed_vector" shows "norm ((a + b) - (c + d)) \ norm (a - c) + norm (b - d)" proof - have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))" by (simp add: diff_minus add_ac) also have "\ \ norm (a - c) + norm (b - d)" by (rule norm_triangle_ineq) finally show ?thesis . qed lemma abs_norm_cancel [simp]: fixes a :: "'a::real_normed_vector" shows "\norm a\ = norm a" by (rule abs_of_nonneg [OF norm_ge_zero]) lemma norm_add_less: fixes x y :: "'a::real_normed_vector" shows "\norm x < r; norm y < s\ \ norm (x + y) < r + s" by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono]) lemma norm_mult_less: fixes x y :: "'a::real_normed_algebra" shows "\norm x < r; norm y < s\ \ norm (x * y) < r * s" apply (rule order_le_less_trans [OF norm_mult_ineq]) apply (simp add: mult_strict_mono') done lemma norm_of_real [simp]: "norm (of_real r :: 'a::real_normed_algebra_1) = \r\" unfolding of_real_def by (simp add: norm_scaleR) lemma norm_number_of [simp]: "norm (number_of w::'a::{number_ring,real_normed_algebra_1}) = \number_of w\" by (subst of_real_number_of_eq [symmetric], rule norm_of_real) lemma norm_of_int [simp]: "norm (of_int z::'a::real_normed_algebra_1) = \of_int z\" by (subst of_real_of_int_eq [symmetric], rule norm_of_real) lemma norm_of_nat [simp]: "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n" apply (subst of_real_of_nat_eq [symmetric]) apply (subst norm_of_real, simp) done lemma nonzero_norm_inverse: fixes a :: "'a::real_normed_div_algebra" shows "a \ 0 \ norm (inverse a) = inverse (norm a)" apply (rule inverse_unique [symmetric]) apply (simp add: norm_mult [symmetric]) done lemma norm_inverse: fixes a :: "'a::{real_normed_div_algebra,division_by_zero}" shows "norm (inverse a) = inverse (norm a)" apply (case_tac "a = 0", simp) apply (erule nonzero_norm_inverse) done lemma nonzero_norm_divide: fixes a b :: "'a::real_normed_field" shows "b \ 0 \ norm (a / b) = norm a / norm b" by (simp add: divide_inverse norm_mult nonzero_norm_inverse) lemma norm_divide: fixes a b :: "'a::{real_normed_field,division_by_zero}" shows "norm (a / b) = norm a / norm b" by (simp add: divide_inverse norm_mult norm_inverse) lemma norm_power_ineq: fixes x :: "'a::{real_normed_algebra_1,recpower}" shows "norm (x ^ n) \ norm x ^ n" proof (induct n) case 0 show "norm (x ^ 0) \ norm x ^ 0" by simp next case (Suc n) have "norm (x * x ^ n) \ norm x * norm (x ^ n)" by (rule norm_mult_ineq) also from Suc have "\ \ norm x * norm x ^ n" using norm_ge_zero by (rule mult_left_mono) finally show "norm (x ^ Suc n) \ norm x ^ Suc n" by simp qed lemma norm_power: fixes x :: "'a::{real_normed_div_algebra,recpower}" shows "norm (x ^ n) = norm x ^ n" by (induct n) (simp_all add: norm_mult) subsection {* Sign function *} lemma norm_sgn: "norm (sgn(x::'a::real_normed_vector)) = (if x = 0 then 0 else 1)" by (simp add: sgn_div_norm norm_scaleR) lemma sgn_zero [simp]: "sgn(0::'a::real_normed_vector) = 0" by (simp add: sgn_div_norm) lemma sgn_zero_iff: "(sgn(x::'a::real_normed_vector) = 0) = (x = 0)" by (simp add: sgn_div_norm) lemma sgn_minus: "sgn (- x) = - sgn(x::'a::real_normed_vector)" by (simp add: sgn_div_norm) lemma sgn_scaleR: "sgn (scaleR r x) = scaleR (sgn r) (sgn(x::'a::real_normed_vector))" by (simp add: sgn_div_norm norm_scaleR mult_ac) lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1" by (simp add: sgn_div_norm) lemma sgn_of_real: "sgn (of_real r::'a::real_normed_algebra_1) = of_real (sgn r)" unfolding of_real_def by (simp only: sgn_scaleR sgn_one) lemma sgn_mult: fixes x y :: "'a::real_normed_div_algebra" shows "sgn (x * y) = sgn x * sgn y" by (simp add: sgn_div_norm norm_mult mult_commute) lemma real_sgn_eq: "sgn (x::real) = x / \x\" by (simp add: sgn_div_norm divide_inverse) lemma real_sgn_pos: "0 < (x::real) \ sgn x = 1" unfolding real_sgn_eq by simp lemma real_sgn_neg: "(x::real) < 0 \ sgn x = -1" unfolding real_sgn_eq by simp subsection {* Bounded Linear and Bilinear Operators *} locale bounded_linear = additive + constrains f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes scaleR: "f (scaleR r x) = scaleR r (f x)" assumes bounded: "\K. \x. norm (f x) \ norm x * K" begin lemma pos_bounded: "\K>0. \x. norm (f x) \ norm x * K" proof - obtain K where K: "\x. norm (f x) \ norm x * K" using bounded by fast show ?thesis proof (intro exI impI conjI allI) show "0 < max 1 K" by (rule order_less_le_trans [OF zero_less_one le_maxI1]) next fix x have "norm (f x) \ norm x * K" using K . also have "\ \ norm x * max 1 K" by (rule mult_left_mono [OF le_maxI2 norm_ge_zero]) finally show "norm (f x) \ norm x * max 1 K" . qed qed lemma nonneg_bounded: "\K\0. \x. norm (f x) \ norm x * K" proof - from pos_bounded show ?thesis by (auto intro: order_less_imp_le) qed end locale bounded_bilinear = fixes prod :: "['a::real_normed_vector, 'b::real_normed_vector] \ 'c::real_normed_vector" (infixl "**" 70) assumes add_left: "prod (a + a') b = prod a b + prod a' b" assumes add_right: "prod a (b + b') = prod a b + prod a b'" assumes scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)" assumes scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)" assumes bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" begin lemma pos_bounded: "\K>0. \a b. norm (a ** b) \ norm a * norm b * K" apply (cut_tac bounded, erule exE) apply (rule_tac x="max 1 K" in exI, safe) apply (rule order_less_le_trans [OF zero_less_one le_maxI1]) apply (drule spec, drule spec, erule order_trans) apply (rule mult_left_mono [OF le_maxI2]) apply (intro mult_nonneg_nonneg norm_ge_zero) done lemma nonneg_bounded: "\K\0. \a b. norm (a ** b) \ norm a * norm b * K" proof - from pos_bounded show ?thesis by (auto intro: order_less_imp_le) qed lemma additive_right: "additive (\b. prod a b)" by (rule additive.intro, rule add_right) lemma additive_left: "additive (\a. prod a b)" by (rule additive.intro, rule add_left) lemma zero_left: "prod 0 b = 0" by (rule additive.zero [OF additive_left]) lemma zero_right: "prod a 0 = 0" by (rule additive.zero [OF additive_right]) lemma minus_left: "prod (- a) b = - prod a b" by (rule additive.minus [OF additive_left]) lemma minus_right: "prod a (- b) = - prod a b" by (rule additive.minus [OF additive_right]) lemma diff_left: "prod (a - a') b = prod a b - prod a' b" by (rule additive.diff [OF additive_left]) lemma diff_right: "prod a (b - b') = prod a b - prod a b'" by (rule additive.diff [OF additive_right]) lemma bounded_linear_left: "bounded_linear (\a. a ** b)" apply (unfold_locales) apply (rule add_left) apply (rule scaleR_left) apply (cut_tac bounded, safe) apply (rule_tac x="norm b * K" in exI) apply (simp add: mult_ac) done lemma bounded_linear_right: "bounded_linear (\b. a ** b)" apply (unfold_locales) apply (rule add_right) apply (rule scaleR_right) apply (cut_tac bounded, safe) apply (rule_tac x="norm a * K" in exI) apply (simp add: mult_ac) done lemma prod_diff_prod: "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)" by (simp add: diff_left diff_right) end -interpretation mult!: +interpretation mult: bounded_bilinear "op * :: 'a \ 'a \ 'a::real_normed_algebra" apply (rule bounded_bilinear.intro) apply (rule left_distrib) apply (rule right_distrib) apply (rule mult_scaleR_left) apply (rule mult_scaleR_right) apply (rule_tac x="1" in exI) apply (simp add: norm_mult_ineq) done -interpretation mult_left!: +interpretation mult_left: bounded_linear "(\x::'a::real_normed_algebra. x * y)" by (rule mult.bounded_linear_left) -interpretation mult_right!: +interpretation mult_right: bounded_linear "(\y::'a::real_normed_algebra. x * y)" by (rule mult.bounded_linear_right) -interpretation divide!: +interpretation divide: bounded_linear "(\x::'a::real_normed_field. x / y)" unfolding divide_inverse by (rule mult.bounded_linear_left) -interpretation scaleR!: bounded_bilinear "scaleR" +interpretation scaleR: bounded_bilinear "scaleR" apply (rule bounded_bilinear.intro) apply (rule scaleR_left_distrib) apply (rule scaleR_right_distrib) apply simp apply (rule scaleR_left_commute) apply (rule_tac x="1" in exI) apply (simp add: norm_scaleR) done -interpretation scaleR_left!: bounded_linear "\r. scaleR r x" +interpretation scaleR_left: bounded_linear "\r. scaleR r x" by (rule scaleR.bounded_linear_left) -interpretation scaleR_right!: bounded_linear "\x. scaleR r x" +interpretation scaleR_right: bounded_linear "\x. scaleR r x" by (rule scaleR.bounded_linear_right) -interpretation of_real!: bounded_linear "\r. of_real r" +interpretation of_real: bounded_linear "\r. of_real r" unfolding of_real_def by (rule scaleR.bounded_linear_left) end diff --git a/src/HOL/Word/TdThs.thy b/src/HOL/Word/TdThs.thy --- a/src/HOL/Word/TdThs.thy +++ b/src/HOL/Word/TdThs.thy @@ -1,228 +1,228 @@ (* Author: Jeremy Dawson and Gerwin Klein, NICTA consequences of type definition theorems, and of extended type definition theorems *) header {* Type Definition Theorems *} theory TdThs imports Main begin section "More lemmas about normal type definitions" lemma tdD1: "type_definition Rep Abs A \ \x. Rep x \ A" and tdD2: "type_definition Rep Abs A \ \x. Abs (Rep x) = x" and tdD3: "type_definition Rep Abs A \ \y. y \ A \ Rep (Abs y) = y" by (auto simp: type_definition_def) lemma td_nat_int: "type_definition int nat (Collect (op <= 0))" unfolding type_definition_def by auto context type_definition begin lemmas Rep' [iff] = Rep [simplified] (* if A is given as Collect .. *) declare Rep_inverse [simp] Rep_inject [simp] lemma Abs_eqD: "Abs x = Abs y ==> x \ A ==> y \ A ==> x = y" by (simp add: Abs_inject) lemma Abs_inverse': "r : A ==> Abs r = a ==> Rep a = r" by (safe elim!: Abs_inverse) lemma Rep_comp_inverse: "Rep o f = g ==> Abs o g = f" using Rep_inverse by (auto intro: ext) lemma Rep_eqD [elim!]: "Rep x = Rep y ==> x = y" by simp lemma Rep_inverse': "Rep a = r ==> Abs r = a" by (safe intro!: Rep_inverse) lemma comp_Abs_inverse: "f o Abs = g ==> g o Rep = f" using Rep_inverse by (auto intro: ext) lemma set_Rep: "A = range Rep" proof (rule set_ext) fix x show "(x \ A) = (x \ range Rep)" by (auto dest: Abs_inverse [of x, symmetric]) qed lemma set_Rep_Abs: "A = range (Rep o Abs)" proof (rule set_ext) fix x show "(x \ A) = (x \ range (Rep o Abs))" by (auto dest: Abs_inverse [of x, symmetric]) qed lemma Abs_inj_on: "inj_on Abs A" unfolding inj_on_def by (auto dest: Abs_inject [THEN iffD1]) lemma image: "Abs ` A = UNIV" by (auto intro!: image_eqI) lemmas td_thm = type_definition_axioms lemma fns1: "Rep o fa = fr o Rep | fa o Abs = Abs o fr ==> Abs o fr o Rep = fa" by (auto dest: Rep_comp_inverse elim: comp_Abs_inverse simp: o_assoc) lemmas fns1a = disjI1 [THEN fns1] lemmas fns1b = disjI2 [THEN fns1] lemma fns4: "Rep o fa o Abs = fr ==> Rep o fa = fr o Rep & fa o Abs = Abs o fr" by (auto intro!: ext) end -interpretation nat_int!: type_definition int nat "Collect (op <= 0)" +interpretation nat_int: type_definition int nat "Collect (op <= 0)" by (rule td_nat_int) declare nat_int.Rep_cases [cases del] nat_int.Abs_cases [cases del] nat_int.Rep_induct [induct del] nat_int.Abs_induct [induct del] subsection "Extended form of type definition predicate" lemma td_conds: "norm o norm = norm ==> (fr o norm = norm o fr) = (norm o fr o norm = fr o norm & norm o fr o norm = norm o fr)" apply safe apply (simp_all add: o_assoc [symmetric]) apply (simp_all add: o_assoc) done lemma fn_comm_power: "fa o tr = tr o fr ==> fa ^ n o tr = tr o fr ^ n" apply (rule ext) apply (induct n) apply (auto dest: fun_cong) done lemmas fn_comm_power' = ext [THEN fn_comm_power, THEN fun_cong, unfolded o_def, standard] locale td_ext = type_definition + fixes norm assumes eq_norm: "\x. Rep (Abs x) = norm x" begin lemma Abs_norm [simp]: "Abs (norm x) = Abs x" using eq_norm [of x] by (auto elim: Rep_inverse') lemma td_th: "g o Abs = f ==> f (Rep x) = g x" by (drule comp_Abs_inverse [symmetric]) simp lemma eq_norm': "Rep o Abs = norm" by (auto simp: eq_norm intro!: ext) lemma norm_Rep [simp]: "norm (Rep x) = Rep x" by (auto simp: eq_norm' intro: td_th) lemmas td = td_thm lemma set_iff_norm: "w : A <-> w = norm w" by (auto simp: set_Rep_Abs eq_norm' eq_norm [symmetric]) lemma inverse_norm: "(Abs n = w) = (Rep w = norm n)" apply (rule iffI) apply (clarsimp simp add: eq_norm) apply (simp add: eq_norm' [symmetric]) done lemma norm_eq_iff: "(norm x = norm y) = (Abs x = Abs y)" by (simp add: eq_norm' [symmetric]) lemma norm_comps: "Abs o norm = Abs" "norm o Rep = Rep" "norm o norm = norm" by (auto simp: eq_norm' [symmetric] o_def) lemmas norm_norm [simp] = norm_comps lemma fns5: "Rep o fa o Abs = fr ==> fr o norm = fr & norm o fr = fr" by (fold eq_norm') (auto intro!: ext) (* following give conditions for converses to td_fns1 the condition (norm o fr o norm = fr o norm) says that fr takes normalised arguments to normalised results, (norm o fr o norm = norm o fr) says that fr takes norm-equivalent arguments to norm-equivalent results, (fr o norm = fr) says that fr takes norm-equivalent arguments to the same result, and (norm o fr = fr) says that fr takes any argument to a normalised result *) lemma fns2: "Abs o fr o Rep = fa ==> (norm o fr o norm = fr o norm) = (Rep o fa = fr o Rep)" apply (fold eq_norm') apply safe prefer 2 apply (simp add: o_assoc) apply (rule ext) apply (drule_tac x="Rep x" in fun_cong) apply auto done lemma fns3: "Abs o fr o Rep = fa ==> (norm o fr o norm = norm o fr) = (fa o Abs = Abs o fr)" apply (fold eq_norm') apply safe prefer 2 apply (simp add: o_assoc [symmetric]) apply (rule ext) apply (drule fun_cong) apply simp done lemma fns: "fr o norm = norm o fr ==> (fa o Abs = Abs o fr) = (Rep o fa = fr o Rep)" apply safe apply (frule fns1b) prefer 2 apply (frule fns1a) apply (rule fns3 [THEN iffD1]) prefer 3 apply (rule fns2 [THEN iffD1]) apply (simp_all add: o_assoc [symmetric]) apply (simp_all add: o_assoc) done lemma range_norm: "range (Rep o Abs) = A" by (simp add: set_Rep_Abs) end lemmas td_ext_def' = td_ext_def [unfolded type_definition_def td_ext_axioms_def] end diff --git a/src/HOL/Word/WordArith.thy b/src/HOL/Word/WordArith.thy --- a/src/HOL/Word/WordArith.thy +++ b/src/HOL/Word/WordArith.thy @@ -1,1262 +1,1262 @@ (* Author: Jeremy Dawson and Gerwin Klein, NICTA contains arithmetic theorems for word, instantiations to arithmetic type classes and tactics for reducing word arithmetic to linear arithmetic on int or nat *) header {* Word Arithmetic *} theory WordArith imports WordDefinition begin lemma word_less_alt: "(a < b) = (uint a < uint b)" unfolding word_less_def word_le_def by (auto simp del: word_uint.Rep_inject simp: word_uint.Rep_inject [symmetric]) lemma signed_linorder: "linorder word_sle word_sless" proof qed (unfold word_sle_def word_sless_def, auto) -interpretation signed!: linorder "word_sle" "word_sless" +interpretation signed: linorder "word_sle" "word_sless" by (rule signed_linorder) lemmas word_arith_wis = word_add_def word_mult_def word_minus_def word_succ_def word_pred_def word_0_wi word_1_wi lemma udvdI: "0 \ n ==> uint b = n * uint a ==> a udvd b" by (auto simp: udvd_def) lemmas word_div_no [simp] = word_div_def [of "number_of a" "number_of b", standard] lemmas word_mod_no [simp] = word_mod_def [of "number_of a" "number_of b", standard] lemmas word_less_no [simp] = word_less_def [of "number_of a" "number_of b", standard] lemmas word_le_no [simp] = word_le_def [of "number_of a" "number_of b", standard] lemmas word_sless_no [simp] = word_sless_def [of "number_of a" "number_of b", standard] lemmas word_sle_no [simp] = word_sle_def [of "number_of a" "number_of b", standard] (* following two are available in class number_ring, but convenient to have them here here; note - the number_ring versions, numeral_0_eq_0 and numeral_1_eq_1 are in the default simpset, so to use the automatic simplifications for (eg) sint (number_of bin) on sint 1, must do (simp add: word_1_no del: numeral_1_eq_1) *) lemmas word_0_wi_Pls = word_0_wi [folded Pls_def] lemmas word_0_no = word_0_wi_Pls [folded word_no_wi] lemma int_one_bin: "(1 :: int) == (Int.Pls BIT bit.B1)" unfolding Pls_def Bit_def by auto lemma word_1_no: "(1 :: 'a :: len0 word) == number_of (Int.Pls BIT bit.B1)" unfolding word_1_wi word_number_of_def int_one_bin by auto lemma word_m1_wi: "-1 == word_of_int -1" by (rule word_number_of_alt) lemma word_m1_wi_Min: "-1 = word_of_int Int.Min" by (simp add: word_m1_wi number_of_eq) lemma word_0_bl: "of_bl [] = 0" unfolding word_0_wi of_bl_def by (simp add : Pls_def) lemma word_1_bl: "of_bl [True] = 1" unfolding word_1_wi of_bl_def by (simp add : bl_to_bin_def Bit_def Pls_def) lemma uint_0 [simp] : "(uint 0 = 0)" unfolding word_0_wi by (simp add: word_ubin.eq_norm Pls_def [symmetric]) lemma of_bl_0 [simp] : "of_bl (replicate n False) = 0" by (simp add : word_0_wi of_bl_def bl_to_bin_rep_False Pls_def) lemma to_bl_0: "to_bl (0::'a::len0 word) = replicate (len_of TYPE('a)) False" unfolding uint_bl by (simp add : word_size bin_to_bl_Pls Pls_def [symmetric]) lemma uint_0_iff: "(uint x = 0) = (x = 0)" by (auto intro!: word_uint.Rep_eqD) lemma unat_0_iff: "(unat x = 0) = (x = 0)" unfolding unat_def by (auto simp add : nat_eq_iff uint_0_iff) lemma unat_0 [simp]: "unat 0 = 0" unfolding unat_def by auto lemma size_0_same': "size w = 0 ==> w = (v :: 'a :: len0 word)" apply (unfold word_size) apply (rule box_equals) defer apply (rule word_uint.Rep_inverse)+ apply (rule word_ubin.norm_eq_iff [THEN iffD1]) apply simp done lemmas size_0_same = size_0_same' [folded word_size] lemmas unat_eq_0 = unat_0_iff lemmas unat_eq_zero = unat_0_iff lemma unat_gt_0: "(0 < unat x) = (x ~= 0)" by (auto simp: unat_0_iff [symmetric]) lemma ucast_0 [simp] : "ucast 0 = 0" unfolding ucast_def by simp (simp add: word_0_wi) lemma sint_0 [simp] : "sint 0 = 0" unfolding sint_uint by (simp add: Pls_def [symmetric]) lemma scast_0 [simp] : "scast 0 = 0" apply (unfold scast_def) apply simp apply (simp add: word_0_wi) done lemma sint_n1 [simp] : "sint -1 = -1" apply (unfold word_m1_wi_Min) apply (simp add: word_sbin.eq_norm) apply (unfold Min_def number_of_eq) apply simp done lemma scast_n1 [simp] : "scast -1 = -1" apply (unfold scast_def sint_n1) apply (unfold word_number_of_alt) apply (rule refl) done lemma uint_1 [simp] : "uint (1 :: 'a :: len word) = 1" unfolding word_1_wi by (simp add: word_ubin.eq_norm int_one_bin bintrunc_minus_simps) lemma unat_1 [simp] : "unat (1 :: 'a :: len word) = 1" by (unfold unat_def uint_1) auto lemma ucast_1 [simp] : "ucast (1 :: 'a :: len word) = 1" unfolding ucast_def word_1_wi by (simp add: word_ubin.eq_norm int_one_bin bintrunc_minus_simps) (* abstraction preserves the operations (the definitions tell this for bins in range uint) *) lemmas arths = bintr_ariths [THEN word_ubin.norm_eq_iff [THEN iffD1], folded word_ubin.eq_norm, standard] lemma wi_homs: shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" and wi_hom_neg: "- word_of_int a = word_of_int (- a)" and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (Int.succ a)" and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (Int.pred a)" by (auto simp: word_arith_wis arths) lemmas wi_hom_syms = wi_homs [symmetric] lemma word_sub_def: "a - b == a + - (b :: 'a :: len0 word)" unfolding word_sub_wi diff_def by (simp only : word_uint.Rep_inverse wi_hom_syms) lemmas word_diff_minus = word_sub_def [THEN meta_eq_to_obj_eq, standard] lemma word_of_int_sub_hom: "(word_of_int a) - word_of_int b = word_of_int (a - b)" unfolding word_sub_def diff_def by (simp only : wi_homs) lemmas new_word_of_int_homs = word_of_int_sub_hom wi_homs word_0_wi word_1_wi lemmas new_word_of_int_hom_syms = new_word_of_int_homs [symmetric, standard] lemmas word_of_int_hom_syms = new_word_of_int_hom_syms [unfolded succ_def pred_def] lemmas word_of_int_homs = new_word_of_int_homs [unfolded succ_def pred_def] lemmas word_of_int_add_hom = word_of_int_homs (2) lemmas word_of_int_mult_hom = word_of_int_homs (3) lemmas word_of_int_minus_hom = word_of_int_homs (4) lemmas word_of_int_succ_hom = word_of_int_homs (5) lemmas word_of_int_pred_hom = word_of_int_homs (6) lemmas word_of_int_0_hom = word_of_int_homs (7) lemmas word_of_int_1_hom = word_of_int_homs (8) (* now, to get the weaker results analogous to word_div/mod_def *) lemmas word_arith_alts = word_sub_wi [unfolded succ_def pred_def, standard] word_arith_wis [unfolded succ_def pred_def, standard] lemmas word_sub_alt = word_arith_alts (1) lemmas word_add_alt = word_arith_alts (2) lemmas word_mult_alt = word_arith_alts (3) lemmas word_minus_alt = word_arith_alts (4) lemmas word_succ_alt = word_arith_alts (5) lemmas word_pred_alt = word_arith_alts (6) lemmas word_0_alt = word_arith_alts (7) lemmas word_1_alt = word_arith_alts (8) subsection "Transferring goals from words to ints" lemma word_ths: shows word_succ_p1: "word_succ a = a + 1" and word_pred_m1: "word_pred a = a - 1" and word_pred_succ: "word_pred (word_succ a) = a" and word_succ_pred: "word_succ (word_pred a) = a" and word_mult_succ: "word_succ a * b = b + a * b" by (rule word_uint.Abs_cases [of b], rule word_uint.Abs_cases [of a], simp add: pred_def succ_def add_commute mult_commute ring_distribs new_word_of_int_homs)+ lemmas uint_cong = arg_cong [where f = uint] lemmas uint_word_ariths = word_arith_alts [THEN trans [OF uint_cong int_word_uint], standard] lemmas uint_word_arith_bintrs = uint_word_ariths [folded bintrunc_mod2p] (* similar expressions for sint (arith operations) *) lemmas sint_word_ariths = uint_word_arith_bintrs [THEN uint_sint [symmetric, THEN trans], unfolded uint_sint bintr_arith1s bintr_ariths len_gt_0 [THEN bin_sbin_eq_iff'] word_sbin.norm_Rep, standard] lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint], standard] lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint], standard] lemma word_pred_0_n1: "word_pred 0 = word_of_int -1" unfolding word_pred_def number_of_eq by (simp add : pred_def word_no_wi) lemma word_pred_0_Min: "word_pred 0 = word_of_int Int.Min" by (simp add: word_pred_0_n1 number_of_eq) lemma word_m1_Min: "- 1 = word_of_int Int.Min" unfolding Min_def by (simp only: word_of_int_hom_syms) lemma succ_pred_no [simp]: "word_succ (number_of bin) = number_of (Int.succ bin) & word_pred (number_of bin) = number_of (Int.pred bin)" unfolding word_number_of_def by (simp add : new_word_of_int_homs) lemma word_sp_01 [simp] : "word_succ -1 = 0 & word_succ 0 = 1 & word_pred 0 = -1 & word_pred 1 = 0" by (unfold word_0_no word_1_no) auto (* alternative approach to lifting arithmetic equalities *) lemma word_of_int_Ex: "\y. x = word_of_int y" by (rule_tac x="uint x" in exI) simp lemma word_arith_eqs: fixes a :: "'a::len0 word" fixes b :: "'a::len0 word" shows word_add_0: "0 + a = a" and word_add_0_right: "a + 0 = a" and word_mult_1: "1 * a = a" and word_mult_1_right: "a * 1 = a" and word_add_commute: "a + b = b + a" and word_add_assoc: "a + b + c = a + (b + c)" and word_add_left_commute: "a + (b + c) = b + (a + c)" and word_mult_commute: "a * b = b * a" and word_mult_assoc: "a * b * c = a * (b * c)" and word_mult_left_commute: "a * (b * c) = b * (a * c)" and word_left_distrib: "(a + b) * c = a * c + b * c" and word_right_distrib: "a * (b + c) = a * b + a * c" and word_left_minus: "- a + a = 0" and word_diff_0_right: "a - 0 = a" and word_diff_self: "a - a = 0" using word_of_int_Ex [of a] word_of_int_Ex [of b] word_of_int_Ex [of c] by (auto simp: word_of_int_hom_syms [symmetric] zadd_0_right add_commute add_assoc add_left_commute mult_commute mult_assoc mult_left_commute left_distrib right_distrib) lemmas word_add_ac = word_add_commute word_add_assoc word_add_left_commute lemmas word_mult_ac = word_mult_commute word_mult_assoc word_mult_left_commute lemmas word_plus_ac0 = word_add_0 word_add_0_right word_add_ac lemmas word_times_ac1 = word_mult_1 word_mult_1_right word_mult_ac subsection "Order on fixed-length words" lemma word_order_trans: "x <= y ==> y <= z ==> x <= (z :: 'a :: len0 word)" unfolding word_le_def by auto lemma word_order_refl: "z <= (z :: 'a :: len0 word)" unfolding word_le_def by auto lemma word_order_antisym: "x <= y ==> y <= x ==> x = (y :: 'a :: len0 word)" unfolding word_le_def by (auto intro!: word_uint.Rep_eqD) lemma word_order_linear: "y <= x | x <= (y :: 'a :: len0 word)" unfolding word_le_def by auto lemma word_zero_le [simp] : "0 <= (y :: 'a :: len0 word)" unfolding word_le_def by auto instance word :: (len0) semigroup_add by intro_classes (simp add: word_add_assoc) instance word :: (len0) linorder by intro_classes (auto simp: word_less_def word_le_def) instance word :: (len0) ring by intro_classes (auto simp: word_arith_eqs word_diff_minus word_diff_self [unfolded word_diff_minus]) lemma word_m1_ge [simp] : "word_pred 0 >= y" unfolding word_le_def by (simp only : word_pred_0_n1 word_uint.eq_norm m1mod2k) auto lemmas word_n1_ge [simp] = word_m1_ge [simplified word_sp_01] lemmas word_not_simps [simp] = word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] lemma word_gt_0: "0 < y = (0 ~= (y :: 'a :: len0 word))" unfolding word_less_def by auto lemmas word_gt_0_no [simp] = word_gt_0 [of "number_of y", standard] lemma word_sless_alt: "(a =0. unat b = n * unat a)" apply (unfold udvd_def) apply safe apply (simp add: unat_def nat_mult_distrib) apply (simp add: uint_nat int_mult) apply (rule exI) apply safe prefer 2 apply (erule notE) apply (rule refl) apply force done lemma udvd_iff_dvd: "x udvd y <-> unat x dvd unat y" unfolding dvd_def udvd_nat_alt by force lemmas unat_mono = word_less_nat_alt [THEN iffD1, standard] lemma word_zero_neq_one: "0 < len_of TYPE ('a :: len0) ==> (0 :: 'a word) ~= 1"; unfolding word_arith_wis by (auto simp add: word_ubin.norm_eq_iff [symmetric] gr0_conv_Suc) lemmas lenw1_zero_neq_one = len_gt_0 [THEN word_zero_neq_one] lemma no_no [simp] : "number_of (number_of b) = number_of b" by (simp add: number_of_eq) lemma unat_minus_one: "x ~= 0 ==> unat (x - 1) = unat x - 1" apply (unfold unat_def) apply (simp only: int_word_uint word_arith_alts rdmods) apply (subgoal_tac "uint x >= 1") prefer 2 apply (drule contrapos_nn) apply (erule word_uint.Rep_inverse' [symmetric]) apply (insert uint_ge_0 [of x])[1] apply arith apply (rule box_equals) apply (rule nat_diff_distrib) prefer 2 apply assumption apply simp apply (subst mod_pos_pos_trivial) apply arith apply (insert uint_lt2p [of x])[1] apply arith apply (rule refl) apply simp done lemma measure_unat: "p ~= 0 ==> unat (p - 1) < unat p" by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0, standard] lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0, standard] lemma uint_sub_lt2p [simp]: "uint (x :: 'a :: len0 word) - uint (y :: 'b :: len0 word) < 2 ^ len_of TYPE('a)" using uint_ge_0 [of y] uint_lt2p [of x] by arith subsection "Conditions for the addition (etc) of two words to overflow" lemma uint_add_lem: "(uint x + uint y < 2 ^ len_of TYPE('a)) = (uint (x + y :: 'a :: len0 word) = uint x + uint y)" by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) lemma uint_mult_lem: "(uint x * uint y < 2 ^ len_of TYPE('a)) = (uint (x * y :: 'a :: len0 word) = uint x * uint y)" by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) lemma uint_sub_lem: "(uint x >= uint y) = (uint (x - y) = uint x - uint y)" by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) lemma uint_add_le: "uint (x + y) <= uint x + uint y" unfolding uint_word_ariths by (auto simp: mod_add_if_z) lemma uint_sub_ge: "uint (x - y) >= uint x - uint y" unfolding uint_word_ariths by (auto simp: mod_sub_if_z) lemmas uint_sub_if' = trans [OF uint_word_ariths(1) mod_sub_if_z, simplified, standard] lemmas uint_plus_if' = trans [OF uint_word_ariths(2) mod_add_if_z, simplified, standard] subsection {* Definition of uint\_arith *} lemma word_of_int_inverse: "word_of_int r = a ==> 0 <= r ==> r < 2 ^ len_of TYPE('a) ==> uint (a::'a::len0 word) = r" apply (erule word_uint.Abs_inverse' [rotated]) apply (simp add: uints_num) done lemma uint_split: fixes x::"'a::len0 word" shows "P (uint x) = (ALL i. word_of_int i = x & 0 <= i & i < 2^len_of TYPE('a) --> P i)" apply (fold word_int_case_def) apply (auto dest!: word_of_int_inverse simp: int_word_uint int_mod_eq' split: word_int_split) done lemma uint_split_asm: fixes x::"'a::len0 word" shows "P (uint x) = (~(EX i. word_of_int i = x & 0 <= i & i < 2^len_of TYPE('a) & ~ P i))" by (auto dest!: word_of_int_inverse simp: int_word_uint int_mod_eq' split: uint_split) lemmas uint_splits = uint_split uint_split_asm lemmas uint_arith_simps = word_le_def word_less_alt word_uint.Rep_inject [symmetric] uint_sub_if' uint_plus_if' (* use this to stop, eg, 2 ^ len_of TYPE (32) being simplified *) lemma power_False_cong: "False ==> a ^ b = c ^ d" by auto (* uint_arith_tac: reduce to arithmetic on int, try to solve by arith *) ML {* fun uint_arith_ss_of ss = ss addsimps @{thms uint_arith_simps} delsimps @{thms word_uint.Rep_inject} addsplits @{thms split_if_asm} addcongs @{thms power_False_cong} fun uint_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.verbose_arith_tac ctxt n t handle COOPER => Seq.empty; val cs = local_claset_of ctxt; val ss = local_simpset_of ctxt; in [ clarify_tac cs 1, full_simp_tac (uint_arith_ss_of ss) 1, ALLGOALS (full_simp_tac (HOL_ss addsplits @{thms uint_splits} addcongs @{thms power_False_cong})), rewrite_goals_tac @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN REPEAT (etac conjE n) THEN REPEAT (dtac @{thm word_of_int_inverse} n THEN atac n THEN atac n)), TRYALL arith_tac' ] end fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) *} method_setup uint_arith = {* Scan.succeed (SIMPLE_METHOD' o uint_arith_tac) *} "solving word arithmetic via integers and arith" subsection "More on overflows and monotonicity" lemma no_plus_overflow_uint_size: "((x :: 'a :: len0 word) <= x + y) = (uint x + uint y < 2 ^ size x)" unfolding word_size by uint_arith lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] lemma no_ulen_sub: "((x :: 'a :: len0 word) >= x - y) = (uint y <= uint x)" by uint_arith lemma no_olen_add': fixes x :: "'a::len0 word" shows "(x \ y + x) = (uint y + uint x < 2 ^ len_of TYPE('a))" by (simp add: word_add_ac add_ac no_olen_add) lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric], standard] lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem, standard] lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1, standard] lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem, standard] lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] lemmas word_sub_le = word_sub_le_iff [THEN iffD2, standard] lemma word_less_sub1: "(x :: 'a :: len word) ~= 0 ==> (1 < x) = (0 < x - 1)" by uint_arith lemma word_le_sub1: "(x :: 'a :: len word) ~= 0 ==> (1 <= x) = (0 <= x - 1)" by uint_arith lemma sub_wrap_lt: "((x :: 'a :: len0 word) < x - z) = (x < z)" by uint_arith lemma sub_wrap: "((x :: 'a :: len0 word) <= x - z) = (z = 0 | x < z)" by uint_arith lemma plus_minus_not_NULL_ab: "(x :: 'a :: len0 word) <= ab - c ==> c <= ab ==> c ~= 0 ==> x + c ~= 0" by uint_arith lemma plus_minus_no_overflow_ab: "(x :: 'a :: len0 word) <= ab - c ==> c <= ab ==> x <= x + c" by uint_arith lemma le_minus': "(a :: 'a :: len0 word) + c <= b ==> a <= a + c ==> c <= b - a" by uint_arith lemma le_plus': "(a :: 'a :: len0 word) <= b ==> c <= b - a ==> a + c <= b" by uint_arith lemmas le_plus = le_plus' [rotated] lemmas le_minus = leD [THEN thin_rl, THEN le_minus', standard] lemma word_plus_mono_right: "(y :: 'a :: len0 word) <= z ==> x <= x + z ==> x + y <= x + z" by uint_arith lemma word_less_minus_cancel: "y - x < z - x ==> x <= z ==> (y :: 'a :: len0 word) < z" by uint_arith lemma word_less_minus_mono_left: "(y :: 'a :: len0 word) < z ==> x <= y ==> y - x < z - x" by uint_arith lemma word_less_minus_mono: "a < c ==> d < b ==> a - b < a ==> c - d < c ==> a - b < c - (d::'a::len word)" by uint_arith lemma word_le_minus_cancel: "y - x <= z - x ==> x <= z ==> (y :: 'a :: len0 word) <= z" by uint_arith lemma word_le_minus_mono_left: "(y :: 'a :: len0 word) <= z ==> x <= y ==> y - x <= z - x" by uint_arith lemma word_le_minus_mono: "a <= c ==> d <= b ==> a - b <= a ==> c - d <= c ==> a - b <= c - (d::'a::len word)" by uint_arith lemma plus_le_left_cancel_wrap: "(x :: 'a :: len0 word) + y' < x ==> x + y < x ==> (x + y' < x + y) = (y' < y)" by uint_arith lemma plus_le_left_cancel_nowrap: "(x :: 'a :: len0 word) <= x + y' ==> x <= x + y ==> (x + y' < x + y) = (y' < y)" by uint_arith lemma word_plus_mono_right2: "(a :: 'a :: len0 word) <= a + b ==> c <= b ==> a <= a + c" by uint_arith lemma word_less_add_right: "(x :: 'a :: len0 word) < y - z ==> z <= y ==> x + z < y" by uint_arith lemma word_less_sub_right: "(x :: 'a :: len0 word) < y + z ==> y <= x ==> x - y < z" by uint_arith lemma word_le_plus_either: "(x :: 'a :: len0 word) <= y | x <= z ==> y <= y + z ==> x <= y + z" by uint_arith lemma word_less_nowrapI: "(x :: 'a :: len0 word) < z - k ==> k <= z ==> 0 < k ==> x < x + k" by uint_arith lemma inc_le: "(i :: 'a :: len word) < m ==> i + 1 <= m" by uint_arith lemma inc_i: "(1 :: 'a :: len word) <= i ==> i < m ==> 1 <= (i + 1) & i + 1 <= m" by uint_arith lemma udvd_incr_lem: "up < uq ==> up = ua + n * uint K ==> uq = ua + n' * uint K ==> up + uint K <= uq" apply clarsimp apply (drule less_le_mult) apply safe done lemma udvd_incr': "p < q ==> uint p = ua + n * uint K ==> uint q = ua + n' * uint K ==> p + K <= q" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (erule uint_add_le [THEN order_trans]) done lemma udvd_decr': "p < q ==> uint p = ua + n * uint K ==> uint q = ua + n' * uint K ==> p <= q - K" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (drule le_diff_eq [THEN iffD2]) apply (erule order_trans) apply (rule uint_sub_ge) done lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, simplified] lemmas udvd_incr0 = udvd_incr' [where ua=0, simplified] lemmas udvd_decr0 = udvd_decr' [where ua=0, simplified] lemma udvd_minus_le': "xy < k ==> z udvd xy ==> z udvd k ==> xy <= k - z" apply (unfold udvd_def) apply clarify apply (erule (2) udvd_decr0) done ML{*Delsimprocs cancel_factors*} lemma udvd_incr2_K: "p < a + s ==> a <= a + s ==> K udvd s ==> K udvd p - a ==> a <= p ==> 0 < K ==> p <= p + K & p + K <= a + s" apply (unfold udvd_def) apply clarify apply (simp add: uint_arith_simps split: split_if_asm) prefer 2 apply (insert uint_range' [of s])[1] apply arith apply (drule add_commute [THEN xtr1]) apply (simp add: diff_less_eq [symmetric]) apply (drule less_le_mult) apply arith apply simp done ML{*Delsimprocs cancel_factors*} (* links with rbl operations *) lemma word_succ_rbl: "to_bl w = bl ==> to_bl (word_succ w) = (rev (rbl_succ (rev bl)))" apply (unfold word_succ_def) apply clarify apply (simp add: to_bl_of_bin) apply (simp add: to_bl_def rbl_succ) done lemma word_pred_rbl: "to_bl w = bl ==> to_bl (word_pred w) = (rev (rbl_pred (rev bl)))" apply (unfold word_pred_def) apply clarify apply (simp add: to_bl_of_bin) apply (simp add: to_bl_def rbl_pred) done lemma word_add_rbl: "to_bl v = vbl ==> to_bl w = wbl ==> to_bl (v + w) = (rev (rbl_add (rev vbl) (rev wbl)))" apply (unfold word_add_def) apply clarify apply (simp add: to_bl_of_bin) apply (simp add: to_bl_def rbl_add) done lemma word_mult_rbl: "to_bl v = vbl ==> to_bl w = wbl ==> to_bl (v * w) = (rev (rbl_mult (rev vbl) (rev wbl)))" apply (unfold word_mult_def) apply clarify apply (simp add: to_bl_of_bin) apply (simp add: to_bl_def rbl_mult) done lemma rtb_rbl_ariths: "rev (to_bl w) = ys \ rev (to_bl (word_succ w)) = rbl_succ ys" "rev (to_bl w) = ys \ rev (to_bl (word_pred w)) = rbl_pred ys" "[| rev (to_bl v) = ys; rev (to_bl w) = xs |] ==> rev (to_bl (v * w)) = rbl_mult ys xs" "[| rev (to_bl v) = ys; rev (to_bl w) = xs |] ==> rev (to_bl (v + w)) = rbl_add ys xs" by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) subsection "Arithmetic type class instantiations" instance word :: (len0) comm_monoid_add .. instance word :: (len0) comm_monoid_mult apply (intro_classes) apply (simp add: word_mult_commute) apply (simp add: word_mult_1) done instance word :: (len0) comm_semiring by (intro_classes) (simp add : word_left_distrib) instance word :: (len0) ab_group_add .. instance word :: (len0) comm_ring .. instance word :: (len) comm_semiring_1 by (intro_classes) (simp add: lenw1_zero_neq_one) instance word :: (len) comm_ring_1 .. instance word :: (len0) comm_semiring_0 .. instance word :: (len0) order .. instance word :: (len) recpower by (intro_classes) simp_all (* note that iszero_def is only for class comm_semiring_1_cancel, which requires word length >= 1, ie 'a :: len word *) lemma zero_bintrunc: "iszero (number_of x :: 'a :: len word) = (bintrunc (len_of TYPE('a)) x = Int.Pls)" apply (unfold iszero_def word_0_wi word_no_wi) apply (rule word_ubin.norm_eq_iff [symmetric, THEN trans]) apply (simp add : Pls_def [symmetric]) done lemmas word_le_0_iff [simp] = word_zero_le [THEN leD, THEN linorder_antisym_conv1] lemma word_of_nat: "of_nat n = word_of_int (int n)" by (induct n) (auto simp add : word_of_int_hom_syms) lemma word_of_int: "of_int = word_of_int" apply (rule ext) apply (unfold of_int_def) apply (rule contentsI) apply safe apply (simp_all add: word_of_nat word_of_int_homs) defer apply (rule Rep_Integ_ne [THEN nonemptyE]) apply (rule bexI) prefer 2 apply assumption apply (auto simp add: RI_eq_diff) done lemma word_of_int_nat: "0 <= x ==> word_of_int x = of_nat (nat x)" by (simp add: of_nat_nat word_of_int) lemma word_number_of_eq: "number_of w = (of_int w :: 'a :: len word)" unfolding word_number_of_def word_of_int by auto instance word :: (len) number_ring by (intro_classes) (simp add : word_number_of_eq) lemma iszero_word_no [simp] : "iszero (number_of bin :: 'a :: len word) = iszero (number_of (bintrunc (len_of TYPE('a)) bin) :: int)" apply (simp add: zero_bintrunc number_of_is_id) apply (unfold iszero_def Pls_def) apply (rule refl) done subsection "Word and nat" lemma td_ext_unat': "n = len_of TYPE ('a :: len) ==> td_ext (unat :: 'a word => nat) of_nat (unats n) (%i. i mod 2 ^ n)" apply (unfold td_ext_def' unat_def word_of_nat unats_uints) apply (auto intro!: imageI simp add : word_of_int_hom_syms) apply (erule word_uint.Abs_inverse [THEN arg_cong]) apply (simp add: int_word_uint nat_mod_distrib nat_power_eq) done lemmas td_ext_unat = refl [THEN td_ext_unat'] lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm, standard] -interpretation word_unat!: +interpretation word_unat: td_ext "unat::'a::len word => nat" of_nat "unats (len_of TYPE('a::len))" "%i. i mod 2 ^ len_of TYPE('a::len)" by (rule td_ext_unat) lemmas td_unat = word_unat.td_thm lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] lemma unat_le: "y <= unat (z :: 'a :: len word) ==> y : unats (len_of TYPE ('a))" apply (unfold unats_def) apply clarsimp apply (rule xtrans, rule unat_lt2p, assumption) done lemma word_nchotomy: "ALL w. EX n. (w :: 'a :: len word) = of_nat n & n < 2 ^ len_of TYPE ('a)" apply (rule allI) apply (rule word_unat.Abs_cases) apply (unfold unats_def) apply auto done lemma of_nat_eq: fixes w :: "'a::len word" shows "(of_nat n = w) = (\q. n = unat w + q * 2 ^ len_of TYPE('a))" apply (rule trans) apply (rule word_unat.inverse_norm) apply (rule iffI) apply (rule mod_eqD) apply simp apply clarsimp done lemma of_nat_eq_size: "(of_nat n = w) = (EX q. n = unat w + q * 2 ^ size w)" unfolding word_size by (rule of_nat_eq) lemma of_nat_0: "(of_nat m = (0::'a::len word)) = (\q. m = q * 2 ^ len_of TYPE('a))" by (simp add: of_nat_eq) lemmas of_nat_2p = mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]] lemma of_nat_gt_0: "of_nat k ~= 0 ==> 0 < k" by (cases k) auto lemma of_nat_neq_0: "0 < k ==> k < 2 ^ len_of TYPE ('a :: len) ==> of_nat k ~= (0 :: 'a word)" by (clarsimp simp add : of_nat_0) lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" by simp lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a :: len word)" by (simp add: word_of_nat word_of_int_mult_hom zmult_int) lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" by (simp add: word_of_nat word_of_int_succ_hom add_ac) lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" by (simp add: word_of_nat word_0_wi) lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" by (simp add: word_of_nat word_1_wi) lemmas Abs_fnat_homs = Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" by simp lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" by (simp add: Abs_fnat_hom_mult [symmetric]) lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" by (subst Abs_fnat_hom_Suc [symmetric]) simp lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" by (simp add: word_div_def word_of_nat zdiv_int uint_nat) lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" by (simp add: word_mod_def word_of_nat zmod_int uint_nat) lemmas word_arith_nat_defs = word_arith_nat_add word_arith_nat_mult word_arith_nat_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 word_arith_nat_div word_arith_nat_mod lemmas unat_cong = arg_cong [where f = "unat"] lemmas unat_word_ariths = word_arith_nat_defs [THEN trans [OF unat_cong unat_of_nat], standard] lemmas word_sub_less_iff = word_sub_le_iff [simplified linorder_not_less [symmetric], simplified] lemma unat_add_lem: "(unat x + unat y < 2 ^ len_of TYPE('a)) = (unat (x + y :: 'a :: len word) = unat x + unat y)" unfolding unat_word_ariths by (auto intro!: trans [OF _ nat_mod_lem]) lemma unat_mult_lem: "(unat x * unat y < 2 ^ len_of TYPE('a)) = (unat (x * y :: 'a :: len word) = unat x * unat y)" unfolding unat_word_ariths by (auto intro!: trans [OF _ nat_mod_lem]) lemmas unat_plus_if' = trans [OF unat_word_ariths(1) mod_nat_add, simplified, standard] lemma le_no_overflow: "x <= b ==> a <= a + b ==> x <= a + (b :: 'a :: len0 word)" apply (erule order_trans) apply (erule olen_add_eqv [THEN iffD1]) done lemmas un_ui_le = trans [OF word_le_nat_alt [symmetric] word_le_def, standard] lemma unat_sub_if_size: "unat (x - y) = (if unat y <= unat x then unat x - unat y else unat x + 2 ^ size x - unat y)" apply (unfold word_size) apply (simp add: un_ui_le) apply (auto simp add: unat_def uint_sub_if') apply (rule nat_diff_distrib) prefer 3 apply (simp add: algebra_simps) apply (rule nat_diff_distrib [THEN trans]) prefer 3 apply (subst nat_add_distrib) prefer 3 apply (simp add: nat_power_eq) apply auto apply uint_arith done lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] lemma unat_div: "unat ((x :: 'a :: len word) div y) = unat x div unat y" apply (simp add : unat_word_ariths) apply (rule unat_lt2p [THEN xtr7, THEN nat_mod_eq']) apply (rule div_le_dividend) done lemma unat_mod: "unat ((x :: 'a :: len word) mod y) = unat x mod unat y" apply (clarsimp simp add : unat_word_ariths) apply (cases "unat y") prefer 2 apply (rule unat_lt2p [THEN xtr7, THEN nat_mod_eq']) apply (rule mod_le_divisor) apply auto done lemma uint_div: "uint ((x :: 'a :: len word) div y) = uint x div uint y" unfolding uint_nat by (simp add : unat_div zdiv_int) lemma uint_mod: "uint ((x :: 'a :: len word) mod y) = uint x mod uint y" unfolding uint_nat by (simp add : unat_mod zmod_int) subsection {* Definition of unat\_arith tactic *} lemma unat_split: fixes x::"'a::len word" shows "P (unat x) = (ALL n. of_nat n = x & n < 2^len_of TYPE('a) --> P n)" by (auto simp: unat_of_nat) lemma unat_split_asm: fixes x::"'a::len word" shows "P (unat x) = (~(EX n. of_nat n = x & n < 2^len_of TYPE('a) & ~ P n))" by (auto simp: unat_of_nat) lemmas of_nat_inverse = word_unat.Abs_inverse' [rotated, unfolded unats_def, simplified] lemmas unat_splits = unat_split unat_split_asm lemmas unat_arith_simps = word_le_nat_alt word_less_nat_alt word_unat.Rep_inject [symmetric] unat_sub_if' unat_plus_if' unat_div unat_mod (* unat_arith_tac: tactic to reduce word arithmetic to nat, try to solve via arith *) ML {* fun unat_arith_ss_of ss = ss addsimps @{thms unat_arith_simps} delsimps @{thms word_unat.Rep_inject} addsplits @{thms split_if_asm} addcongs @{thms power_False_cong} fun unat_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.verbose_arith_tac ctxt n t handle COOPER => Seq.empty; val cs = local_claset_of ctxt; val ss = local_simpset_of ctxt; in [ clarify_tac cs 1, full_simp_tac (unat_arith_ss_of ss) 1, ALLGOALS (full_simp_tac (HOL_ss addsplits @{thms unat_splits} addcongs @{thms power_False_cong})), rewrite_goals_tac @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN REPEAT (etac conjE n) THEN REPEAT (dtac @{thm of_nat_inverse} n THEN atac n)), TRYALL arith_tac' ] end fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) *} method_setup unat_arith = {* Scan.succeed (SIMPLE_METHOD' o unat_arith_tac) *} "solving word arithmetic via natural numbers and arith" lemma no_plus_overflow_unat_size: "((x :: 'a :: len word) <= x + y) = (unat x + unat y < 2 ^ size x)" unfolding word_size by unat_arith lemma unat_sub: "b <= a ==> unat (a - b) = unat a - unat (b :: 'a :: len word)" by unat_arith lemmas no_olen_add_nat = no_plus_overflow_unat_size [unfolded word_size] lemmas unat_plus_simple = trans [OF no_olen_add_nat unat_add_lem, standard] lemma word_div_mult: "(0 :: 'a :: len word) < y ==> unat x * unat y < 2 ^ len_of TYPE('a) ==> x * y div y = x" apply unat_arith apply clarsimp apply (subst unat_mult_lem [THEN iffD1]) apply auto done lemma div_lt': "(i :: 'a :: len word) <= k div x ==> unat i * unat x < 2 ^ len_of TYPE('a)" apply unat_arith apply clarsimp apply (drule mult_le_mono1) apply (erule order_le_less_trans) apply (rule xtr7 [OF unat_lt2p div_mult_le]) done lemmas div_lt'' = order_less_imp_le [THEN div_lt'] lemma div_lt_mult: "(i :: 'a :: len word) < k div x ==> 0 < x ==> i * x < k" apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule (1) mult_less_mono1) apply (erule order_less_le_trans) apply (rule div_mult_le) done lemma div_le_mult: "(i :: 'a :: len word) <= k div x ==> 0 < x ==> i * x <= k" apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule mult_le_mono1) apply (erule order_trans) apply (rule div_mult_le) done lemma div_lt_uint': "(i :: 'a :: len word) <= k div x ==> uint i * uint x < 2 ^ len_of TYPE('a)" apply (unfold uint_nat) apply (drule div_lt') apply (simp add: zmult_int zless_nat_eq_int_zless [symmetric] nat_power_eq) done lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] lemma word_le_exists': "(x :: 'a :: len0 word) <= y ==> (EX z. y = x + z & uint x + uint z < 2 ^ len_of TYPE('a))" apply (rule exI) apply (rule conjI) apply (rule zadd_diff_inverse) apply uint_arith done lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] lemmas plus_minus_no_overflow = order_less_imp_le [THEN plus_minus_no_overflow_ab] lemmas mcs = word_less_minus_cancel word_less_minus_mono_left word_le_minus_cancel word_le_minus_mono_left lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel, standard] lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel, standard] lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel, standard] lemmas le_unat_uoi = unat_le [THEN word_unat.Abs_inverse] lemmas thd = refl [THEN [2] split_div_lemma [THEN iffD2], THEN conjunct1] lemma thd1: "a div b * b \ (a::nat)" using gt_or_eq_0 [of b] apply (rule disjE) apply (erule xtr4 [OF thd mult_commute]) apply clarsimp done lemmas uno_simps [THEN le_unat_uoi, standard] = mod_le_divisor div_le_dividend thd1 lemma word_mod_div_equality: "(n div b) * b + (n mod b) = (n :: 'a :: len word)" apply (unfold word_less_nat_alt word_arith_nat_defs) apply (cut_tac y="unat b" in gt_or_eq_0) apply (erule disjE) apply (simp add: mod_div_equality uno_simps) apply simp done lemma word_div_mult_le: "a div b * b <= (a::'a::len word)" apply (unfold word_le_nat_alt word_arith_nat_defs) apply (cut_tac y="unat b" in gt_or_eq_0) apply (erule disjE) apply (simp add: div_mult_le uno_simps) apply simp done lemma word_mod_less_divisor: "0 < n ==> m mod n < (n :: 'a :: len word)" apply (simp only: word_less_nat_alt word_arith_nat_defs) apply (clarsimp simp add : uno_simps) done lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a :: len word)" by (induct n) (simp_all add : word_of_int_hom_syms power_Suc) lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a :: len word)" by (simp add : word_of_int_power_hom [symmetric]) lemma of_bl_length_less: "length x = k ==> k < len_of TYPE('a) ==> (of_bl x :: 'a :: len word) < 2 ^ k" apply (unfold of_bl_no [unfolded word_number_of_def] word_less_alt word_number_of_alt) apply safe apply (simp (no_asm) add: word_of_int_power_hom word_uint.eq_norm del: word_of_int_bin) apply (simp add: mod_pos_pos_trivial) apply (subst mod_pos_pos_trivial) apply (rule bl_to_bin_ge0) apply (rule order_less_trans) apply (rule bl_to_bin_lt2p) apply simp apply (rule bl_to_bin_lt2p) done subsection "Cardinality, finiteness of set of words" lemmas card_lessThan' = card_lessThan [unfolded lessThan_def] lemmas card_eq = word_unat.Abs_inj_on [THEN card_image, unfolded word_unat.image, unfolded unats_def, standard] lemmas card_word = trans [OF card_eq card_lessThan', standard] lemma finite_word_UNIV: "finite (UNIV :: 'a :: len word set)" apply (rule contrapos_np) prefer 2 apply (erule card_infinite) apply (simp add: card_word) done lemma card_word_size: "card (UNIV :: 'a :: len word set) = (2 ^ size (x :: 'a word))" unfolding word_size by (rule card_word) end diff --git a/src/HOL/Word/WordBitwise.thy b/src/HOL/Word/WordBitwise.thy --- a/src/HOL/Word/WordBitwise.thy +++ b/src/HOL/Word/WordBitwise.thy @@ -1,503 +1,503 @@ (* Author: Jeremy Dawson and Gerwin Klein, NICTA contains theorems to do with bit-wise (logical) operations on words *) header {* Bitwise Operations on Words *} theory WordBitwise imports WordArith begin lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or (* following definitions require both arithmetic and bit-wise word operations *) (* to get word_no_log_defs from word_log_defs, using bin_log_bintrs *) lemmas wils1 = bin_log_bintrs [THEN word_ubin.norm_eq_iff [THEN iffD1], folded word_ubin.eq_norm, THEN eq_reflection, standard] (* the binary operations only *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemmas word_no_log_defs [simp] = word_not_def [where a="number_of a", unfolded word_no_wi wils1, folded word_no_wi, standard] word_log_binary_defs [where a="number_of a" and b="number_of b", unfolded word_no_wi wils1, folded word_no_wi, standard] lemmas word_wi_log_defs = word_no_log_defs [unfolded word_no_wi] lemma uint_or: "uint (x OR y) = (uint x) OR (uint y)" by (simp add: word_or_def word_no_wi [symmetric] number_of_is_id bin_trunc_ao(2) [symmetric]) lemma uint_and: "uint (x AND y) = (uint x) AND (uint y)" by (simp add: word_and_def number_of_is_id word_no_wi [symmetric] bin_trunc_ao(1) [symmetric]) lemma word_ops_nth_size: "n < size (x::'a::len0 word) ==> (x OR y) !! n = (x !! n | y !! n) & (x AND y) !! n = (x !! n & y !! n) & (x XOR y) !! n = (x !! n ~= y !! n) & (NOT x) !! n = (~ x !! n)" unfolding word_size word_no_wi word_test_bit_def word_log_defs by (clarsimp simp add : word_ubin.eq_norm nth_bintr bin_nth_ops) lemma word_ao_nth: fixes x :: "'a::len0 word" shows "(x OR y) !! n = (x !! n | y !! n) & (x AND y) !! n = (x !! n & y !! n)" apply (cases "n < size x") apply (drule_tac y = "y" in word_ops_nth_size) apply simp apply (simp add : test_bit_bin word_size) done (* get from commutativity, associativity etc of int_and etc to same for word_and etc *) lemmas bwsimps = word_of_int_homs(2) word_0_wi_Pls word_m1_wi_Min word_wi_log_defs lemma word_bw_assocs: fixes x :: "'a::len0 word" shows "(x AND y) AND z = x AND y AND z" "(x OR y) OR z = x OR y OR z" "(x XOR y) XOR z = x XOR y XOR z" using word_of_int_Ex [where x=x] word_of_int_Ex [where x=y] word_of_int_Ex [where x=z] by (auto simp: bwsimps bbw_assocs) lemma word_bw_comms: fixes x :: "'a::len0 word" shows "x AND y = y AND x" "x OR y = y OR x" "x XOR y = y XOR x" using word_of_int_Ex [where x=x] word_of_int_Ex [where x=y] by (auto simp: bwsimps bin_ops_comm) lemma word_bw_lcs: fixes x :: "'a::len0 word" shows "y AND x AND z = x AND y AND z" "y OR x OR z = x OR y OR z" "y XOR x XOR z = x XOR y XOR z" using word_of_int_Ex [where x=x] word_of_int_Ex [where x=y] word_of_int_Ex [where x=z] by (auto simp: bwsimps) lemma word_log_esimps [simp]: fixes x :: "'a::len0 word" shows "x AND 0 = 0" "x AND -1 = x" "x OR 0 = x" "x OR -1 = -1" "x XOR 0 = x" "x XOR -1 = NOT x" "0 AND x = 0" "-1 AND x = x" "0 OR x = x" "-1 OR x = -1" "0 XOR x = x" "-1 XOR x = NOT x" using word_of_int_Ex [where x=x] by (auto simp: bwsimps) lemma word_not_dist: fixes x :: "'a::len0 word" shows "NOT (x OR y) = NOT x AND NOT y" "NOT (x AND y) = NOT x OR NOT y" using word_of_int_Ex [where x=x] word_of_int_Ex [where x=y] by (auto simp: bwsimps bbw_not_dist) lemma word_bw_same: fixes x :: "'a::len0 word" shows "x AND x = x" "x OR x = x" "x XOR x = 0" using word_of_int_Ex [where x=x] by (auto simp: bwsimps) lemma word_ao_absorbs [simp]: fixes x :: "'a::len0 word" shows "x AND (y OR x) = x" "x OR y AND x = x" "x AND (x OR y) = x" "y AND x OR x = x" "(y OR x) AND x = x" "x OR x AND y = x" "(x OR y) AND x = x" "x AND y OR x = x" using word_of_int_Ex [where x=x] word_of_int_Ex [where x=y] by (auto simp: bwsimps) lemma word_not_not [simp]: "NOT NOT (x::'a::len0 word) = x" using word_of_int_Ex [where x=x] by (auto simp: bwsimps) lemma word_ao_dist: fixes x :: "'a::len0 word" shows "(x OR y) AND z = x AND z OR y AND z" using word_of_int_Ex [where x=x] word_of_int_Ex [where x=y] word_of_int_Ex [where x=z] by (auto simp: bwsimps bbw_ao_dist simp del: bin_ops_comm) lemma word_oa_dist: fixes x :: "'a::len0 word" shows "x AND y OR z = (x OR z) AND (y OR z)" using word_of_int_Ex [where x=x] word_of_int_Ex [where x=y] word_of_int_Ex [where x=z] by (auto simp: bwsimps bbw_oa_dist simp del: bin_ops_comm) lemma word_add_not [simp]: fixes x :: "'a::len0 word" shows "x + NOT x = -1" using word_of_int_Ex [where x=x] by (auto simp: bwsimps bin_add_not) lemma word_plus_and_or [simp]: fixes x :: "'a::len0 word" shows "(x AND y) + (x OR y) = x + y" using word_of_int_Ex [where x=x] word_of_int_Ex [where x=y] by (auto simp: bwsimps plus_and_or) lemma leoa: fixes x :: "'a::len0 word" shows "(w = (x OR y)) ==> (y = (w AND y))" by auto lemma leao: fixes x' :: "'a::len0 word" shows "(w' = (x' AND y')) ==> (x' = (x' OR w'))" by auto lemmas word_ao_equiv = leao [COMP leoa [COMP iffI]] lemma le_word_or2: "x <= x OR (y::'a::len0 word)" unfolding word_le_def uint_or by (auto intro: le_int_or) lemmas le_word_or1 = xtr3 [OF word_bw_comms (2) le_word_or2, standard] lemmas word_and_le1 = xtr3 [OF word_ao_absorbs (4) [symmetric] le_word_or2, standard] lemmas word_and_le2 = xtr3 [OF word_ao_absorbs (8) [symmetric] le_word_or2, standard] lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" unfolding to_bl_def word_log_defs by (simp add: bl_not_bin number_of_is_id word_no_wi [symmetric]) lemma bl_word_xor: "to_bl (v XOR w) = map2 op ~= (to_bl v) (to_bl w)" unfolding to_bl_def word_log_defs bl_xor_bin by (simp add: number_of_is_id word_no_wi [symmetric]) lemma bl_word_or: "to_bl (v OR w) = map2 op | (to_bl v) (to_bl w)" unfolding to_bl_def word_log_defs by (simp add: bl_or_bin number_of_is_id word_no_wi [symmetric]) lemma bl_word_and: "to_bl (v AND w) = map2 op & (to_bl v) (to_bl w)" unfolding to_bl_def word_log_defs by (simp add: bl_and_bin number_of_is_id word_no_wi [symmetric]) lemma word_lsb_alt: "lsb (w::'a::len0 word) = test_bit w 0" by (auto simp: word_test_bit_def word_lsb_def) lemma word_lsb_1_0: "lsb (1::'a::len word) & ~ lsb (0::'b::len0 word)" unfolding word_lsb_def word_1_no word_0_no by auto lemma word_lsb_last: "lsb (w::'a::len word) = last (to_bl w)" apply (unfold word_lsb_def uint_bl bin_to_bl_def) apply (rule_tac bin="uint w" in bin_exhaust) apply (cases "size w") apply auto apply (auto simp add: bin_to_bl_aux_alt) done lemma word_lsb_int: "lsb w = (uint w mod 2 = 1)" unfolding word_lsb_def bin_last_mod by auto lemma word_msb_sint: "msb w = (sint w < 0)" unfolding word_msb_def by (simp add : sign_Min_lt_0 number_of_is_id) lemma word_msb_no': "w = number_of bin ==> msb (w::'a::len word) = bin_nth bin (size w - 1)" unfolding word_msb_def word_number_of_def by (clarsimp simp add: word_sbin.eq_norm word_size bin_sign_lem) lemmas word_msb_no = refl [THEN word_msb_no', unfolded word_size] lemma word_msb_nth': "msb (w::'a::len word) = bin_nth (uint w) (size w - 1)" apply (unfold word_size) apply (rule trans [OF _ word_msb_no]) apply (simp add : word_number_of_def) done lemmas word_msb_nth = word_msb_nth' [unfolded word_size] lemma word_msb_alt: "msb (w::'a::len word) = hd (to_bl w)" apply (unfold word_msb_nth uint_bl) apply (subst hd_conv_nth) apply (rule length_greater_0_conv [THEN iffD1]) apply simp apply (simp add : nth_bin_to_bl word_size) done lemma word_set_nth: "set_bit w n (test_bit w n) = (w::'a::len0 word)" unfolding word_test_bit_def word_set_bit_def by auto lemma bin_nth_uint': "bin_nth (uint w) n = (rev (bin_to_bl (size w) (uint w)) ! n & n < size w)" apply (unfold word_size) apply (safe elim!: bin_nth_uint_imp) apply (frule bin_nth_uint_imp) apply (fast dest!: bin_nth_bl)+ done lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] lemma test_bit_bl: "w !! n = (rev (to_bl w) ! n & n < size w)" unfolding to_bl_def word_test_bit_def word_size by (rule bin_nth_uint) lemma to_bl_nth: "n < size w ==> to_bl w ! n = w !! (size w - Suc n)" apply (unfold test_bit_bl) apply clarsimp apply (rule trans) apply (rule nth_rev_alt) apply (auto simp add: word_size) done lemma test_bit_set: fixes w :: "'a::len0 word" shows "(set_bit w n x) !! n = (n < size w & x)" unfolding word_size word_test_bit_def word_set_bit_def by (clarsimp simp add : word_ubin.eq_norm nth_bintr) lemma test_bit_set_gen: fixes w :: "'a::len0 word" shows "test_bit (set_bit w n x) m = (if m = n then n < size w & x else test_bit w m)" apply (unfold word_size word_test_bit_def word_set_bit_def) apply (clarsimp simp add: word_ubin.eq_norm nth_bintr bin_nth_sc_gen) apply (auto elim!: test_bit_size [unfolded word_size] simp add: word_test_bit_def [symmetric]) done lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" unfolding of_bl_def bl_to_bin_rep_F by auto lemma msb_nth': fixes w :: "'a::len word" shows "msb w = w !! (size w - 1)" unfolding word_msb_nth' word_test_bit_def by simp lemmas msb_nth = msb_nth' [unfolded word_size] lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size], standard] lemmas msb1 = msb0 [where i = 0] lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size], standard] lemmas word_ops_lsb = lsb0 [unfolded word_lsb_alt] lemma td_ext_nth': "n = size (w::'a::len0 word) ==> ofn = set_bits ==> [w, ofn g] = l ==> td_ext test_bit ofn {f. ALL i. f i --> i < n} (%h i. h i & i < n)" apply (unfold word_size td_ext_def') apply (safe del: subset_antisym) apply (rule_tac [3] ext) apply (rule_tac [4] ext) apply (unfold word_size of_nth_def test_bit_bl) apply safe defer apply (clarsimp simp: word_bl.Abs_inverse)+ apply (rule word_bl.Rep_inverse') apply (rule sym [THEN trans]) apply (rule bl_of_nth_nth) apply simp apply (rule bl_of_nth_inj) apply (clarsimp simp add : test_bit_bl word_size) done lemmas td_ext_nth = td_ext_nth' [OF refl refl refl, unfolded word_size] -interpretation test_bit!: +interpretation test_bit: td_ext "op !! :: 'a::len0 word => nat => bool" set_bits "{f. \i. f i \ i < len_of TYPE('a::len0)}" "(\h i. h i \ i < len_of TYPE('a::len0))" by (rule td_ext_nth) declare test_bit.Rep' [simp del] declare test_bit.Rep' [rule del] lemmas td_nth = test_bit.td_thm lemma word_set_set_same: fixes w :: "'a::len0 word" shows "set_bit (set_bit w n x) n y = set_bit w n y" by (rule word_eqI) (simp add : test_bit_set_gen word_size) lemma word_set_set_diff: fixes w :: "'a::len0 word" assumes "m ~= n" shows "set_bit (set_bit w m x) n y = set_bit (set_bit w n y) m x" by (rule word_eqI) (clarsimp simp add : test_bit_set_gen word_size prems) lemma test_bit_no': fixes w :: "'a::len0 word" shows "w = number_of bin ==> test_bit w n = (n < size w & bin_nth bin n)" unfolding word_test_bit_def word_number_of_def word_size by (simp add : nth_bintr [symmetric] word_ubin.eq_norm) lemmas test_bit_no = refl [THEN test_bit_no', unfolded word_size, THEN eq_reflection, standard] lemma nth_0: "~ (0::'a::len0 word) !! n" unfolding test_bit_no word_0_no by auto lemma nth_sint: fixes w :: "'a::len word" defines "l \ len_of TYPE ('a)" shows "bin_nth (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" unfolding sint_uint l_def by (clarsimp simp add: nth_sbintr word_test_bit_def [symmetric]) lemma word_lsb_no: "lsb (number_of bin :: 'a :: len word) = (bin_last bin = bit.B1)" unfolding word_lsb_alt test_bit_no by auto lemma word_set_no: "set_bit (number_of bin::'a::len0 word) n b = number_of (bin_sc n (if b then bit.B1 else bit.B0) bin)" apply (unfold word_set_bit_def word_number_of_def [symmetric]) apply (rule word_eqI) apply (clarsimp simp: word_size bin_nth_sc_gen number_of_is_id test_bit_no nth_bintr) done lemmas setBit_no = setBit_def [THEN trans [OF meta_eq_to_obj_eq word_set_no], simplified if_simps, THEN eq_reflection, standard] lemmas clearBit_no = clearBit_def [THEN trans [OF meta_eq_to_obj_eq word_set_no], simplified if_simps, THEN eq_reflection, standard] lemma to_bl_n1: "to_bl (-1::'a::len0 word) = replicate (len_of TYPE ('a)) True" apply (rule word_bl.Abs_inverse') apply simp apply (rule word_eqI) apply (clarsimp simp add: word_size test_bit_no) apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) done lemma word_msb_n1: "msb (-1::'a::len word)" unfolding word_msb_alt word_msb_alt to_bl_n1 by simp declare word_set_set_same [simp] word_set_nth [simp] test_bit_no [simp] word_set_no [simp] nth_0 [simp] setBit_no [simp] clearBit_no [simp] word_lsb_no [simp] word_msb_no [simp] word_msb_n1 [simp] word_lsb_1_0 [simp] lemma word_set_nth_iff: "(set_bit w n b = w) = (w !! n = b | n >= size (w::'a::len0 word))" apply (rule iffI) apply (rule disjCI) apply (drule word_eqD) apply (erule sym [THEN trans]) apply (simp add: test_bit_set) apply (erule disjE) apply clarsimp apply (rule word_eqI) apply (clarsimp simp add : test_bit_set_gen) apply (drule test_bit_size) apply force done lemma test_bit_2p': "w = word_of_int (2 ^ n) ==> w !! m = (m = n & m < size (w :: 'a :: len word))" unfolding word_test_bit_def word_size by (auto simp add: word_ubin.eq_norm nth_bintr nth_2p_bin) lemmas test_bit_2p = refl [THEN test_bit_2p', unfolded word_size] lemmas nth_w2p = test_bit_2p [unfolded of_int_number_of_eq word_of_int [symmetric] Int.of_int_power] lemma uint_2p: "(0::'a::len word) < 2 ^ n ==> uint (2 ^ n::'a::len word) = 2 ^ n" apply (unfold word_arith_power_alt) apply (case_tac "len_of TYPE ('a)") apply clarsimp apply (case_tac "nat") apply clarsimp apply (case_tac "n") apply (clarsimp simp add : word_1_wi [symmetric]) apply (clarsimp simp add : word_0_wi [symmetric]) apply (drule word_gt_0 [THEN iffD1]) apply (safe intro!: word_eqI bin_nth_lem ext) apply (auto simp add: test_bit_2p nth_2p_bin word_test_bit_def [symmetric]) done lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a :: len word) = 2 ^ n" apply (unfold word_arith_power_alt) apply (case_tac "len_of TYPE ('a)") apply clarsimp apply (case_tac "nat") apply (rule word_ubin.norm_eq_iff [THEN iffD1]) apply (rule box_equals) apply (rule_tac [2] bintr_ariths (1))+ apply (clarsimp simp add : number_of_is_id) apply simp done lemma bang_is_le: "x !! m ==> 2 ^ m <= (x :: 'a :: len word)" apply (rule xtr3) apply (rule_tac [2] y = "x" in le_word_or2) apply (rule word_eqI) apply (auto simp add: word_ao_nth nth_w2p word_size) done lemma word_clr_le: fixes w :: "'a::len0 word" shows "w >= set_bit w n False" apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm) apply simp apply (rule order_trans) apply (rule bintr_bin_clr_le) apply simp done lemma word_set_ge: fixes w :: "'a::len word" shows "w <= set_bit w n True" apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm) apply simp apply (rule order_trans [OF _ bintr_bin_set_ge]) apply simp done end diff --git a/src/HOL/Word/WordDefinition.thy b/src/HOL/Word/WordDefinition.thy --- a/src/HOL/Word/WordDefinition.thy +++ b/src/HOL/Word/WordDefinition.thy @@ -1,958 +1,958 @@ (* Author: Jeremy Dawson and Gerwin Klein, NICTA Basic definition of word type and basic theorems following from the definition of the word type *) header {* Definition of Word Type *} theory WordDefinition imports Size BinBoolList TdThs begin subsection {* Type definition *} typedef (open word) 'a word = "{(0::int) ..< 2^len_of TYPE('a::len0)}" morphisms uint Abs_word by auto definition word_of_int :: "int \ 'a\len0 word" where -- {* representation of words using unsigned or signed bins, only difference in these is the type class *} [code del]: "word_of_int w = Abs_word (bintrunc (len_of TYPE ('a)) w)" lemma uint_word_of_int [code]: "uint (word_of_int w \ 'a\len0 word) = w mod 2 ^ len_of TYPE('a)" by (auto simp add: word_of_int_def bintrunc_mod2p intro: Abs_word_inverse) code_datatype word_of_int subsection {* Type conversions and casting *} definition sint :: "'a :: len word => int" where -- {* treats the most-significant-bit as a sign bit *} sint_uint: "sint w = sbintrunc (len_of TYPE ('a) - 1) (uint w)" definition unat :: "'a :: len0 word => nat" where "unat w = nat (uint w)" definition uints :: "nat => int set" where -- "the sets of integers representing the words" "uints n = range (bintrunc n)" definition sints :: "nat => int set" where "sints n = range (sbintrunc (n - 1))" definition unats :: "nat => nat set" where "unats n = {i. i < 2 ^ n}" definition norm_sint :: "nat => int => int" where "norm_sint n w = (w + 2 ^ (n - 1)) mod 2 ^ n - 2 ^ (n - 1)" definition scast :: "'a :: len word => 'b :: len word" where -- "cast a word to a different length" "scast w = word_of_int (sint w)" definition ucast :: "'a :: len0 word => 'b :: len0 word" where "ucast w = word_of_int (uint w)" instantiation word :: (len0) size begin definition word_size: "size (w :: 'a word) = len_of TYPE('a)" instance .. end definition source_size :: "('a :: len0 word => 'b) => nat" where -- "whether a cast (or other) function is to a longer or shorter length" "source_size c = (let arb = undefined ; x = c arb in size arb)" definition target_size :: "('a => 'b :: len0 word) => nat" where "target_size c = size (c undefined)" definition is_up :: "('a :: len0 word => 'b :: len0 word) => bool" where "is_up c \ source_size c <= target_size c" definition is_down :: "('a :: len0 word => 'b :: len0 word) => bool" where "is_down c \ target_size c <= source_size c" definition of_bl :: "bool list => 'a :: len0 word" where "of_bl bl = word_of_int (bl_to_bin bl)" definition to_bl :: "'a :: len0 word => bool list" where "to_bl w = bin_to_bl (len_of TYPE ('a)) (uint w)" definition word_reverse :: "'a :: len0 word => 'a word" where "word_reverse w = of_bl (rev (to_bl w))" definition word_int_case :: "(int => 'b) => ('a :: len0 word) => 'b" where "word_int_case f w = f (uint w)" syntax of_int :: "int => 'a" translations "case x of of_int y => b" == "CONST word_int_case (%y. b) x" subsection "Arithmetic operations" instantiation word :: (len0) "{number, uminus, minus, plus, one, zero, times, Divides.div, power, ord, bit}" begin definition word_0_wi: "0 = word_of_int 0" definition word_1_wi: "1 = word_of_int 1" definition word_add_def: "a + b = word_of_int (uint a + uint b)" definition word_sub_wi: "a - b = word_of_int (uint a - uint b)" definition word_minus_def: "- a = word_of_int (- uint a)" definition word_mult_def: "a * b = word_of_int (uint a * uint b)" definition word_div_def: "a div b = word_of_int (uint a div uint b)" definition word_mod_def: "a mod b = word_of_int (uint a mod uint b)" primrec power_word where "(a\'a word) ^ 0 = 1" | "(a\'a word) ^ Suc n = a * a ^ n" definition word_number_of_def: "number_of w = word_of_int w" definition word_le_def: "a \ b \ uint a \ uint b" definition word_less_def: "x < y \ x \ y \ x \ (y \ 'a word)" definition word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" definition word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" definition word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)" definition word_not_def: "NOT (a::'a word) = word_of_int (NOT (uint a))" instance .. end definition word_succ :: "'a :: len0 word => 'a word" where "word_succ a = word_of_int (Int.succ (uint a))" definition word_pred :: "'a :: len0 word => 'a word" where "word_pred a = word_of_int (Int.pred (uint a))" constdefs udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) "a udvd b == EX n>=0. uint b = n * uint a" word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50) "a <=s b == sint a <= sint b" word_sless :: "'a :: len word => 'a word => bool" ("(_/ bin_last (uint a) = bit.B1" definition shiftl1 :: "'a word \ 'a word" where "shiftl1 w = word_of_int (uint w BIT bit.B0)" definition shiftr1 :: "'a word \ 'a word" where -- "shift right as unsigned or as signed, ie logical or arithmetic" "shiftr1 w = word_of_int (bin_rest (uint w))" definition shiftl_def: "w << n = (shiftl1 ^ n) w" definition shiftr_def: "w >> n = (shiftr1 ^ n) w" instance .. end instantiation word :: (len) bitss begin definition word_msb_def: "msb a \ bin_sign (sint a) = Int.Min" instance .. end constdefs setBit :: "'a :: len0 word => nat => 'a word" "setBit w n == set_bit w n True" clearBit :: "'a :: len0 word => nat => 'a word" "clearBit w n == set_bit w n False" subsection "Shift operations" constdefs sshiftr1 :: "'a :: len word => 'a word" "sshiftr1 w == word_of_int (bin_rest (sint w))" bshiftr1 :: "bool => 'a :: len word => 'a word" "bshiftr1 b w == of_bl (b # butlast (to_bl w))" sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) "w >>> n == (sshiftr1 ^ n) w" mask :: "nat => 'a::len word" "mask n == (1 << n) - 1" revcast :: "'a :: len0 word => 'b :: len0 word" "revcast w == of_bl (takefill False (len_of TYPE('b)) (to_bl w))" slice1 :: "nat => 'a :: len0 word => 'b :: len0 word" "slice1 n w == of_bl (takefill False n (to_bl w))" slice :: "nat => 'a :: len0 word => 'b :: len0 word" "slice n w == slice1 (size w - n) w" subsection "Rotation" constdefs rotater1 :: "'a list => 'a list" "rotater1 ys == case ys of [] => [] | x # xs => last ys # butlast ys" rotater :: "nat => 'a list => 'a list" "rotater n == rotater1 ^ n" word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" "word_rotr n w == of_bl (rotater n (to_bl w))" word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word" "word_rotl n w == of_bl (rotate n (to_bl w))" word_roti :: "int => 'a :: len0 word => 'a :: len0 word" "word_roti i w == if i >= 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w" subsection "Split and cat operations" constdefs word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word" "word_cat a b == word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))" word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)" "word_split a == case bin_split (len_of TYPE ('c)) (uint a) of (u, v) => (word_of_int u, word_of_int v)" word_rcat :: "'a :: len0 word list => 'b :: len0 word" "word_rcat ws == word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))" word_rsplit :: "'a :: len0 word => 'b :: len word list" "word_rsplit w == map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))" constdefs -- "Largest representable machine integer." max_word :: "'a::len word" "max_word \ word_of_int (2^len_of TYPE('a) - 1)" consts of_bool :: "bool \ 'a::len word" primrec "of_bool False = 0" "of_bool True = 1" lemmas of_nth_def = word_set_bits_def lemmas word_size_gt_0 [iff] = xtr1 [OF word_size len_gt_0, standard] lemmas lens_gt_0 = word_size_gt_0 len_gt_0 lemmas lens_not_0 [iff] = lens_gt_0 [THEN gr_implies_not0, standard] lemma uints_num: "uints n = {i. 0 \ i \ i < 2 ^ n}" by (simp add: uints_def range_bintrunc) lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \ i \ i < 2 ^ (n - 1)}" by (simp add: sints_def range_sbintrunc) lemmas atLeastLessThan_alt = atLeastLessThan_def [unfolded atLeast_def lessThan_def Collect_conj_eq [symmetric]] lemma mod_in_reps: "m > 0 ==> y mod m : {0::int ..< m}" unfolding atLeastLessThan_alt by auto lemma uint_0:"0 <= uint x" and uint_lt: "uint (x::'a::len0 word) < 2 ^ len_of TYPE('a)" by (auto simp: uint [simplified]) lemma uint_mod_same: "uint x mod 2 ^ len_of TYPE('a) = uint (x::'a::len0 word)" by (simp add: int_mod_eq uint_lt uint_0) lemma td_ext_uint: "td_ext (uint :: 'a word => int) word_of_int (uints (len_of TYPE('a::len0))) (%w::int. w mod 2 ^ len_of TYPE('a))" apply (unfold td_ext_def') apply (simp add: uints_num word_of_int_def bintrunc_mod2p) apply (simp add: uint_mod_same uint_0 uint_lt word.uint_inverse word.Abs_word_inverse int_mod_lem) done lemmas int_word_uint = td_ext_uint [THEN td_ext.eq_norm, standard] -interpretation word_uint!: +interpretation word_uint: td_ext "uint::'a::len0 word \ int" word_of_int "uints (len_of TYPE('a::len0))" "\w. w mod 2 ^ len_of TYPE('a::len0)" by (rule td_ext_uint) lemmas td_uint = word_uint.td_thm lemmas td_ext_ubin = td_ext_uint [simplified len_gt_0 no_bintr_alt1 [symmetric]] -interpretation word_ubin!: +interpretation word_ubin: td_ext "uint::'a::len0 word \ int" word_of_int "uints (len_of TYPE('a::len0))" "bintrunc (len_of TYPE('a::len0))" by (rule td_ext_ubin) lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = (sbintrunc (len_of TYPE ('a :: len) - 1) bin)" unfolding sint_uint by (auto simp: word_ubin.eq_norm sbintrunc_bintrunc_lt) lemma uint_sint: "uint w = bintrunc (len_of TYPE('a)) (sint (w :: 'a :: len word))" unfolding sint_uint by (auto simp: bintrunc_sbintrunc_le) lemma bintr_uint': "n >= size w ==> bintrunc n (uint w) = uint w" apply (unfold word_size) apply (subst word_ubin.norm_Rep [symmetric]) apply (simp only: bintrunc_bintrunc_min word_size min_def) apply simp done lemma wi_bintr': "wb = word_of_int bin ==> n >= size wb ==> word_of_int (bintrunc n bin) = wb" unfolding word_size by (clarsimp simp add : word_ubin.norm_eq_iff [symmetric] min_def) lemmas bintr_uint = bintr_uint' [unfolded word_size] lemmas wi_bintr = wi_bintr' [unfolded word_size] lemma td_ext_sbin: "td_ext (sint :: 'a word => int) word_of_int (sints (len_of TYPE('a::len))) (sbintrunc (len_of TYPE('a) - 1))" apply (unfold td_ext_def' sint_uint) apply (simp add : word_ubin.eq_norm) apply (cases "len_of TYPE('a)") apply (auto simp add : sints_def) apply (rule sym [THEN trans]) apply (rule word_ubin.Abs_norm) apply (simp only: bintrunc_sbintrunc) apply (drule sym) apply simp done lemmas td_ext_sint = td_ext_sbin [simplified len_gt_0 no_sbintr_alt2 Suc_pred' [symmetric]] (* We do sint before sbin, before sint is the user version and interpretations do not produce thm duplicates. I.e. we get the name word_sint.Rep_eqD, but not word_sbin.Req_eqD, because the latter is the same thm as the former *) -interpretation word_sint!: +interpretation word_sint: td_ext "sint ::'a::len word => int" word_of_int "sints (len_of TYPE('a::len))" "%w. (w + 2^(len_of TYPE('a::len) - 1)) mod 2^len_of TYPE('a::len) - 2 ^ (len_of TYPE('a::len) - 1)" by (rule td_ext_sint) -interpretation word_sbin!: +interpretation word_sbin: td_ext "sint ::'a::len word => int" word_of_int "sints (len_of TYPE('a::len))" "sbintrunc (len_of TYPE('a::len) - 1)" by (rule td_ext_sbin) lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm, standard] lemmas td_sint = word_sint.td lemma word_number_of_alt: "number_of b == word_of_int (number_of b)" unfolding word_number_of_def by (simp add: number_of_eq) lemma word_no_wi: "number_of = word_of_int" by (auto simp: word_number_of_def intro: ext) lemma to_bl_def': "(to_bl :: 'a :: len0 word => bool list) = bin_to_bl (len_of TYPE('a)) o uint" by (auto simp: to_bl_def intro: ext) lemmas word_reverse_no_def [simp] = word_reverse_def [of "number_of w", standard] lemmas uints_mod = uints_def [unfolded no_bintr_alt1] lemma uint_bintrunc: "uint (number_of bin :: 'a word) = number_of (bintrunc (len_of TYPE ('a :: len0)) bin)" unfolding word_number_of_def number_of_eq by (auto intro: word_ubin.eq_norm) lemma sint_sbintrunc: "sint (number_of bin :: 'a word) = number_of (sbintrunc (len_of TYPE ('a :: len) - 1) bin)" unfolding word_number_of_def number_of_eq by (subst word_sbin.eq_norm) simp lemma unat_bintrunc: "unat (number_of bin :: 'a :: len0 word) = number_of (bintrunc (len_of TYPE('a)) bin)" unfolding unat_def nat_number_of_def by (simp only: uint_bintrunc) (* WARNING - these may not always be helpful *) declare uint_bintrunc [simp] sint_sbintrunc [simp] unat_bintrunc [simp] lemma size_0_eq: "size (w :: 'a :: len0 word) = 0 ==> v = w" apply (unfold word_size) apply (rule word_uint.Rep_eqD) apply (rule box_equals) defer apply (rule word_ubin.norm_Rep)+ apply simp done lemmas uint_lem = word_uint.Rep [unfolded uints_num mem_Collect_eq] lemmas sint_lem = word_sint.Rep [unfolded sints_num mem_Collect_eq] lemmas uint_ge_0 [iff] = uint_lem [THEN conjunct1, standard] lemmas uint_lt2p [iff] = uint_lem [THEN conjunct2, standard] lemmas sint_ge = sint_lem [THEN conjunct1, standard] lemmas sint_lt = sint_lem [THEN conjunct2, standard] lemma sign_uint_Pls [simp]: "bin_sign (uint x) = Int.Pls" by (simp add: sign_Pls_ge_0 number_of_eq) lemmas uint_m2p_neg = iffD2 [OF diff_less_0_iff_less uint_lt2p, standard] lemmas uint_m2p_not_non_neg = iffD2 [OF linorder_not_le uint_m2p_neg, standard] lemma lt2p_lem: "len_of TYPE('a) <= n ==> uint (w :: 'a :: len0 word) < 2 ^ n" by (rule xtr8 [OF _ uint_lt2p]) simp lemmas uint_le_0_iff [simp] = uint_ge_0 [THEN leD, THEN linorder_antisym_conv1, standard] lemma uint_nat: "uint w == int (unat w)" unfolding unat_def by auto lemma uint_number_of: "uint (number_of b :: 'a :: len0 word) = number_of b mod 2 ^ len_of TYPE('a)" unfolding word_number_of_alt by (simp only: int_word_uint) lemma unat_number_of: "bin_sign b = Int.Pls ==> unat (number_of b::'a::len0 word) = number_of b mod 2 ^ len_of TYPE ('a)" apply (unfold unat_def) apply (clarsimp simp only: uint_number_of) apply (rule nat_mod_distrib [THEN trans]) apply (erule sign_Pls_ge_0 [THEN iffD1]) apply (simp_all add: nat_power_eq) done lemma sint_number_of: "sint (number_of b :: 'a :: len word) = (number_of b + 2 ^ (len_of TYPE('a) - 1)) mod 2 ^ len_of TYPE('a) - 2 ^ (len_of TYPE('a) - 1)" unfolding word_number_of_alt by (rule int_word_sint) lemma word_of_int_bin [simp] : "(word_of_int (number_of bin) :: 'a :: len0 word) = (number_of bin)" unfolding word_number_of_alt by auto lemma word_int_case_wi: "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ len_of TYPE('b::len0))" unfolding word_int_case_def by (simp add: word_uint.eq_norm) lemma word_int_split: "P (word_int_case f x) = (ALL i. x = (word_of_int i :: 'b :: len0 word) & 0 <= i & i < 2 ^ len_of TYPE('b) --> P (f i))" unfolding word_int_case_def by (auto simp: word_uint.eq_norm int_mod_eq') lemma word_int_split_asm: "P (word_int_case f x) = (~ (EX n. x = (word_of_int n :: 'b::len0 word) & 0 <= n & n < 2 ^ len_of TYPE('b::len0) & ~ P (f n)))" unfolding word_int_case_def by (auto simp: word_uint.eq_norm int_mod_eq') lemmas uint_range' = word_uint.Rep [unfolded uints_num mem_Collect_eq, standard] lemmas sint_range' = word_sint.Rep [unfolded One_nat_def sints_num mem_Collect_eq, standard] lemma uint_range_size: "0 <= uint w & uint w < 2 ^ size w" unfolding word_size by (rule uint_range') lemma sint_range_size: "- (2 ^ (size w - Suc 0)) <= sint w & sint w < 2 ^ (size w - Suc 0)" unfolding word_size by (rule sint_range') lemmas sint_above_size = sint_range_size [THEN conjunct2, THEN [2] xtr8, folded One_nat_def, standard] lemmas sint_below_size = sint_range_size [THEN conjunct1, THEN [2] order_trans, folded One_nat_def, standard] lemma test_bit_eq_iff: "(test_bit (u::'a::len0 word) = test_bit v) = (u = v)" unfolding word_test_bit_def by (simp add: bin_nth_eq_iff) lemma test_bit_size [rule_format] : "(w::'a::len0 word) !! n --> n < size w" apply (unfold word_test_bit_def) apply (subst word_ubin.norm_Rep [symmetric]) apply (simp only: nth_bintr word_size) apply fast done lemma word_eqI [rule_format] : fixes u :: "'a::len0 word" shows "(ALL n. n < size u --> u !! n = v !! n) ==> u = v" apply (rule test_bit_eq_iff [THEN iffD1]) apply (rule ext) apply (erule allE) apply (erule impCE) prefer 2 apply assumption apply (auto dest!: test_bit_size simp add: word_size) done lemmas word_eqD = test_bit_eq_iff [THEN iffD2, THEN fun_cong, standard] lemma test_bit_bin': "w !! n = (n < size w & bin_nth (uint w) n)" unfolding word_test_bit_def word_size by (simp add: nth_bintr [symmetric]) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma bin_nth_uint_imp': "bin_nth (uint w) n --> n < size w" apply (unfold word_size) apply (rule impI) apply (rule nth_bintr [THEN iffD1, THEN conjunct1]) apply (subst word_ubin.norm_Rep) apply assumption done lemma bin_nth_sint': "n >= size w --> bin_nth (sint w) n = bin_nth (sint w) (size w - 1)" apply (rule impI) apply (subst word_sbin.norm_Rep [symmetric]) apply (simp add : nth_sbintr word_size) apply auto done lemmas bin_nth_uint_imp = bin_nth_uint_imp' [rule_format, unfolded word_size] lemmas bin_nth_sint = bin_nth_sint' [rule_format, unfolded word_size] (* type definitions theorem for in terms of equivalent bool list *) lemma td_bl: "type_definition (to_bl :: 'a::len0 word => bool list) of_bl {bl. length bl = len_of TYPE('a)}" apply (unfold type_definition_def of_bl_def to_bl_def) apply (simp add: word_ubin.eq_norm) apply safe apply (drule sym) apply simp done -interpretation word_bl!: +interpretation word_bl: type_definition "to_bl :: 'a::len0 word => bool list" of_bl "{bl. length bl = len_of TYPE('a::len0)}" by (rule td_bl) lemma word_size_bl: "size w == size (to_bl w)" unfolding word_size by auto lemma to_bl_use_of_bl: "(to_bl w = bl) = (w = of_bl bl \ length bl = length (to_bl w))" by (fastsimp elim!: word_bl.Abs_inverse [simplified]) lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" unfolding word_reverse_def by (simp add: word_bl.Abs_inverse) lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" unfolding word_reverse_def by (simp add : word_bl.Abs_inverse) lemma word_rev_gal: "word_reverse w = u ==> word_reverse u = w" by auto lemmas word_rev_gal' = sym [THEN word_rev_gal, symmetric, standard] lemmas length_bl_gt_0 [iff] = xtr1 [OF word_bl.Rep' len_gt_0, standard] lemmas bl_not_Nil [iff] = length_bl_gt_0 [THEN length_greater_0_conv [THEN iffD1], standard] lemmas length_bl_neq_0 [iff] = length_bl_gt_0 [THEN gr_implies_not0] lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = Int.Min)" apply (unfold to_bl_def sint_uint) apply (rule trans [OF _ bl_sbin_sign]) apply simp done lemma of_bl_drop': "lend = length bl - len_of TYPE ('a :: len0) ==> of_bl (drop lend bl) = (of_bl bl :: 'a word)" apply (unfold of_bl_def) apply (clarsimp simp add : trunc_bl2bin [symmetric]) done lemmas of_bl_no = of_bl_def [folded word_number_of_def] lemma test_bit_of_bl: "(of_bl bl::'a::len0 word) !! n = (rev bl ! n \ n < len_of TYPE('a) \ n < length bl)" apply (unfold of_bl_def word_test_bit_def) apply (auto simp add: word_size word_ubin.eq_norm nth_bintr bin_nth_of_bl) done lemma no_of_bl: "(number_of bin ::'a::len0 word) = of_bl (bin_to_bl (len_of TYPE ('a)) bin)" unfolding word_size of_bl_no by (simp add : word_number_of_def) lemma uint_bl: "to_bl w == bin_to_bl (size w) (uint w)" unfolding word_size to_bl_def by auto lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" unfolding uint_bl by (simp add : word_size) lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len0 word) = bin_to_bl (len_of TYPE('a)) bin" unfolding uint_bl by (clarsimp simp add: word_ubin.eq_norm word_size) lemmas to_bl_no_bin [simp] = to_bl_of_bin [folded word_number_of_def] lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" unfolding uint_bl by (simp add : word_size) lemmas uint_bl_bin [simp] = trans [OF bin_bl_bin word_ubin.norm_Rep, standard] lemmas num_AB_u [simp] = word_uint.Rep_inverse [unfolded o_def word_number_of_def [symmetric], standard] lemmas num_AB_s [simp] = word_sint.Rep_inverse [unfolded o_def word_number_of_def [symmetric], standard] (* naturals *) lemma uints_unats: "uints n = int ` unats n" apply (unfold unats_def uints_num) apply safe apply (rule_tac image_eqI) apply (erule_tac nat_0_le [symmetric]) apply auto apply (erule_tac nat_less_iff [THEN iffD2]) apply (rule_tac [2] zless_nat_eq_int_zless [THEN iffD1]) apply (auto simp add : nat_power_eq int_power) done lemma unats_uints: "unats n = nat ` uints n" by (auto simp add : uints_unats image_iff) lemmas bintr_num = word_ubin.norm_eq_iff [symmetric, folded word_number_of_def, standard] lemmas sbintr_num = word_sbin.norm_eq_iff [symmetric, folded word_number_of_def, standard] lemmas num_of_bintr = word_ubin.Abs_norm [folded word_number_of_def, standard] lemmas num_of_sbintr = word_sbin.Abs_norm [folded word_number_of_def, standard]; (* don't add these to simpset, since may want bintrunc n w to be simplified; may want these in reverse, but loop as simp rules, so use following *) lemma num_of_bintr': "bintrunc (len_of TYPE('a :: len0)) a = b ==> number_of a = (number_of b :: 'a word)" apply safe apply (rule_tac num_of_bintr [symmetric]) done lemma num_of_sbintr': "sbintrunc (len_of TYPE('a :: len) - 1) a = b ==> number_of a = (number_of b :: 'a word)" apply safe apply (rule_tac num_of_sbintr [symmetric]) done lemmas num_abs_bintr = sym [THEN trans, OF num_of_bintr word_number_of_def, standard] lemmas num_abs_sbintr = sym [THEN trans, OF num_of_sbintr word_number_of_def, standard] (** cast - note, no arg for new length, as it's determined by type of result, thus in "cast w = w, the type means cast to length of w! **) lemma ucast_id: "ucast w = w" unfolding ucast_def by auto lemma scast_id: "scast w = w" unfolding scast_def by auto lemma ucast_bl: "ucast w == of_bl (to_bl w)" unfolding ucast_def of_bl_def uint_bl by (auto simp add : word_size) lemma nth_ucast: "(ucast w::'a::len0 word) !! n = (w !! n & n < len_of TYPE('a))" apply (unfold ucast_def test_bit_bin) apply (simp add: word_ubin.eq_norm nth_bintr word_size) apply (fast elim!: bin_nth_uint_imp) done (* for literal u(s)cast *) lemma ucast_bintr [simp]: "ucast (number_of w ::'a::len0 word) = number_of (bintrunc (len_of TYPE('a)) w)" unfolding ucast_def by simp lemma scast_sbintr [simp]: "scast (number_of w ::'a::len word) = number_of (sbintrunc (len_of TYPE('a) - Suc 0) w)" unfolding scast_def by simp lemmas source_size = source_size_def [unfolded Let_def word_size] lemmas target_size = target_size_def [unfolded Let_def word_size] lemmas is_down = is_down_def [unfolded source_size target_size] lemmas is_up = is_up_def [unfolded source_size target_size] lemmas is_up_down = trans [OF is_up is_down [symmetric], standard] lemma down_cast_same': "uc = ucast ==> is_down uc ==> uc = scast" apply (unfold is_down) apply safe apply (rule ext) apply (unfold ucast_def scast_def uint_sint) apply (rule word_ubin.norm_eq_iff [THEN iffD1]) apply simp done lemma word_rev_tf': "r = to_bl (of_bl bl) ==> r = rev (takefill False (length r) (rev bl))" unfolding of_bl_def uint_bl by (clarsimp simp add: bl_bin_bl_rtf word_ubin.eq_norm word_size) lemmas word_rev_tf = refl [THEN word_rev_tf', unfolded word_bl.Rep', standard] lemmas word_rep_drop = word_rev_tf [simplified takefill_alt, simplified, simplified rev_take, simplified] lemma to_bl_ucast: "to_bl (ucast (w::'b::len0 word) ::'a::len0 word) = replicate (len_of TYPE('a) - len_of TYPE('b)) False @ drop (len_of TYPE('b) - len_of TYPE('a)) (to_bl w)" apply (unfold ucast_bl) apply (rule trans) apply (rule word_rep_drop) apply simp done lemma ucast_up_app': "uc = ucast ==> source_size uc + n = target_size uc ==> to_bl (uc w) = replicate n False @ (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma ucast_down_drop': "uc = ucast ==> source_size uc = target_size uc + n ==> to_bl (uc w) = drop n (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma scast_down_drop': "sc = scast ==> source_size sc = target_size sc + n ==> to_bl (sc w) = drop n (to_bl w)" apply (subgoal_tac "sc = ucast") apply safe apply simp apply (erule refl [THEN ucast_down_drop']) apply (rule refl [THEN down_cast_same', symmetric]) apply (simp add : source_size target_size is_down) done lemma sint_up_scast': "sc = scast ==> is_up sc ==> sint (sc w) = sint w" apply (unfold is_up) apply safe apply (simp add: scast_def word_sbin.eq_norm) apply (rule box_equals) prefer 3 apply (rule word_sbin.norm_Rep) apply (rule sbintrunc_sbintrunc_l) defer apply (subst word_sbin.norm_Rep) apply (rule refl) apply simp done lemma uint_up_ucast': "uc = ucast ==> is_up uc ==> uint (uc w) = uint w" apply (unfold is_up) apply safe apply (rule bin_eqI) apply (fold word_test_bit_def) apply (auto simp add: nth_ucast) apply (auto simp add: test_bit_bin) done lemmas down_cast_same = refl [THEN down_cast_same'] lemmas ucast_up_app = refl [THEN ucast_up_app'] lemmas ucast_down_drop = refl [THEN ucast_down_drop'] lemmas scast_down_drop = refl [THEN scast_down_drop'] lemmas uint_up_ucast = refl [THEN uint_up_ucast'] lemmas sint_up_scast = refl [THEN sint_up_scast'] lemma ucast_up_ucast': "uc = ucast ==> is_up uc ==> ucast (uc w) = ucast w" apply (simp (no_asm) add: ucast_def) apply (clarsimp simp add: uint_up_ucast) done lemma scast_up_scast': "sc = scast ==> is_up sc ==> scast (sc w) = scast w" apply (simp (no_asm) add: scast_def) apply (clarsimp simp add: sint_up_scast) done lemma ucast_of_bl_up': "w = of_bl bl ==> size bl <= size w ==> ucast w = of_bl bl" by (auto simp add : nth_ucast word_size test_bit_of_bl intro!: word_eqI) lemmas ucast_up_ucast = refl [THEN ucast_up_ucast'] lemmas scast_up_scast = refl [THEN scast_up_scast'] lemmas ucast_of_bl_up = refl [THEN ucast_of_bl_up'] lemmas ucast_up_ucast_id = trans [OF ucast_up_ucast ucast_id] lemmas scast_up_scast_id = trans [OF scast_up_scast scast_id] lemmas isduu = is_up_down [where c = "ucast", THEN iffD2] lemmas isdus = is_up_down [where c = "scast", THEN iffD2] lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] lemmas scast_down_scast_id = isdus [THEN ucast_up_ucast_id] lemma up_ucast_surj: "is_up (ucast :: 'b::len0 word => 'a::len0 word) ==> surj (ucast :: 'a word => 'b word)" by (rule surjI, erule ucast_up_ucast_id) lemma up_scast_surj: "is_up (scast :: 'b::len word => 'a::len word) ==> surj (scast :: 'a word => 'b word)" by (rule surjI, erule scast_up_scast_id) lemma down_scast_inj: "is_down (scast :: 'b::len word => 'a::len word) ==> inj_on (ucast :: 'a word => 'b word) A" by (rule inj_on_inverseI, erule scast_down_scast_id) lemma down_ucast_inj: "is_down (ucast :: 'b::len0 word => 'a::len0 word) ==> inj_on (ucast :: 'a word => 'b word) A" by (rule inj_on_inverseI, erule ucast_down_ucast_id) lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" by (rule word_bl.Rep_eqD) (simp add: word_rep_drop) lemma ucast_down_no': "uc = ucast ==> is_down uc ==> uc (number_of bin) = number_of bin" apply (unfold word_number_of_def is_down) apply (clarsimp simp add: ucast_def word_ubin.eq_norm) apply (rule word_ubin.norm_eq_iff [THEN iffD1]) apply (erule bintrunc_bintrunc_ge) done lemmas ucast_down_no = ucast_down_no' [OF refl] lemma ucast_down_bl': "uc = ucast ==> is_down uc ==> uc (of_bl bl) = of_bl bl" unfolding of_bl_no by clarify (erule ucast_down_no) lemmas ucast_down_bl = ucast_down_bl' [OF refl] lemmas slice_def' = slice_def [unfolded word_size] lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def lemmas word_log_bin_defs = word_log_defs text {* Executable equality *} instantiation word :: ("{len0}") eq begin definition eq_word :: "'a word \ 'a word \ bool" where "eq_word k l \ HOL.eq (uint k) (uint l)" instance proof qed (simp add: eq eq_word_def) end end diff --git a/src/HOL/Word/WordGenLib.thy b/src/HOL/Word/WordGenLib.thy --- a/src/HOL/Word/WordGenLib.thy +++ b/src/HOL/Word/WordGenLib.thy @@ -1,463 +1,463 @@ (* Author: Gerwin Klein, Jeremy Dawson Miscellaneous additional library definitions and lemmas for the word type. Instantiation to boolean algebras, definition of recursion and induction patterns for words. *) header {* Miscellaneous Library for Words *} theory WordGenLib imports WordShift Boolean_Algebra begin declare of_nat_2p [simp] lemma word_int_cases: "\\n. \(x ::'a::len0 word) = word_of_int n; 0 \ n; n < 2^len_of TYPE('a)\ \ P\ \ P" by (cases x rule: word_uint.Abs_cases) (simp add: uints_num) lemma word_nat_cases [cases type: word]: "\\n. \(x ::'a::len word) = of_nat n; n < 2^len_of TYPE('a)\ \ P\ \ P" by (cases x rule: word_unat.Abs_cases) (simp add: unats_def) lemma max_word_eq: "(max_word::'a::len word) = 2^len_of TYPE('a) - 1" by (simp add: max_word_def word_of_int_hom_syms word_of_int_2p) lemma max_word_max [simp,intro!]: "n \ max_word" by (cases n rule: word_int_cases) (simp add: max_word_def word_le_def int_word_uint int_mod_eq') lemma word_of_int_2p_len: "word_of_int (2 ^ len_of TYPE('a)) = (0::'a::len0 word)" by (subst word_uint.Abs_norm [symmetric]) (simp add: word_of_int_hom_syms) lemma word_pow_0: "(2::'a::len word) ^ len_of TYPE('a) = 0" proof - have "word_of_int (2 ^ len_of TYPE('a)) = (0::'a word)" by (rule word_of_int_2p_len) thus ?thesis by (simp add: word_of_int_2p) qed lemma max_word_wrap: "x + 1 = 0 \ x = max_word" apply (simp add: max_word_eq) apply uint_arith apply auto apply (simp add: word_pow_0) done lemma max_word_minus: "max_word = (-1::'a::len word)" proof - have "-1 + 1 = (0::'a word)" by simp thus ?thesis by (rule max_word_wrap [symmetric]) qed lemma max_word_bl [simp]: "to_bl (max_word::'a::len word) = replicate (len_of TYPE('a)) True" by (subst max_word_minus to_bl_n1)+ simp lemma max_test_bit [simp]: "(max_word::'a::len word) !! n = (n < len_of TYPE('a))" by (auto simp add: test_bit_bl word_size) lemma word_and_max [simp]: "x AND max_word = x" by (rule word_eqI) (simp add: word_ops_nth_size word_size) lemma word_or_max [simp]: "x OR max_word = max_word" by (rule word_eqI) (simp add: word_ops_nth_size word_size) lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND (z::'a::len0 word)" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR (z::'a::len0 word))" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_and_not [simp]: "x AND NOT x = (0::'a::len0 word)" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_or_not [simp]: "x OR NOT x = max_word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_boolean: "boolean (op AND) (op OR) bitNOT 0 max_word" apply (rule boolean.intro) apply (rule word_bw_assocs) apply (rule word_bw_assocs) apply (rule word_bw_comms) apply (rule word_bw_comms) apply (rule word_ao_dist2) apply (rule word_oa_dist2) apply (rule word_and_max) apply (rule word_log_esimps) apply (rule word_and_not) apply (rule word_or_not) done -interpretation word_bool_alg!: +interpretation word_bool_alg: boolean "op AND" "op OR" bitNOT 0 max_word by (rule word_boolean) lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND (y::'a::len0 word)" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) -interpretation word_bool_alg!: +interpretation word_bool_alg: boolean_xor "op AND" "op OR" bitNOT 0 max_word "op XOR" apply (rule boolean_xor.intro) apply (rule word_boolean) apply (rule boolean_xor_axioms.intro) apply (rule word_xor_and_or) done lemma shiftr_0 [iff]: "(x::'a::len0 word) >> 0 = x" by (simp add: shiftr_bl) lemma shiftl_0 [simp]: "(x :: 'a :: len word) << 0 = x" by (simp add: shiftl_t2n) lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" by (simp add: shiftl_t2n) lemma uint_lt_0 [simp]: "uint x < 0 = False" by (simp add: linorder_not_less) lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" by (simp add: shiftr1_def word_0_alt) lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (induct n) (auto simp: shiftr_def) lemma word_less_1 [simp]: "((x::'a::len word) < 1) = (x = 0)" by (simp add: word_less_nat_alt unat_0_iff) lemma to_bl_mask: "to_bl (mask n :: 'a::len word) = replicate (len_of TYPE('a) - n) False @ replicate (min (len_of TYPE('a)) n) True" by (simp add: mask_bl word_rep_drop min_def) lemma map_replicate_True: "n = length xs ==> map (\(x,y). x & y) (zip xs (replicate n True)) = xs" by (induct xs arbitrary: n) auto lemma map_replicate_False: "n = length xs ==> map (\(x,y). x & y) (zip xs (replicate n False)) = replicate n False" by (induct xs arbitrary: n) auto lemma bl_and_mask: fixes w :: "'a::len word" fixes n defines "n' \ len_of TYPE('a) - n" shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" proof - note [simp] = map_replicate_True map_replicate_False have "to_bl (w AND mask n) = map2 op & (to_bl w) (to_bl (mask n::'a::len word))" by (simp add: bl_word_and) also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" by simp also have "map2 op & \ (to_bl (mask n::'a::len word)) = replicate n' False @ drop n' (to_bl w)" unfolding to_bl_mask n'_def map2_def by (subst zip_append) auto finally show ?thesis . qed lemma drop_rev_takefill: "length xs \ n ==> drop (n - length xs) (rev (takefill False n (rev xs))) = xs" by (simp add: takefill_alt rev_take) lemma map_nth_0 [simp]: "map (op !! (0::'a::len0 word)) xs = replicate (length xs) False" by (induct xs) auto lemma uint_plus_if_size: "uint (x + y) = (if uint x + uint y < 2^size x then uint x + uint y else uint x + uint y - 2^size x)" by (simp add: word_arith_alts int_word_uint mod_add_if_z word_size) lemma unat_plus_if_size: "unat (x + (y::'a::len word)) = (if unat x + unat y < 2^size x then unat x + unat y else unat x + unat y - 2^size x)" apply (subst word_arith_nat_defs) apply (subst unat_of_nat) apply (simp add: mod_nat_add word_size) done lemma word_neq_0_conv [simp]: fixes w :: "'a :: len word" shows "(w \ 0) = (0 < w)" proof - have "0 \ w" by (rule word_zero_le) thus ?thesis by (auto simp add: word_less_def) qed lemma max_lt: "unat (max a b div c) = unat (max a b) div unat (c:: 'a :: len word)" apply (subst word_arith_nat_defs) apply (subst word_unat.eq_norm) apply (subst mod_if) apply clarsimp apply (erule notE) apply (insert div_le_dividend [of "unat (max a b)" "unat c"]) apply (erule order_le_less_trans) apply (insert unat_lt2p [of "max a b"]) apply simp done lemma uint_sub_if_size: "uint (x - y) = (if uint y \ uint x then uint x - uint y else uint x - uint y + 2^size x)" by (simp add: word_arith_alts int_word_uint mod_sub_if_z word_size) lemma unat_sub_simple: "x \ y ==> unat (y - x) = unat y - unat x" by (simp add: unat_def uint_sub_if_size word_le_def nat_diff_distrib) lemmas unat_sub = unat_sub_simple lemma word_less_sub1: fixes x :: "'a :: len word" shows "x \ 0 ==> 1 < x = (0 < x - 1)" by (simp add: unat_sub_if_size word_less_nat_alt) lemma word_le_sub1: fixes x :: "'a :: len word" shows "x \ 0 ==> 1 \ x = (0 \ x - 1)" by (simp add: unat_sub_if_size order_le_less word_less_nat_alt) lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "number_of w", standard] lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "number_of w", standard] lemma word_of_int_minus: "word_of_int (2^len_of TYPE('a) - i) = (word_of_int (-i)::'a::len word)" proof - have x: "2^len_of TYPE('a) - i = -i + 2^len_of TYPE('a)" by simp show ?thesis apply (subst x) apply (subst word_uint.Abs_norm [symmetric], subst mod_add_self2) apply simp done qed lemmas word_of_int_inj = word_uint.Abs_inject [unfolded uints_num, simplified] lemma word_le_less_eq: "(x ::'z::len word) \ y = (x = y \ x < y)" by (auto simp add: word_less_def) lemma mod_plus_cong: assumes 1: "(b::int) = b'" and 2: "x mod b' = x' mod b'" and 3: "y mod b' = y' mod b'" and 4: "x' + y' = z'" shows "(x + y) mod b = z' mod b'" proof - from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" by (simp add: mod_add_eq[symmetric]) also have "\ = (x' + y') mod b'" by (simp add: mod_add_eq[symmetric]) finally show ?thesis by (simp add: 4) qed lemma mod_minus_cong: assumes 1: "(b::int) = b'" and 2: "x mod b' = x' mod b'" and 3: "y mod b' = y' mod b'" and 4: "x' - y' = z'" shows "(x - y) mod b = z' mod b'" using assms apply (subst zmod_zsub_left_eq) apply (subst zmod_zsub_right_eq) apply (simp add: zmod_zsub_left_eq [symmetric] zmod_zsub_right_eq [symmetric]) done lemma word_induct_less: "\P (0::'a::len word); \n. \n < m; P n\ \ P (1 + n)\ \ P m" apply (cases m) apply atomize apply (erule rev_mp)+ apply (rule_tac x=m in spec) apply (induct_tac n) apply simp apply clarsimp apply (erule impE) apply clarsimp apply (erule_tac x=n in allE) apply (erule impE) apply (simp add: unat_arith_simps) apply (clarsimp simp: unat_of_nat) apply simp apply (erule_tac x="of_nat na" in allE) apply (erule impE) apply (simp add: unat_arith_simps) apply (clarsimp simp: unat_of_nat) apply simp done lemma word_induct: "\P (0::'a::len word); \n. P n \ P (1 + n)\ \ P m" by (erule word_induct_less, simp) lemma word_induct2 [induct type]: "\P 0; \n. \1 + n \ 0; P n\ \ P (1 + n)\ \ P (n::'b::len word)" apply (rule word_induct, simp) apply (case_tac "1+n = 0", auto) done constdefs word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" "word_rec forZero forSuc n \ nat_rec forZero (forSuc \ of_nat) (unat n)" lemma word_rec_0: "word_rec z s 0 = z" by (simp add: word_rec_def) lemma word_rec_Suc: "1 + n \ (0::'a::len word) \ word_rec z s (1 + n) = s n (word_rec z s n)" apply (simp add: word_rec_def unat_word_ariths) apply (subst nat_mod_eq') apply (cut_tac x=n in unat_lt2p) apply (drule Suc_mono) apply (simp add: less_Suc_eq_le) apply (simp only: order_less_le, simp) apply (erule contrapos_pn, simp) apply (drule arg_cong[where f=of_nat]) apply simp apply (subst (asm) word_unat.Rep_inverse[of n]) apply simp apply simp done lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" apply (rule subst[where t="n" and s="1 + (n - 1)"]) apply simp apply (subst word_rec_Suc) apply simp apply simp done lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \ op + 1) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_twice: "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ op + (n - m)) m" apply (erule rev_mp) apply (rule_tac x=z in spec) apply (rule_tac x=f in spec) apply (induct n) apply (simp add: word_rec_0) apply clarsimp apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) apply simp apply (case_tac "1 + (n - m) = 0") apply (simp add: word_rec_0) apply (rule_tac f = "word_rec ?a ?b" in arg_cong) apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) apply simp apply (simp (no_asm_use)) apply (simp add: word_rec_Suc word_rec_in2) apply (erule impE) apply uint_arith apply (drule_tac x="x \ op + 1" in spec) apply (drule_tac x="x 0 xa" in spec) apply simp apply (rule_tac t="\a. x (1 + (n - m + a))" and s="\a. x (1 + (n - m) + a)" in subst) apply (clarsimp simp add: expand_fun_eq) apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) apply simp apply (rule refl) apply (rule refl) done lemma word_rec_id: "word_rec z (\_. id) n = z" by (induct n) (auto simp add: word_rec_0 word_rec_Suc) lemma word_rec_id_eq: "\m < n. f m = id \ word_rec z f n = z" apply (erule rev_mp) apply (induct n) apply (auto simp add: word_rec_0 word_rec_Suc) apply (drule spec, erule mp) apply uint_arith apply (drule_tac x=n in spec, erule impE) apply uint_arith apply simp done lemma word_rec_max: "\m\n. m \ -1 \ f m = id \ word_rec z f -1 = word_rec z f n" apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) apply simp apply simp apply (rule word_rec_id_eq) apply clarsimp apply (drule spec, rule mp, erule mp) apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) prefer 2 apply assumption apply simp apply (erule contrapos_pn) apply simp apply (drule arg_cong[where f="\x. x - n"]) apply simp done lemma unatSuc: "1 + n \ (0::'a::len word) \ unat (1 + n) = Suc (unat n)" by unat_arith lemmas word_no_1 [simp] = word_1_no [symmetric, unfolded BIT_simps] lemmas word_no_0 [simp] = word_0_no [symmetric] declare word_0_bl [simp] declare bin_to_bl_def [simp] declare to_bl_0 [simp] declare of_bl_True [simp] end diff --git a/src/HOL/ex/LocaleTest2.thy b/src/HOL/ex/LocaleTest2.thy --- a/src/HOL/ex/LocaleTest2.thy +++ b/src/HOL/ex/LocaleTest2.thy @@ -1,915 +1,915 @@ (* Title: HOL/ex/LocaleTest2.thy Author: Clemens Ballarin Copyright (c) 2007 by Clemens Ballarin More regression tests for locales. Definitions are less natural in FOL, since there is no selection operator. Hence we do them here in HOL, not in the main test suite for locales, which is FOL/ex/LocaleTest.thy *) header {* Test of Locale Interpretation *} theory LocaleTest2 imports Main GCD begin section {* Interpretation of Defined Concepts *} text {* Naming convention for global objects: prefixes D and d *} subsection {* Lattices *} text {* Much of the lattice proofs are from HOL/Lattice. *} subsubsection {* Definitions *} locale dpo = 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" begin theorem circular: "[| x \ y; y \ z; z \ x |] ==> x = y & y = z" by (blast intro: trans) definition less :: "['a, 'a] => bool" (infixl "\" 50) where "(x \ y) = (x \ y & x ~= y)" theorem abs_test: "op \ = (%x y. x \ y)" by simp definition is_inf :: "['a, 'a, 'a] => bool" where "is_inf x y i = (i \ x \ i \ y \ (\z. z \ x \ z \ y \ z \ i))" definition is_sup :: "['a, 'a, 'a] => bool" where "is_sup x y s = (x \ s \ y \ s \ (\z. x \ z \ y \ z \ s \ z))" end locale dlat = dpo + assumes ex_inf: "EX inf. dpo.is_inf le x y inf" and ex_sup: "EX sup. dpo.is_sup le x y sup" begin definition meet :: "['a, 'a] => 'a" (infixl "\" 70) where "x \ y = (THE inf. is_inf x y inf)" definition join :: "['a, 'a] => 'a" (infixl "\" 65) where "x \ y = (THE sup. is_sup x y sup)" lemma 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 is_inf_lower [elim?]: "is_inf x y i \ (i \ x \ i \ y \ C) \ C" by (unfold is_inf_def) blast lemma 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 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 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 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 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 meet_right [intro?]: "x \ y \ y" by (rule is_inf_lower) (rule is_inf_meet) lemma meet_le [intro?]: "[| z \ x; z \ y |] ==> z \ x \ y" by (rule is_inf_greatest) (rule is_inf_meet) lemma 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 is_sup_least [elim?]: "is_sup x y s \ x \ z \ y \ z \ s \ z" by (unfold is_sup_def) blast lemma 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 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 lemma 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 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 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 join_left [intro?]: "x \ x \ y" by (rule is_sup_upper) (rule is_sup_join) lemma join_right [intro?]: "y \ x \ y" by (rule is_sup_upper) (rule is_sup_join) lemma join_le [intro?]: "[| x \ z; y \ z |] ==> x \ y \ z" by (rule is_sup_least) (rule is_sup_join) theorem 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 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 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 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 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 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 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 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 meet_related2 [elim?]: "y \ x \ x \ y = y" by (drule meet_related) (simp add: meet_commute) theorem 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 join_related2 [elim?]: "y \ x \ x \ y = x" by (drule join_related) (simp add: join_commute) text {* Additional theorems *} theorem 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 meet_connection2: "(x \ y) = (y \ x = x)" using meet_commute meet_connection by simp theorem 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 join_connection2: "(x \ y) = (x \ y = y)" using join_commute join_connection by simp text {* Naming according to Jacobson I, p.\ 459. *} lemmas L1 = join_commute meet_commute lemmas L2 = join_assoc meet_assoc (*lemmas L3 = join_idem meet_idem*) lemmas L4 = join_meet_absorb meet_join_absorb end locale ddlat = dlat + assumes meet_distr: "dlat.meet le x (dlat.join le y z) = dlat.join le (dlat.meet le x y) (dlat.meet le x z)" begin lemma 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 end locale dlo = dpo + assumes total: "x \ y | y \ x" begin lemma less_total: "x \ y | x = y | y \ x" using total by (unfold less_def) blast end sublocale dlo < dlat proof 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 "EX inf. is_inf x y inf" by blast next 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 "EX sup. is_sup x y sup" by blast qed sublocale dlo < ddlat proof fix 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_commute*) join_connection2 join_related2 (*meet_commute*) meet_connection meet_related2 total) also from c have "... = ?r" by (metis (*c*) (*join_commute*) meet_related2) finally have "?l = ?r" . } moreover { assume c: "x \ y | x \ z" from c have "?l = x" by (metis (*anti_sym*) (*c*) (*circular*) (*join_assoc*)(* join_commute *) join_connection2 (*join_left*) join_related2 meet_connection(* meet_related2*) 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 subsubsection {* Total order @{text "<="} on @{typ int} *} -interpretation int!: dpo "op <= :: [int, int] => bool" +interpretation int: dpo "op <= :: [int, int] => bool" where "(dpo.less (op <=) (x::int) y) = (x < y)" txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *} proof - show "dpo (op <= :: [int, int] => bool)" proof qed auto then interpret int: dpo "op <= :: [int, int] => bool" . txt {* Gives interpreted version of @{text less_def} (without condition). *} show "(dpo.less (op <=) (x::int) y) = (x < y)" by (unfold int.less_def) auto qed thm int.circular lemma "\ (x::int) \ y; y \ z; z \ x\ \ x = y \ y = z" apply (rule int.circular) apply assumption apply assumption apply assumption done thm int.abs_test lemma "(op < :: [int, int] => bool) = op <" apply (rule int.abs_test) done -interpretation int!: dlat "op <= :: [int, int] => bool" +interpretation int: dlat "op <= :: [int, int] => bool" where meet_eq: "dlat.meet (op <=) (x::int) y = min x y" and join_eq: "dlat.join (op <=) (x::int) y = max x y" proof - show "dlat (op <= :: [int, int] => bool)" apply unfold_locales apply (unfold int.is_inf_def int.is_sup_def) apply arith+ done then interpret int: dlat "op <= :: [int, int] => bool" . txt {* Interpretation to ease use of definitions, which are conditional in general but unconditional after interpretation. *} show "dlat.meet (op <=) (x::int) y = min x y" apply (unfold int.meet_def) apply (rule the_equality) apply (unfold int.is_inf_def) by auto show "dlat.join (op <=) (x::int) y = max x y" apply (unfold int.join_def) apply (rule the_equality) apply (unfold int.is_sup_def) by auto qed -interpretation int!: dlo "op <= :: [int, int] => bool" +interpretation int: dlo "op <= :: [int, int] => bool" proof qed arith text {* Interpreted theorems from the locales, involving defined terms. *} thm int.less_def text {* from dpo *} thm int.meet_left text {* from dlat *} thm int.meet_distr text {* from ddlat *} thm int.less_total text {* from dlo *} subsubsection {* Total order @{text "<="} on @{typ nat} *} -interpretation nat!: dpo "op <= :: [nat, nat] => bool" +interpretation nat: dpo "op <= :: [nat, nat] => bool" where "dpo.less (op <=) (x::nat) y = (x < y)" txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *} proof - show "dpo (op <= :: [nat, nat] => bool)" proof qed auto then interpret nat: dpo "op <= :: [nat, nat] => bool" . txt {* Gives interpreted version of @{text less_def} (without condition). *} show "dpo.less (op <=) (x::nat) y = (x < y)" apply (unfold nat.less_def) apply auto done qed -interpretation nat!: dlat "op <= :: [nat, nat] => bool" +interpretation nat: dlat "op <= :: [nat, nat] => bool" where "dlat.meet (op <=) (x::nat) y = min x y" and "dlat.join (op <=) (x::nat) y = max x y" proof - show "dlat (op <= :: [nat, nat] => bool)" apply unfold_locales apply (unfold nat.is_inf_def nat.is_sup_def) apply arith+ done then interpret nat: dlat "op <= :: [nat, nat] => bool" . txt {* Interpretation to ease use of definitions, which are conditional in general but unconditional after interpretation. *} show "dlat.meet (op <=) (x::nat) y = min x y" apply (unfold nat.meet_def) apply (rule the_equality) apply (unfold nat.is_inf_def) by auto show "dlat.join (op <=) (x::nat) y = max x y" apply (unfold nat.join_def) apply (rule the_equality) apply (unfold nat.is_sup_def) by auto qed -interpretation nat!: dlo "op <= :: [nat, nat] => bool" +interpretation nat: dlo "op <= :: [nat, nat] => bool" proof qed arith text {* Interpreted theorems from the locales, involving defined terms. *} thm nat.less_def text {* from dpo *} thm nat.meet_left text {* from dlat *} thm nat.meet_distr text {* from ddlat *} thm nat.less_total text {* from ldo *} subsubsection {* Lattice @{text "dvd"} on @{typ nat} *} -interpretation nat_dvd!: dpo "op dvd :: [nat, nat] => bool" +interpretation nat_dvd: dpo "op dvd :: [nat, nat] => bool" where "dpo.less (op dvd) (x::nat) y = (x dvd y & x ~= y)" txt {* We give interpretation for less, but not @{text is_inf} and @{text is_sub}. *} proof - show "dpo (op dvd :: [nat, nat] => bool)" proof qed (auto simp: dvd_def) then interpret nat_dvd: dpo "op dvd :: [nat, nat] => bool" . txt {* Gives interpreted version of @{text less_def} (without condition). *} show "dpo.less (op dvd) (x::nat) y = (x dvd y & x ~= y)" apply (unfold nat_dvd.less_def) apply auto done qed -interpretation nat_dvd!: dlat "op dvd :: [nat, nat] => bool" +interpretation nat_dvd: dlat "op dvd :: [nat, nat] => bool" where "dlat.meet (op dvd) (x::nat) y = gcd x y" and "dlat.join (op dvd) (x::nat) y = lcm x y" proof - show "dlat (op dvd :: [nat, nat] => bool)" apply unfold_locales apply (unfold nat_dvd.is_inf_def nat_dvd.is_sup_def) apply (rule_tac x = "gcd x y" in exI) apply auto [1] apply (rule_tac x = "lcm x y" in exI) apply (auto intro: lcm_dvd1 lcm_dvd2 lcm_least) done then interpret nat_dvd: dlat "op dvd :: [nat, nat] => bool" . txt {* Interpretation to ease use of definitions, which are conditional in general but unconditional after interpretation. *} show "dlat.meet (op dvd) (x::nat) y = gcd x y" apply (unfold nat_dvd.meet_def) apply (rule the_equality) apply (unfold nat_dvd.is_inf_def) by auto show "dlat.join (op dvd) (x::nat) y = lcm x y" apply (unfold nat_dvd.join_def) apply (rule the_equality) apply (unfold nat_dvd.is_sup_def) by (auto intro: lcm_dvd1 lcm_dvd2 lcm_least) qed text {* Interpreted theorems from the locales, involving defined terms. *} thm nat_dvd.less_def text {* from dpo *} lemma "((x::nat) dvd y & x ~= y) = (x dvd y & x ~= y)" apply (rule nat_dvd.less_def) done thm nat_dvd.meet_left text {* from dlat *} lemma "gcd x y dvd x" apply (rule nat_dvd.meet_left) done subsection {* Group example with defined operations @{text inv} and @{text unit} *} subsubsection {* Locale declarations and lemmas *} locale Dsemi = fixes prod (infixl "**" 65) assumes assoc: "(x ** y) ** z = x ** (y ** z)" locale Dmonoid = Dsemi + fixes one assumes l_one [simp]: "one ** x = x" and r_one [simp]: "x ** one = x" begin definition inv where "inv x = (THE y. x ** y = one & y ** x = one)" definition unit where "unit x = (EX y. x ** y = one & y ** x = one)" lemma inv_unique: assumes eq: "y ** x = one" "x ** y' = one" shows "y = y'" proof - from eq have "y = y ** (x ** y')" by (simp add: r_one) also have "... = (y ** x) ** y'" by (simp add: assoc) also from eq have "... = y'" by (simp add: l_one) finally show ?thesis . qed lemma unit_one [intro, simp]: "unit one" by (unfold unit_def) auto lemma unit_l_inv_ex: "unit x ==> \y. y ** x = one" by (unfold unit_def) auto lemma unit_r_inv_ex: "unit x ==> \y. x ** y = one" by (unfold unit_def) auto lemma unit_l_inv: "unit x ==> inv x ** x = one" apply (simp add: unit_def inv_def) apply (erule exE) apply (rule theI2, fast) apply (rule inv_unique) apply fast+ done lemma unit_r_inv: "unit x ==> x ** inv x = one" apply (simp add: unit_def inv_def) apply (erule exE) apply (rule theI2, fast) apply (rule inv_unique) apply fast+ done lemma unit_inv_unit [intro, simp]: "unit x ==> unit (inv x)" proof - assume x: "unit x" show "unit (inv x)" by (auto simp add: unit_def intro: unit_l_inv unit_r_inv x) qed lemma unit_l_cancel [simp]: "unit x ==> (x ** y = x ** z) = (y = z)" proof assume eq: "x ** y = x ** z" and G: "unit x" then have "(inv x ** x) ** y = (inv x ** x) ** z" by (simp add: assoc) with G show "y = z" by (simp add: unit_l_inv) next assume eq: "y = z" and G: "unit x" then show "x ** y = x ** z" by simp qed lemma unit_inv_inv [simp]: "unit x ==> inv (inv x) = x" proof - assume x: "unit x" then have "inv x ** inv (inv x) = inv x ** x" by (simp add: unit_l_inv unit_r_inv) with x show ?thesis by simp qed lemma inv_inj_on_unit: "inj_on inv {x. unit x}" proof (rule inj_onI, simp) fix x y assume G: "unit x" "unit y" 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 unit_inv_comm: assumes inv: "x ** y = one" and G: "unit x" "unit y" shows "y ** x = one" proof - from G have "x ** y ** x = x ** one" by (auto simp add: inv) with G show ?thesis by (simp del: r_one add: assoc) qed end locale Dgrp = Dmonoid + assumes unit [intro, simp]: "Dmonoid.unit (op **) one x" begin lemma l_inv_ex [simp]: "\y. y ** x = one" using unit_l_inv_ex by simp lemma r_inv_ex [simp]: "\y. x ** y = one" using unit_r_inv_ex by simp lemma l_inv [simp]: "inv x ** x = one" using unit_l_inv by simp lemma l_cancel [simp]: "(x ** y = x ** z) = (y = z)" using unit_l_inv by simp lemma r_inv [simp]: "x ** inv x = one" proof - have "inv x ** (x ** inv x) = inv x ** one" by (simp add: assoc [symmetric] l_inv) then show ?thesis by (simp del: r_one) qed lemma r_cancel [simp]: "(y ** x = z ** x) = (y = z)" proof assume eq: "y ** x = z ** x" then have "y ** (x ** inv x) = z ** (x ** inv x)" by (simp add: assoc [symmetric] del: r_inv) then show "y = z" by simp qed simp lemma inv_one [simp]: "inv one = one" proof - have "inv one = one ** (inv one)" by (simp del: r_inv) moreover have "... = one" by simp finally show ?thesis . qed lemma inv_inv [simp]: "inv (inv x) = x" using unit_inv_inv by simp lemma inv_inj: "inj_on inv UNIV" using inv_inj_on_unit by simp lemma inv_mult_group: "inv (x ** y) = inv y ** inv x" proof - have "inv (x ** y) ** (x ** y) = (inv y ** inv x) ** (x ** y)" by (simp add: assoc l_inv) (simp add: assoc [symmetric]) then show ?thesis by (simp del: l_inv) qed lemma inv_comm: "x ** y = one ==> y ** x = one" by (rule unit_inv_comm) auto lemma inv_equality: "y ** x = one ==> inv x = y" apply (simp add: inv_def) apply (rule the_equality) apply (simp add: inv_comm [of y x]) apply (rule r_cancel [THEN iffD1], auto) done end locale Dhom = prod: Dgrp prod one + sum: Dgrp sum zero for prod (infixl "**" 65) and one and sum (infixl "+++" 60) and zero + fixes hom assumes hom_mult [simp]: "hom (x ** y) = hom x +++ hom y" begin lemma hom_one [simp]: "hom one = zero" proof - have "hom one +++ zero = hom one +++ hom one" by (simp add: hom_mult [symmetric] del: hom_mult) then show ?thesis by (simp del: r_one) qed end subsubsection {* Interpretation of Functions *} -interpretation Dfun!: Dmonoid "op o" "id :: 'a => 'a" +interpretation Dfun: Dmonoid "op o" "id :: 'a => 'a" where "Dmonoid.unit (op o) id f = bij (f::'a => 'a)" (* and "Dmonoid.inv (op o) id" = "inv :: ('a => 'a) => ('a => 'a)" *) proof - show "Dmonoid op o (id :: 'a => 'a)" proof qed (simp_all add: o_assoc) note Dmonoid = this (* from this interpret Dmonoid "op o" "id :: 'a => 'a" . *) show "Dmonoid.unit (op o) (id :: 'a => 'a) f = bij f" apply (unfold Dmonoid.unit_def [OF Dmonoid]) apply rule apply clarify proof - fix f g assume id1: "f o g = id" and id2: "g o f = id" show "bij f" proof (rule bijI) show "inj f" proof (rule inj_onI) fix x y assume "f x = f y" then have "(g o f) x = (g o f) y" by simp with id2 show "x = y" by simp qed next show "surj f" proof (rule surjI) fix x from id1 have "(f o g) x = x" by simp then show "f (g x) = x" by simp qed qed next fix f assume bij: "bij f" then have inv: "f o Hilbert_Choice.inv f = id & Hilbert_Choice.inv f o f = id" by (simp add: bij_def surj_iff inj_iff) show "EX g. f o g = id & g o f = id" by rule (rule inv) qed qed thm Dmonoid.unit_def Dfun.unit_def thm Dmonoid.inv_inj_on_unit Dfun.inv_inj_on_unit lemma unit_id: "(f :: unit => unit) = id" by rule simp -interpretation Dfun!: Dgrp "op o" "id :: unit => unit" +interpretation Dfun: Dgrp "op o" "id :: unit => unit" where "Dmonoid.inv (op o) id f = inv (f :: unit => unit)" proof - have "Dmonoid op o (id :: 'a => 'a)" .. note Dmonoid = this show "Dgrp (op o) (id :: unit => unit)" apply unfold_locales apply (unfold Dmonoid.unit_def [OF Dmonoid]) apply (insert unit_id) apply simp done show "Dmonoid.inv (op o) id f = inv (f :: unit => unit)" apply (unfold Dmonoid.inv_def [OF Dmonoid] inv_def) apply (insert unit_id) apply simp apply (rule the_equality) apply rule apply rule apply simp done qed thm Dfun.unit_l_inv Dfun.l_inv thm Dfun.inv_equality thm Dfun.inv_equality end diff --git a/src/HOLCF/Algebraic.thy b/src/HOLCF/Algebraic.thy --- a/src/HOLCF/Algebraic.thy +++ b/src/HOLCF/Algebraic.thy @@ -1,537 +1,537 @@ (* Title: HOLCF/Algebraic.thy Author: Brian Huffman *) header {* Algebraic deflations *} theory Algebraic imports Completion Fix Eventual begin subsection {* Constructing finite deflations by iteration *} lemma finite_deflation_imp_deflation: "finite_deflation d \ deflation d" unfolding finite_deflation_def by simp lemma le_Suc_induct: assumes le: "i \ j" assumes step: "\i. P i (Suc i)" assumes refl: "\i. P i i" assumes trans: "\i j k. \P i j; P j k\ \ P i k" shows "P i j" proof (cases "i = j") assume "i = j" thus "P i j" by (simp add: refl) next assume "i \ j" with le have "i < j" by simp thus "P i j" using step trans by (rule less_Suc_induct) qed text {* A pre-deflation is like a deflation, but not idempotent. *} locale pre_deflation = fixes f :: "'a \ 'a::cpo" assumes less: "\x. f\x \ x" assumes finite_range: "finite (range (\x. f\x))" begin lemma iterate_less: "iterate i\f\x \ x" by (induct i, simp_all add: trans_less [OF less]) lemma iterate_fixed: "f\x = x \ iterate i\f\x = x" by (induct i, simp_all) lemma antichain_iterate_app: "i \ j \ iterate j\f\x \ iterate i\f\x" apply (erule le_Suc_induct) apply (simp add: less) apply (rule refl_less) apply (erule (1) trans_less) done lemma finite_range_iterate_app: "finite (range (\i. iterate i\f\x))" proof (rule finite_subset) show "range (\i. iterate i\f\x) \ insert x (range (\x. f\x))" by (clarify, case_tac i, simp_all) show "finite (insert x (range (\x. f\x)))" by (simp add: finite_range) qed lemma eventually_constant_iterate_app: "eventually_constant (\i. iterate i\f\x)" unfolding eventually_constant_def MOST_nat_le proof - let ?Y = "\i. iterate i\f\x" have "\j. \k. ?Y j \ ?Y k" apply (rule finite_range_has_max) apply (erule antichain_iterate_app) apply (rule finite_range_iterate_app) done then obtain j where j: "\k. ?Y j \ ?Y k" by fast show "\z m. \n\m. ?Y n = z" proof (intro exI allI impI) fix k assume "j \ k" hence "?Y k \ ?Y j" by (rule antichain_iterate_app) also have "?Y j \ ?Y k" by (rule j) finally show "?Y k = ?Y j" . qed qed lemma eventually_constant_iterate: "eventually_constant (\n. iterate n\f)" proof - have "\y\range (\x. f\x). eventually_constant (\i. iterate i\f\y)" by (simp add: eventually_constant_iterate_app) hence "\y\range (\x. f\x). MOST i. MOST j. iterate j\f\y = iterate i\f\y" unfolding eventually_constant_MOST_MOST . hence "MOST i. MOST j. \y\range (\x. f\x). iterate j\f\y = iterate i\f\y" by (simp only: MOST_finite_Ball_distrib [OF finite_range]) hence "MOST i. MOST j. \x. iterate j\f\(f\x) = iterate i\f\(f\x)" by simp hence "MOST i. MOST j. \x. iterate (Suc j)\f\x = iterate (Suc i)\f\x" by (simp only: iterate_Suc2) hence "MOST i. MOST j. iterate (Suc j)\f = iterate (Suc i)\f" by (simp only: expand_cfun_eq) hence "eventually_constant (\i. iterate (Suc i)\f)" unfolding eventually_constant_MOST_MOST . thus "eventually_constant (\i. iterate i\f)" by (rule eventually_constant_SucD) qed abbreviation d :: "'a \ 'a" where "d \ eventual (\n. iterate n\f)" lemma MOST_d: "MOST n. P (iterate n\f) \ P d" using eventually_constant_iterate by (rule MOST_eventual) lemma f_d: "f\(d\x) = d\x" apply (rule MOST_d) apply (subst iterate_Suc [symmetric]) apply (rule eventually_constant_MOST_Suc_eq) apply (rule eventually_constant_iterate_app) done lemma d_fixed_iff: "d\x = x \ f\x = x" proof assume "d\x = x" with f_d [where x=x] show "f\x = x" by simp next assume f: "f\x = x" have "\n. iterate n\f\x = x" by (rule allI, rule nat.induct, simp, simp add: f) hence "MOST n. iterate n\f\x = x" by (rule ALL_MOST) thus "d\x = x" by (rule MOST_d) qed lemma finite_deflation_d: "finite_deflation d" proof fix x :: 'a have "d \ range (\n. iterate n\f)" using eventually_constant_iterate by (rule eventual_mem_range) then obtain n where n: "d = iterate n\f" .. have "iterate n\f\(d\x) = d\x" using f_d by (rule iterate_fixed) thus "d\(d\x) = d\x" by (simp add: n) next fix x :: 'a show "d\x \ x" by (rule MOST_d, simp add: iterate_less) next from finite_range have "finite {x. f\x = x}" by (rule finite_range_imp_finite_fixes) thus "finite {x. d\x = x}" by (simp add: d_fixed_iff) qed end lemma pre_deflation_d_f: assumes "finite_deflation d" assumes f: "\x. f\x \ x" shows "pre_deflation (d oo f)" proof interpret d: finite_deflation d by fact fix x show "\x. (d oo f)\x \ x" by (simp, rule trans_less [OF d.less f]) show "finite (range (\x. (d oo f)\x))" by (rule finite_subset [OF _ d.finite_range], auto) qed lemma eventual_iterate_oo_fixed_iff: assumes "finite_deflation d" assumes f: "\x. f\x \ x" shows "eventual (\n. iterate n\(d oo f))\x = x \ d\x = x \ f\x = x" proof - interpret d: finite_deflation d by fact let ?e = "d oo f" interpret e: pre_deflation "d oo f" using `finite_deflation d` f by (rule pre_deflation_d_f) let ?g = "eventual (\n. iterate n\?e)" show ?thesis apply (subst e.d_fixed_iff) apply simp apply safe apply (erule subst) apply (rule d.idem) apply (rule antisym_less) apply (rule f) apply (erule subst, rule d.less) apply simp done qed subsection {* Type constructor for finite deflations *} defaultsort profinite typedef (open) 'a fin_defl = "{d::'a \ 'a. finite_deflation d}" by (fast intro: finite_deflation_approx) instantiation fin_defl :: (profinite) sq_ord begin definition sq_le_fin_defl_def: "op \ \ \x y. Rep_fin_defl x \ Rep_fin_defl y" instance .. end instance fin_defl :: (profinite) po by (rule typedef_po [OF type_definition_fin_defl sq_le_fin_defl_def]) lemma finite_deflation_Rep_fin_defl: "finite_deflation (Rep_fin_defl d)" using Rep_fin_defl by simp -interpretation Rep_fin_defl!: finite_deflation "Rep_fin_defl d" +interpretation Rep_fin_defl: finite_deflation "Rep_fin_defl d" by (rule finite_deflation_Rep_fin_defl) lemma fin_defl_lessI: "(\x. Rep_fin_defl a\x = x \ Rep_fin_defl b\x = x) \ a \ b" unfolding sq_le_fin_defl_def by (rule Rep_fin_defl.lessI) lemma fin_defl_lessD: "\a \ b; Rep_fin_defl a\x = x\ \ Rep_fin_defl b\x = x" unfolding sq_le_fin_defl_def by (rule Rep_fin_defl.lessD) lemma fin_defl_eqI: "(\x. Rep_fin_defl a\x = x \ Rep_fin_defl b\x = x) \ a = b" apply (rule antisym_less) apply (rule fin_defl_lessI, simp) apply (rule fin_defl_lessI, simp) done lemma Abs_fin_defl_mono: "\finite_deflation a; finite_deflation b; a \ b\ \ Abs_fin_defl a \ Abs_fin_defl b" unfolding sq_le_fin_defl_def by (simp add: Abs_fin_defl_inverse) subsection {* Take function for finite deflations *} definition fd_take :: "nat \ 'a fin_defl \ 'a fin_defl" where "fd_take i d = Abs_fin_defl (eventual (\n. iterate n\(approx i oo Rep_fin_defl d)))" lemma Rep_fin_defl_fd_take: "Rep_fin_defl (fd_take i d) = eventual (\n. iterate n\(approx i oo Rep_fin_defl d))" unfolding fd_take_def apply (rule Abs_fin_defl_inverse [unfolded mem_Collect_eq]) apply (rule pre_deflation.finite_deflation_d) apply (rule pre_deflation_d_f) apply (rule finite_deflation_approx) apply (rule Rep_fin_defl.less) done lemma fd_take_fixed_iff: "Rep_fin_defl (fd_take i d)\x = x \ approx i\x = x \ Rep_fin_defl d\x = x" unfolding Rep_fin_defl_fd_take by (rule eventual_iterate_oo_fixed_iff [OF finite_deflation_approx Rep_fin_defl.less]) lemma fd_take_less: "fd_take n d \ d" apply (rule fin_defl_lessI) apply (simp add: fd_take_fixed_iff) done lemma fd_take_idem: "fd_take n (fd_take n d) = fd_take n d" apply (rule fin_defl_eqI) apply (simp add: fd_take_fixed_iff) done lemma fd_take_mono: "a \ b \ fd_take n a \ fd_take n b" apply (rule fin_defl_lessI) apply (simp add: fd_take_fixed_iff) apply (simp add: fin_defl_lessD) done lemma approx_fixed_le_lemma: "\i \ j; approx i\x = x\ \ approx j\x = x" by (erule subst, simp add: min_def) lemma fd_take_chain: "m \ n \ fd_take m a \ fd_take n a" apply (rule fin_defl_lessI) apply (simp add: fd_take_fixed_iff) apply (simp add: approx_fixed_le_lemma) done lemma finite_range_fd_take: "finite (range (fd_take n))" apply (rule finite_imageD [where f="\a. {x. Rep_fin_defl a\x = x}"]) apply (rule finite_subset [where B="Pow {x. approx n\x = x}"]) apply (clarify, simp add: fd_take_fixed_iff) apply (simp add: finite_fixes_approx) apply (rule inj_onI, clarify) apply (simp add: expand_set_eq fin_defl_eqI) done lemma fd_take_covers: "\n. fd_take n a = a" apply (rule_tac x= "Max ((\x. LEAST n. approx n\x = x) ` {x. Rep_fin_defl a\x = x})" in exI) apply (rule antisym_less) apply (rule fd_take_less) apply (rule fin_defl_lessI) apply (simp add: fd_take_fixed_iff) apply (rule approx_fixed_le_lemma) apply (rule Max_ge) apply (rule finite_imageI) apply (rule Rep_fin_defl.finite_fixes) apply (rule imageI) apply (erule CollectI) apply (rule LeastI_ex) apply (rule profinite_compact_eq_approx) apply (erule subst) apply (rule Rep_fin_defl.compact) done -interpretation fin_defl!: basis_take sq_le fd_take +interpretation fin_defl: basis_take sq_le fd_take apply default apply (rule fd_take_less) apply (rule fd_take_idem) apply (erule fd_take_mono) apply (rule fd_take_chain, simp) apply (rule finite_range_fd_take) apply (rule fd_take_covers) done subsection {* Defining algebraic deflations by ideal completion *} typedef (open) 'a alg_defl = "{S::'a fin_defl set. sq_le.ideal S}" by (fast intro: sq_le.ideal_principal) instantiation alg_defl :: (profinite) sq_ord begin definition "x \ y \ Rep_alg_defl x \ Rep_alg_defl y" instance .. end instance alg_defl :: (profinite) po by (rule sq_le.typedef_ideal_po [OF type_definition_alg_defl sq_le_alg_defl_def]) instance alg_defl :: (profinite) cpo by (rule sq_le.typedef_ideal_cpo [OF type_definition_alg_defl sq_le_alg_defl_def]) lemma Rep_alg_defl_lub: "chain Y \ Rep_alg_defl (\i. Y i) = (\i. Rep_alg_defl (Y i))" by (rule sq_le.typedef_ideal_rep_contlub [OF type_definition_alg_defl sq_le_alg_defl_def]) lemma ideal_Rep_alg_defl: "sq_le.ideal (Rep_alg_defl xs)" by (rule Rep_alg_defl [unfolded mem_Collect_eq]) definition alg_defl_principal :: "'a fin_defl \ 'a alg_defl" where "alg_defl_principal t = Abs_alg_defl {u. u \ t}" lemma Rep_alg_defl_principal: "Rep_alg_defl (alg_defl_principal t) = {u. u \ t}" unfolding alg_defl_principal_def by (simp add: Abs_alg_defl_inverse sq_le.ideal_principal) -interpretation alg_defl!: +interpretation alg_defl: ideal_completion sq_le fd_take alg_defl_principal Rep_alg_defl apply default apply (rule ideal_Rep_alg_defl) apply (erule Rep_alg_defl_lub) apply (rule Rep_alg_defl_principal) apply (simp only: sq_le_alg_defl_def) done text {* Algebraic deflations are pointed *} lemma finite_deflation_UU: "finite_deflation \" by default simp_all lemma alg_defl_minimal: "alg_defl_principal (Abs_fin_defl \) \ x" apply (induct x rule: alg_defl.principal_induct, simp) apply (rule alg_defl.principal_mono) apply (induct_tac a) apply (rule Abs_fin_defl_mono) apply (rule finite_deflation_UU) apply simp apply (rule minimal) done instance alg_defl :: (bifinite) pcpo by intro_classes (fast intro: alg_defl_minimal) lemma inst_alg_defl_pcpo: "\ = alg_defl_principal (Abs_fin_defl \)" by (rule alg_defl_minimal [THEN UU_I, symmetric]) text {* Algebraic deflations are profinite *} instantiation alg_defl :: (profinite) profinite begin definition approx_alg_defl_def: "approx = alg_defl.completion_approx" instance apply (intro_classes, unfold approx_alg_defl_def) apply (rule alg_defl.chain_completion_approx) apply (rule alg_defl.lub_completion_approx) apply (rule alg_defl.completion_approx_idem) apply (rule alg_defl.finite_fixes_completion_approx) done end instance alg_defl :: (bifinite) bifinite .. lemma approx_alg_defl_principal [simp]: "approx n\(alg_defl_principal t) = alg_defl_principal (fd_take n t)" unfolding approx_alg_defl_def by (rule alg_defl.completion_approx_principal) lemma approx_eq_alg_defl_principal: "\t\Rep_alg_defl xs. approx n\xs = alg_defl_principal (fd_take n t)" unfolding approx_alg_defl_def by (rule alg_defl.completion_approx_eq_principal) subsection {* Applying algebraic deflations *} definition cast :: "'a alg_defl \ 'a \ 'a" where "cast = alg_defl.basis_fun Rep_fin_defl" lemma cast_alg_defl_principal: "cast\(alg_defl_principal a) = Rep_fin_defl a" unfolding cast_def apply (rule alg_defl.basis_fun_principal) apply (simp only: sq_le_fin_defl_def) done lemma deflation_cast: "deflation (cast\d)" apply (induct d rule: alg_defl.principal_induct) apply (rule adm_subst [OF _ adm_deflation], simp) apply (simp add: cast_alg_defl_principal) apply (rule finite_deflation_imp_deflation) apply (rule finite_deflation_Rep_fin_defl) done lemma finite_deflation_cast: "compact d \ finite_deflation (cast\d)" apply (drule alg_defl.compact_imp_principal, clarify) apply (simp add: cast_alg_defl_principal) apply (rule finite_deflation_Rep_fin_defl) done -interpretation cast!: deflation "cast\d" +interpretation cast: deflation "cast\d" by (rule deflation_cast) lemma "cast\(\i. alg_defl_principal (Abs_fin_defl (approx i)))\x = x" apply (subst contlub_cfun_arg) apply (rule chainI) apply (rule alg_defl.principal_mono) apply (rule Abs_fin_defl_mono) apply (rule finite_deflation_approx) apply (rule finite_deflation_approx) apply (rule chainE) apply (rule chain_approx) apply (simp add: cast_alg_defl_principal Abs_fin_defl_inverse finite_deflation_approx) done text {* This lemma says that if we have an ep-pair from a bifinite domain into a universal domain, then e oo p is an algebraic deflation. *} lemma assumes "ep_pair e p" constrains e :: "'a::profinite \ 'b::profinite" shows "\d. cast\d = e oo p" proof interpret ep_pair e p by fact let ?a = "\i. e oo approx i oo p" have a: "\i. finite_deflation (?a i)" apply (rule finite_deflation_e_d_p) apply (rule finite_deflation_approx) done let ?d = "\i. alg_defl_principal (Abs_fin_defl (?a i))" show "cast\?d = e oo p" apply (subst contlub_cfun_arg) apply (rule chainI) apply (rule alg_defl.principal_mono) apply (rule Abs_fin_defl_mono [OF a a]) apply (rule chainE, simp) apply (subst cast_alg_defl_principal) apply (simp add: Abs_fin_defl_inverse a) apply (simp add: expand_cfun_eq lub_distribs) done qed text {* This lemma says that if we have an ep-pair from a cpo into a bifinite domain, and e oo p is an algebraic deflation, then the cpo is bifinite. *} lemma assumes "ep_pair e p" constrains e :: "'a::cpo \ 'b::profinite" assumes d: "\x. cast\d\x = e\(p\x)" obtains a :: "nat \ 'a \ 'a" where "\i. finite_deflation (a i)" "(\i. a i) = ID" proof interpret ep_pair e p by fact let ?a = "\i. p oo cast\(approx i\d) oo e" show "\i. finite_deflation (?a i)" apply (rule finite_deflation_p_d_e) apply (rule finite_deflation_cast) apply (rule compact_approx) apply (rule sq_ord_less_eq_trans [OF _ d]) apply (rule monofun_cfun_fun) apply (rule monofun_cfun_arg) apply (rule approx_less) done show "(\i. ?a i) = ID" apply (rule ext_cfun, simp) apply (simp add: lub_distribs) apply (simp add: d) done qed end diff --git a/src/HOLCF/Bifinite.thy b/src/HOLCF/Bifinite.thy --- a/src/HOLCF/Bifinite.thy +++ b/src/HOLCF/Bifinite.thy @@ -1,173 +1,173 @@ (* Title: HOLCF/Bifinite.thy Author: Brian Huffman *) header {* Bifinite domains and approximation *} theory Bifinite imports Deflation begin subsection {* Omega-profinite and bifinite domains *} class profinite = fixes approx :: "nat \ 'a \ 'a" assumes chain_approx [simp]: "chain approx" assumes lub_approx_app [simp]: "(\i. approx i\x) = x" assumes approx_idem: "approx i\(approx i\x) = approx i\x" assumes finite_fixes_approx: "finite {x. approx i\x = x}" class bifinite = profinite + pcpo lemma approx_less: "approx i\x \ x" proof - have "chain (\i. approx i\x)" by simp hence "approx i\x \ (\i. approx i\x)" by (rule is_ub_thelub) thus "approx i\x \ x" by simp qed lemma finite_deflation_approx: "finite_deflation (approx i)" proof fix x :: 'a show "approx i\(approx i\x) = approx i\x" by (rule approx_idem) show "approx i\x \ x" by (rule approx_less) show "finite {x. approx i\x = x}" by (rule finite_fixes_approx) qed -interpretation approx!: finite_deflation "approx i" +interpretation approx: finite_deflation "approx i" by (rule finite_deflation_approx) lemma (in deflation) deflation: "deflation d" .. lemma deflation_approx: "deflation (approx i)" by (rule approx.deflation) lemma lub_approx [simp]: "(\i. approx i) = (\ x. x)" by (rule ext_cfun, simp add: contlub_cfun_fun) lemma approx_strict [simp]: "approx i\\ = \" by (rule UU_I, rule approx_less) lemma approx_approx1: "i \ j \ approx i\(approx j\x) = approx i\x" apply (rule deflation_less_comp1 [OF deflation_approx deflation_approx]) apply (erule chain_mono [OF chain_approx]) done lemma approx_approx2: "j \ i \ approx i\(approx j\x) = approx j\x" apply (rule deflation_less_comp2 [OF deflation_approx deflation_approx]) apply (erule chain_mono [OF chain_approx]) done lemma approx_approx [simp]: "approx i\(approx j\x) = approx (min i j)\x" apply (rule_tac x=i and y=j in linorder_le_cases) apply (simp add: approx_approx1 min_def) apply (simp add: approx_approx2 min_def) done lemma finite_image_approx: "finite ((\x. approx n\x) ` A)" by (rule approx.finite_image) lemma finite_range_approx: "finite (range (\x. approx i\x))" by (rule approx.finite_range) lemma compact_approx [simp]: "compact (approx n\x)" by (rule approx.compact) lemma profinite_compact_eq_approx: "compact x \ \i. approx i\x = x" by (rule admD2, simp_all) lemma profinite_compact_iff: "compact x \ (\n. approx n\x = x)" apply (rule iffI) apply (erule profinite_compact_eq_approx) apply (erule exE) apply (erule subst) apply (rule compact_approx) done lemma approx_induct: assumes adm: "adm P" and P: "\n x. P (approx n\x)" shows "P x" proof - have "P (\n. approx n\x)" by (rule admD [OF adm], simp, simp add: P) thus "P x" by simp qed lemma profinite_less_ext: "(\i. approx i\x \ approx i\y) \ x \ y" apply (subgoal_tac "(\i. approx i\x) \ (\i. approx i\y)", simp) apply (rule lub_mono, simp, simp, simp) done subsection {* Instance for continuous function space *} lemma finite_range_cfun_lemma: assumes a: "finite (range (\x. a\x))" assumes b: "finite (range (\y. b\y))" shows "finite (range (\f. \ x. b\(f\(a\x))))" (is "finite (range ?h)") proof (rule finite_imageD) let ?f = "\g. range (\x. (a\x, g\x))" show "finite (?f ` range ?h)" proof (rule finite_subset) let ?B = "Pow (range (\x. a\x) \ range (\y. b\y))" show "?f ` range ?h \ ?B" by clarsimp show "finite ?B" by (simp add: a b) qed show "inj_on ?f (range ?h)" proof (rule inj_onI, rule ext_cfun, clarsimp) fix x f g assume "range (\x. (a\x, b\(f\(a\x)))) = range (\x. (a\x, b\(g\(a\x))))" hence "range (\x. (a\x, b\(f\(a\x)))) \ range (\x. (a\x, b\(g\(a\x))))" by (rule equalityD1) hence "(a\x, b\(f\(a\x))) \ range (\x. (a\x, b\(g\(a\x))))" by (simp add: subset_eq) then obtain y where "(a\x, b\(f\(a\x))) = (a\y, b\(g\(a\y)))" by (rule rangeE) thus "b\(f\(a\x)) = b\(g\(a\x))" by clarsimp qed qed instantiation "->" :: (profinite, profinite) profinite begin definition approx_cfun_def: "approx = (\n. \ f x. approx n\(f\(approx n\x)))" instance proof show "chain (approx :: nat \ ('a \ 'b) \ ('a \ 'b))" unfolding approx_cfun_def by simp next fix x :: "'a \ 'b" show "(\i. approx i\x) = x" unfolding approx_cfun_def by (simp add: lub_distribs eta_cfun) next fix i :: nat and x :: "'a \ 'b" show "approx i\(approx i\x) = approx i\x" unfolding approx_cfun_def by simp next fix i :: nat show "finite {x::'a \ 'b. approx i\x = x}" apply (rule finite_range_imp_finite_fixes) apply (simp add: approx_cfun_def) apply (intro finite_range_cfun_lemma finite_range_approx) done qed end instance "->" :: (profinite, bifinite) bifinite .. lemma approx_cfun: "approx n\f\x = approx n\(f\(approx n\x))" by (simp add: approx_cfun_def) end diff --git a/src/HOLCF/CompactBasis.thy b/src/HOLCF/CompactBasis.thy --- a/src/HOLCF/CompactBasis.thy +++ b/src/HOLCF/CompactBasis.thy @@ -1,299 +1,299 @@ (* Title: HOLCF/CompactBasis.thy Author: Brian Huffman *) header {* Compact bases of domains *} theory CompactBasis imports Completion begin subsection {* Compact bases of bifinite domains *} defaultsort profinite typedef (open) 'a compact_basis = "{x::'a::profinite. compact x}" by (fast intro: compact_approx) lemma compact_Rep_compact_basis: "compact (Rep_compact_basis a)" by (rule Rep_compact_basis [unfolded mem_Collect_eq]) instantiation compact_basis :: (profinite) sq_ord begin definition compact_le_def: "(op \) \ (\x y. Rep_compact_basis x \ Rep_compact_basis y)" instance .. end instance compact_basis :: (profinite) po by (rule typedef_po [OF type_definition_compact_basis compact_le_def]) text {* Take function for compact basis *} definition compact_take :: "nat \ 'a compact_basis \ 'a compact_basis" where "compact_take = (\n a. Abs_compact_basis (approx n\(Rep_compact_basis a)))" lemma Rep_compact_take: "Rep_compact_basis (compact_take n a) = approx n\(Rep_compact_basis a)" unfolding compact_take_def by (simp add: Abs_compact_basis_inverse) lemmas approx_Rep_compact_basis = Rep_compact_take [symmetric] -interpretation compact_basis!: +interpretation compact_basis: basis_take sq_le compact_take proof fix n :: nat and a :: "'a compact_basis" show "compact_take n a \ a" unfolding compact_le_def by (simp add: Rep_compact_take approx_less) next fix n :: nat and a :: "'a compact_basis" show "compact_take n (compact_take n a) = compact_take n a" by (simp add: Rep_compact_basis_inject [symmetric] Rep_compact_take) next fix n :: nat and a b :: "'a compact_basis" assume "a \ b" thus "compact_take n a \ compact_take n b" unfolding compact_le_def Rep_compact_take by (rule monofun_cfun_arg) next fix n :: nat and a :: "'a compact_basis" show "\n a. compact_take n a \ compact_take (Suc n) a" unfolding compact_le_def Rep_compact_take by (rule chainE, simp) next fix n :: nat show "finite (range (compact_take n))" apply (rule finite_imageD [where f="Rep_compact_basis"]) apply (rule finite_subset [where B="range (\x. approx n\x)"]) apply (clarsimp simp add: Rep_compact_take) apply (rule finite_range_approx) apply (rule inj_onI, simp add: Rep_compact_basis_inject) done next fix a :: "'a compact_basis" show "\n. compact_take n a = a" apply (simp add: Rep_compact_basis_inject [symmetric]) apply (simp add: Rep_compact_take) apply (rule profinite_compact_eq_approx) apply (rule compact_Rep_compact_basis) done qed text {* Ideal completion *} definition approximants :: "'a \ 'a compact_basis set" where "approximants = (\x. {a. Rep_compact_basis a \ x})" -interpretation compact_basis!: +interpretation compact_basis: ideal_completion sq_le compact_take Rep_compact_basis approximants proof fix w :: 'a show "preorder.ideal sq_le (approximants w)" proof (rule sq_le.idealI) show "\x. x \ approximants w" unfolding approximants_def apply (rule_tac x="Abs_compact_basis (approx 0\w)" in exI) apply (simp add: Abs_compact_basis_inverse approx_less) done next fix x y :: "'a compact_basis" assume "x \ approximants w" "y \ approximants w" thus "\z \ approximants w. x \ z \ y \ z" unfolding approximants_def apply simp apply (cut_tac a=x in compact_Rep_compact_basis) apply (cut_tac a=y in compact_Rep_compact_basis) apply (drule profinite_compact_eq_approx) apply (drule profinite_compact_eq_approx) apply (clarify, rename_tac i j) apply (rule_tac x="Abs_compact_basis (approx (max i j)\w)" in exI) apply (simp add: compact_le_def) apply (simp add: Abs_compact_basis_inverse approx_less) apply (erule subst, erule subst) apply (simp add: monofun_cfun chain_mono [OF chain_approx]) done next fix x y :: "'a compact_basis" assume "x \ y" "y \ approximants w" thus "x \ approximants w" unfolding approximants_def apply simp apply (simp add: compact_le_def) apply (erule (1) trans_less) done qed next fix Y :: "nat \ 'a" assume Y: "chain Y" show "approximants (\i. Y i) = (\i. approximants (Y i))" unfolding approximants_def apply safe apply (simp add: compactD2 [OF compact_Rep_compact_basis Y]) apply (erule trans_less, rule is_ub_thelub [OF Y]) done next fix a :: "'a compact_basis" show "approximants (Rep_compact_basis a) = {b. b \ a}" unfolding approximants_def compact_le_def .. next fix x y :: "'a" assume "approximants x \ approximants y" thus "x \ y" apply (subgoal_tac "(\i. approx i\x) \ y", simp) apply (rule admD, simp, simp) apply (drule_tac c="Abs_compact_basis (approx i\x)" in subsetD) apply (simp add: approximants_def Abs_compact_basis_inverse approx_less) apply (simp add: approximants_def Abs_compact_basis_inverse) done qed text {* minimal compact element *} definition compact_bot :: "'a::bifinite compact_basis" where "compact_bot = Abs_compact_basis \" lemma Rep_compact_bot: "Rep_compact_basis compact_bot = \" unfolding compact_bot_def by (simp add: Abs_compact_basis_inverse) lemma compact_bot_minimal [simp]: "compact_bot \ a" unfolding compact_le_def Rep_compact_bot by simp subsection {* A compact basis for powerdomains *} typedef 'a pd_basis = "{S::'a::profinite compact_basis set. finite S \ S \ {}}" by (rule_tac x="{arbitrary}" in exI, simp) lemma finite_Rep_pd_basis [simp]: "finite (Rep_pd_basis u)" by (insert Rep_pd_basis [of u, unfolded pd_basis_def]) simp lemma Rep_pd_basis_nonempty [simp]: "Rep_pd_basis u \ {}" by (insert Rep_pd_basis [of u, unfolded pd_basis_def]) simp text {* unit and plus *} definition PDUnit :: "'a compact_basis \ 'a pd_basis" where "PDUnit = (\x. Abs_pd_basis {x})" definition PDPlus :: "'a pd_basis \ 'a pd_basis \ 'a pd_basis" where "PDPlus t u = Abs_pd_basis (Rep_pd_basis t \ Rep_pd_basis u)" lemma Rep_PDUnit: "Rep_pd_basis (PDUnit x) = {x}" unfolding PDUnit_def by (rule Abs_pd_basis_inverse) (simp add: pd_basis_def) lemma Rep_PDPlus: "Rep_pd_basis (PDPlus u v) = Rep_pd_basis u \ Rep_pd_basis v" unfolding PDPlus_def by (rule Abs_pd_basis_inverse) (simp add: pd_basis_def) lemma PDUnit_inject [simp]: "(PDUnit a = PDUnit b) = (a = b)" unfolding Rep_pd_basis_inject [symmetric] Rep_PDUnit by simp lemma PDPlus_assoc: "PDPlus (PDPlus t u) v = PDPlus t (PDPlus u v)" unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_assoc) lemma PDPlus_commute: "PDPlus t u = PDPlus u t" unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_commute) lemma PDPlus_absorb: "PDPlus t t = t" unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_absorb) lemma pd_basis_induct1: assumes PDUnit: "\a. P (PDUnit a)" assumes PDPlus: "\a t. P t \ P (PDPlus (PDUnit a) t)" shows "P x" apply (induct x, unfold pd_basis_def, clarify) apply (erule (1) finite_ne_induct) apply (cut_tac a=x in PDUnit) apply (simp add: PDUnit_def) apply (drule_tac a=x in PDPlus) apply (simp add: PDUnit_def PDPlus_def Abs_pd_basis_inverse [unfolded pd_basis_def]) done lemma pd_basis_induct: assumes PDUnit: "\a. P (PDUnit a)" assumes PDPlus: "\t u. \P t; P u\ \ P (PDPlus t u)" shows "P x" apply (induct x rule: pd_basis_induct1) apply (rule PDUnit, erule PDPlus [OF PDUnit]) done text {* fold-pd *} definition fold_pd :: "('a compact_basis \ 'b::type) \ ('b \ 'b \ 'b) \ 'a pd_basis \ 'b" where "fold_pd g f t = fold1 f (g ` Rep_pd_basis t)" lemma fold_pd_PDUnit: assumes "ab_semigroup_idem_mult f" shows "fold_pd g f (PDUnit x) = g x" unfolding fold_pd_def Rep_PDUnit by simp lemma fold_pd_PDPlus: assumes "ab_semigroup_idem_mult f" shows "fold_pd g f (PDPlus t u) = f (fold_pd g f t) (fold_pd g f u)" proof - interpret ab_semigroup_idem_mult f by fact show ?thesis unfolding fold_pd_def Rep_PDPlus by (simp add: image_Un fold1_Un2) qed text {* Take function for powerdomain basis *} definition pd_take :: "nat \ 'a pd_basis \ 'a pd_basis" where "pd_take n = (\t. Abs_pd_basis (compact_take n ` Rep_pd_basis t))" lemma Rep_pd_take: "Rep_pd_basis (pd_take n t) = compact_take n ` Rep_pd_basis t" unfolding pd_take_def apply (rule Abs_pd_basis_inverse) apply (simp add: pd_basis_def) done lemma pd_take_simps [simp]: "pd_take n (PDUnit a) = PDUnit (compact_take n a)" "pd_take n (PDPlus t u) = PDPlus (pd_take n t) (pd_take n u)" apply (simp_all add: Rep_pd_basis_inject [symmetric]) apply (simp_all add: Rep_pd_take Rep_PDUnit Rep_PDPlus image_Un) done lemma pd_take_idem: "pd_take n (pd_take n t) = pd_take n t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_basis.take_take) apply simp done lemma finite_range_pd_take: "finite (range (pd_take n))" apply (rule finite_imageD [where f="Rep_pd_basis"]) apply (rule finite_subset [where B="Pow (range (compact_take n))"]) apply (clarsimp simp add: Rep_pd_take) apply (simp add: compact_basis.finite_range_take) apply (rule inj_onI, simp add: Rep_pd_basis_inject) done lemma pd_take_covers: "\n. pd_take n t = t" apply (subgoal_tac "\n. \m\n. pd_take m t = t", fast) apply (induct t rule: pd_basis_induct) apply (cut_tac a=a in compact_basis.take_covers) apply (clarify, rule_tac x=n in exI) apply (clarify, simp) apply (rule antisym_less) apply (rule compact_basis.take_less) apply (drule_tac a=a in compact_basis.take_chain_le) apply simp apply (clarify, rename_tac i j) apply (rule_tac x="max i j" in exI, simp) done end diff --git a/src/HOLCF/Completion.thy b/src/HOLCF/Completion.thy --- a/src/HOLCF/Completion.thy +++ b/src/HOLCF/Completion.thy @@ -1,483 +1,483 @@ (* Title: HOLCF/Completion.thy Author: Brian Huffman *) header {* Defining bifinite domains by ideal completion *} theory Completion imports Bifinite begin subsection {* Ideals over a preorder *} locale preorder = fixes r :: "'a::type \ 'a \ bool" (infix "\" 50) assumes r_refl: "x \ x" assumes r_trans: "\x \ y; y \ z\ \ x \ z" begin definition ideal :: "'a set \ bool" where "ideal A = ((\x. x \ A) \ (\x\A. \y\A. \z\A. x \ z \ y \ z) \ (\x y. x \ y \ y \ A \ x \ A))" lemma idealI: assumes "\x. x \ A" assumes "\x y. \x \ A; y \ A\ \ \z\A. x \ z \ y \ z" assumes "\x y. \x \ y; y \ A\ \ x \ A" shows "ideal A" unfolding ideal_def using prems by fast lemma idealD1: "ideal A \ \x. x \ A" unfolding ideal_def by fast lemma idealD2: "\ideal A; x \ A; y \ A\ \ \z\A. x \ z \ y \ z" unfolding ideal_def by fast lemma idealD3: "\ideal A; x \ y; y \ A\ \ x \ A" unfolding ideal_def by fast lemma ideal_directed_finite: assumes A: "ideal A" shows "\finite U; U \ A\ \ \z\A. \x\U. x \ z" apply (induct U set: finite) apply (simp add: idealD1 [OF A]) apply (simp, clarify, rename_tac y) apply (drule (1) idealD2 [OF A]) apply (clarify, erule_tac x=z in rev_bexI) apply (fast intro: r_trans) done lemma ideal_principal: "ideal {x. x \ z}" apply (rule idealI) apply (rule_tac x=z in exI) apply (fast intro: r_refl) apply (rule_tac x=z in bexI, fast) apply (fast intro: r_refl) apply (fast intro: r_trans) done lemma ex_ideal: "\A. ideal A" by (rule exI, rule ideal_principal) lemma directed_image_ideal: assumes A: "ideal A" assumes f: "\x y. x \ y \ f x \ f y" shows "directed (f ` A)" apply (rule directedI) apply (cut_tac idealD1 [OF A], fast) apply (clarify, rename_tac a b) apply (drule (1) idealD2 [OF A]) apply (clarify, rename_tac c) apply (rule_tac x="f c" in rev_bexI) apply (erule imageI) apply (simp add: f) done lemma lub_image_principal: assumes f: "\x y. x \ y \ f x \ f y" shows "(\x\{x. x \ y}. f x) = f y" apply (rule thelubI) apply (rule is_lub_maximal) apply (rule ub_imageI) apply (simp add: f) apply (rule imageI) apply (simp add: r_refl) done text {* The set of ideals is a cpo *} lemma ideal_UN: fixes A :: "nat \ 'a set" assumes ideal_A: "\i. ideal (A i)" assumes chain_A: "\i j. i \ j \ A i \ A j" shows "ideal (\i. A i)" apply (rule idealI) apply (cut_tac idealD1 [OF ideal_A], fast) apply (clarify, rename_tac i j) apply (drule subsetD [OF chain_A [OF le_maxI1]]) apply (drule subsetD [OF chain_A [OF le_maxI2]]) apply (drule (1) idealD2 [OF ideal_A]) apply blast apply clarify apply (drule (1) idealD3 [OF ideal_A]) apply fast done lemma typedef_ideal_po: fixes Abs :: "'a set \ 'b::sq_ord" assumes type: "type_definition Rep Abs {S. ideal S}" assumes less: "\x y. x \ y \ Rep x \ Rep y" shows "OFCLASS('b, po_class)" apply (intro_classes, unfold less) apply (rule subset_refl) apply (erule (1) subset_trans) apply (rule type_definition.Rep_inject [OF type, THEN iffD1]) apply (erule (1) subset_antisym) done lemma fixes Abs :: "'a set \ 'b::po" assumes type: "type_definition Rep Abs {S. ideal S}" assumes less: "\x y. x \ y \ Rep x \ Rep y" assumes S: "chain S" shows typedef_ideal_lub: "range S <<| Abs (\i. Rep (S i))" and typedef_ideal_rep_contlub: "Rep (\i. S i) = (\i. Rep (S i))" proof - have 1: "ideal (\i. Rep (S i))" apply (rule ideal_UN) apply (rule type_definition.Rep [OF type, unfolded mem_Collect_eq]) apply (subst less [symmetric]) apply (erule chain_mono [OF S]) done hence 2: "Rep (Abs (\i. Rep (S i))) = (\i. Rep (S i))" by (simp add: type_definition.Abs_inverse [OF type]) show 3: "range S <<| Abs (\i. Rep (S i))" apply (rule is_lubI) apply (rule is_ubI) apply (simp add: less 2, fast) apply (simp add: less 2 is_ub_def, fast) done hence 4: "(\i. S i) = Abs (\i. Rep (S i))" by (rule thelubI) show 5: "Rep (\i. S i) = (\i. Rep (S i))" by (simp add: 4 2) qed lemma typedef_ideal_cpo: fixes Abs :: "'a set \ 'b::po" assumes type: "type_definition Rep Abs {S. ideal S}" assumes less: "\x y. x \ y \ Rep x \ Rep y" shows "OFCLASS('b, cpo_class)" by (default, rule exI, erule typedef_ideal_lub [OF type less]) end -interpretation sq_le!: preorder "sq_le :: 'a::po \ 'a \ bool" +interpretation sq_le: preorder "sq_le :: 'a::po \ 'a \ bool" apply unfold_locales apply (rule refl_less) apply (erule (1) trans_less) done subsection {* Lemmas about least upper bounds *} lemma finite_directed_contains_lub: "\finite S; directed S\ \ \u\S. S <<| u" apply (drule (1) directed_finiteD, rule subset_refl) apply (erule bexE) apply (rule rev_bexI, assumption) apply (erule (1) is_lub_maximal) done lemma lub_finite_directed_in_self: "\finite S; directed S\ \ lub S \ S" apply (drule (1) finite_directed_contains_lub, clarify) apply (drule thelubI, simp) done lemma finite_directed_has_lub: "\finite S; directed S\ \ \u. S <<| u" by (drule (1) finite_directed_contains_lub, fast) lemma is_ub_thelub0: "\\u. S <<| u; x \ S\ \ x \ lub S" apply (erule exE, drule lubI) apply (drule is_lubD1) apply (erule (1) is_ubD) done lemma is_lub_thelub0: "\\u. S <<| u; S <| x\ \ lub S \ x" by (erule exE, drule lubI, erule is_lub_lub) subsection {* Locale for ideal completion *} locale basis_take = preorder + fixes take :: "nat \ 'a::type \ 'a" assumes take_less: "take n a \ a" assumes take_take: "take n (take n a) = take n a" assumes take_mono: "a \ b \ take n a \ take n b" assumes take_chain: "take n a \ take (Suc n) a" assumes finite_range_take: "finite (range (take n))" assumes take_covers: "\n. take n a = a" begin lemma take_chain_less: "m < n \ take m a \ take n a" by (erule less_Suc_induct, rule take_chain, erule (1) r_trans) lemma take_chain_le: "m \ n \ take m a \ take n a" by (cases "m = n", simp add: r_refl, simp add: take_chain_less) end locale ideal_completion = basis_take + fixes principal :: "'a::type \ 'b::cpo" fixes rep :: "'b::cpo \ 'a::type set" assumes ideal_rep: "\x. preorder.ideal r (rep x)" assumes rep_contlub: "\Y. chain Y \ rep (\i. Y i) = (\i. rep (Y i))" assumes rep_principal: "\a. rep (principal a) = {b. b \ a}" assumes subset_repD: "\x y. rep x \ rep y \ x \ y" begin lemma finite_take_rep: "finite (take n ` rep x)" by (rule finite_subset [OF image_mono [OF subset_UNIV] finite_range_take]) lemma rep_mono: "x \ y \ rep x \ rep y" apply (frule bin_chain) apply (drule rep_contlub) apply (simp only: thelubI [OF lub_bin_chain]) apply (rule subsetI, rule UN_I [where a=0], simp_all) done lemma less_def: "x \ y \ rep x \ rep y" by (rule iffI [OF rep_mono subset_repD]) lemma rep_eq: "rep x = {a. principal a \ x}" unfolding less_def rep_principal apply safe apply (erule (1) idealD3 [OF ideal_rep]) apply (erule subsetD, simp add: r_refl) done lemma mem_rep_iff_principal_less: "a \ rep x \ principal a \ x" by (simp add: rep_eq) lemma principal_less_iff_mem_rep: "principal a \ x \ a \ rep x" by (simp add: rep_eq) lemma principal_less_iff [simp]: "principal a \ principal b \ a \ b" by (simp add: principal_less_iff_mem_rep rep_principal) lemma principal_eq_iff: "principal a = principal b \ a \ b \ b \ a" unfolding po_eq_conv [where 'a='b] principal_less_iff .. lemma repD: "a \ rep x \ principal a \ x" by (simp add: rep_eq) lemma principal_mono: "a \ b \ principal a \ principal b" by (simp only: principal_less_iff) lemma lessI: "(\a. principal a \ x \ principal a \ u) \ x \ u" unfolding principal_less_iff_mem_rep by (simp add: less_def subset_eq) lemma lub_principal_rep: "principal ` rep x <<| x" apply (rule is_lubI) apply (rule ub_imageI) apply (erule repD) apply (subst less_def) apply (rule subsetI) apply (drule (1) ub_imageD) apply (simp add: rep_eq) done subsection {* Defining functions in terms of basis elements *} definition basis_fun :: "('a::type \ 'c::cpo) \ 'b \ 'c" where "basis_fun = (\f. (\ x. lub (f ` rep x)))" lemma basis_fun_lemma0: fixes f :: "'a::type \ 'c::cpo" assumes f_mono: "\a b. a \ b \ f a \ f b" shows "\u. f ` take i ` rep x <<| u" apply (rule finite_directed_has_lub) apply (rule finite_imageI) apply (rule finite_take_rep) apply (subst image_image) apply (rule directed_image_ideal) apply (rule ideal_rep) apply (rule f_mono) apply (erule take_mono) done lemma basis_fun_lemma1: fixes f :: "'a::type \ 'c::cpo" assumes f_mono: "\a b. a \ b \ f a \ f b" shows "chain (\i. lub (f ` take i ` rep x))" apply (rule chainI) apply (rule is_lub_thelub0) apply (rule basis_fun_lemma0, erule f_mono) apply (rule is_ubI, clarsimp, rename_tac a) apply (rule trans_less [OF f_mono [OF take_chain]]) apply (rule is_ub_thelub0) apply (rule basis_fun_lemma0, erule f_mono) apply simp done lemma basis_fun_lemma2: fixes f :: "'a::type \ 'c::cpo" assumes f_mono: "\a b. a \ b \ f a \ f b" shows "f ` rep x <<| (\i. lub (f ` take i ` rep x))" apply (rule is_lubI) apply (rule ub_imageI, rename_tac a) apply (cut_tac a=a in take_covers, erule exE, rename_tac i) apply (erule subst) apply (rule rev_trans_less) apply (rule_tac x=i in is_ub_thelub) apply (rule basis_fun_lemma1, erule f_mono) apply (rule is_ub_thelub0) apply (rule basis_fun_lemma0, erule f_mono) apply simp apply (rule is_lub_thelub [OF _ ub_rangeI]) apply (rule basis_fun_lemma1, erule f_mono) apply (rule is_lub_thelub0) apply (rule basis_fun_lemma0, erule f_mono) apply (rule is_ubI, clarsimp, rename_tac a) apply (rule trans_less [OF f_mono [OF take_less]]) apply (erule (1) ub_imageD) done lemma basis_fun_lemma: fixes f :: "'a::type \ 'c::cpo" assumes f_mono: "\a b. a \ b \ f a \ f b" shows "\u. f ` rep x <<| u" by (rule exI, rule basis_fun_lemma2, erule f_mono) lemma basis_fun_beta: fixes f :: "'a::type \ 'c::cpo" assumes f_mono: "\a b. a \ b \ f a \ f b" shows "basis_fun f\x = lub (f ` rep x)" unfolding basis_fun_def proof (rule beta_cfun) have lub: "\x. \u. f ` rep x <<| u" using f_mono by (rule basis_fun_lemma) show cont: "cont (\x. lub (f ` rep x))" apply (rule contI2) apply (rule monofunI) apply (rule is_lub_thelub0 [OF lub ub_imageI]) apply (rule is_ub_thelub0 [OF lub imageI]) apply (erule (1) subsetD [OF rep_mono]) apply (rule is_lub_thelub0 [OF lub ub_imageI]) apply (simp add: rep_contlub, clarify) apply (erule rev_trans_less [OF is_ub_thelub]) apply (erule is_ub_thelub0 [OF lub imageI]) done qed lemma basis_fun_principal: fixes f :: "'a::type \ 'c::cpo" assumes f_mono: "\a b. a \ b \ f a \ f b" shows "basis_fun f\(principal a) = f a" apply (subst basis_fun_beta, erule f_mono) apply (subst rep_principal) apply (rule lub_image_principal, erule f_mono) done lemma basis_fun_mono: assumes f_mono: "\a b. a \ b \ f a \ f b" assumes g_mono: "\a b. a \ b \ g a \ g b" assumes less: "\a. f a \ g a" shows "basis_fun f \ basis_fun g" apply (rule less_cfun_ext) apply (simp only: basis_fun_beta f_mono g_mono) apply (rule is_lub_thelub0) apply (rule basis_fun_lemma, erule f_mono) apply (rule ub_imageI, rename_tac a) apply (rule trans_less [OF less]) apply (rule is_ub_thelub0) apply (rule basis_fun_lemma, erule g_mono) apply (erule imageI) done lemma compact_principal [simp]: "compact (principal a)" by (rule compactI2, simp add: principal_less_iff_mem_rep rep_contlub) subsection {* Bifiniteness of ideal completions *} definition completion_approx :: "nat \ 'b \ 'b" where "completion_approx = (\i. basis_fun (\a. principal (take i a)))" lemma completion_approx_beta: "completion_approx i\x = (\a\rep x. principal (take i a))" unfolding completion_approx_def by (simp add: basis_fun_beta principal_mono take_mono) lemma completion_approx_principal: "completion_approx i\(principal a) = principal (take i a)" unfolding completion_approx_def by (simp add: basis_fun_principal principal_mono take_mono) lemma chain_completion_approx: "chain completion_approx" unfolding completion_approx_def apply (rule chainI) apply (rule basis_fun_mono) apply (erule principal_mono [OF take_mono]) apply (erule principal_mono [OF take_mono]) apply (rule principal_mono [OF take_chain]) done lemma lub_completion_approx: "(\i. completion_approx i\x) = x" unfolding completion_approx_beta apply (subst image_image [where f=principal, symmetric]) apply (rule unique_lub [OF _ lub_principal_rep]) apply (rule basis_fun_lemma2, erule principal_mono) done lemma completion_approx_eq_principal: "\a\rep x. completion_approx i\x = principal (take i a)" unfolding completion_approx_beta apply (subst image_image [where f=principal, symmetric]) apply (subgoal_tac "finite (principal ` take i ` rep x)") apply (subgoal_tac "directed (principal ` take i ` rep x)") apply (drule (1) lub_finite_directed_in_self, fast) apply (subst image_image) apply (rule directed_image_ideal) apply (rule ideal_rep) apply (erule principal_mono [OF take_mono]) apply (rule finite_imageI) apply (rule finite_take_rep) done lemma completion_approx_idem: "completion_approx i\(completion_approx i\x) = completion_approx i\x" using completion_approx_eq_principal [where i=i and x=x] by (auto simp add: completion_approx_principal take_take) lemma finite_fixes_completion_approx: "finite {x. completion_approx i\x = x}" (is "finite ?S") apply (subgoal_tac "?S \ principal ` range (take i)") apply (erule finite_subset) apply (rule finite_imageI) apply (rule finite_range_take) apply (clarify, erule subst) apply (cut_tac x=x and i=i in completion_approx_eq_principal) apply fast done lemma principal_induct: assumes adm: "adm P" assumes P: "\a. P (principal a)" shows "P x" apply (subgoal_tac "P (\i. completion_approx i\x)") apply (simp add: lub_completion_approx) apply (rule admD [OF adm]) apply (simp add: chain_completion_approx) apply (cut_tac x=x and i=i in completion_approx_eq_principal) apply (clarify, simp add: P) done lemma principal_induct2: "\\y. adm (\x. P x y); \x. adm (\y. P x y); \a b. P (principal a) (principal b)\ \ P x y" apply (rule_tac x=y in spec) apply (rule_tac x=x in principal_induct, simp) apply (rule allI, rename_tac y) apply (rule_tac x=y in principal_induct, simp) apply simp done lemma compact_imp_principal: "compact x \ \a. x = principal a" apply (drule adm_compact_neq [OF _ cont_id]) apply (drule admD2 [where Y="\n. completion_approx n\x"]) apply (simp add: chain_completion_approx) apply (simp add: lub_completion_approx) apply (erule exE, erule ssubst) apply (cut_tac i=i and x=x in completion_approx_eq_principal) apply (clarify, erule exI) done end end diff --git a/src/HOLCF/ConvexPD.thy b/src/HOLCF/ConvexPD.thy --- a/src/HOLCF/ConvexPD.thy +++ b/src/HOLCF/ConvexPD.thy @@ -1,638 +1,638 @@ (* Title: HOLCF/ConvexPD.thy Author: Brian Huffman *) header {* Convex powerdomain *} theory ConvexPD imports UpperPD LowerPD begin subsection {* Basis preorder *} definition convex_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix "\\" 50) where "convex_le = (\u v. u \\ v \ u \\ v)" lemma convex_le_refl [simp]: "t \\ t" unfolding convex_le_def by (fast intro: upper_le_refl lower_le_refl) lemma convex_le_trans: "\t \\ u; u \\ v\ \ t \\ v" unfolding convex_le_def by (fast intro: upper_le_trans lower_le_trans) -interpretation convex_le!: preorder convex_le +interpretation convex_le: preorder convex_le by (rule preorder.intro, rule convex_le_refl, rule convex_le_trans) lemma upper_le_minimal [simp]: "PDUnit compact_bot \\ t" unfolding convex_le_def Rep_PDUnit by simp lemma PDUnit_convex_mono: "x \ y \ PDUnit x \\ PDUnit y" unfolding convex_le_def by (fast intro: PDUnit_upper_mono PDUnit_lower_mono) lemma PDPlus_convex_mono: "\s \\ t; u \\ v\ \ PDPlus s u \\ PDPlus t v" unfolding convex_le_def by (fast intro: PDPlus_upper_mono PDPlus_lower_mono) lemma convex_le_PDUnit_PDUnit_iff [simp]: "(PDUnit a \\ PDUnit b) = a \ b" unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit by fast lemma convex_le_PDUnit_lemma1: "(PDUnit a \\ t) = (\b\Rep_pd_basis t. a \ b)" unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast lemma convex_le_PDUnit_PDPlus_iff [simp]: "(PDUnit a \\ PDPlus t u) = (PDUnit a \\ t \ PDUnit a \\ u)" unfolding convex_le_PDUnit_lemma1 Rep_PDPlus by fast lemma convex_le_PDUnit_lemma2: "(t \\ PDUnit b) = (\a\Rep_pd_basis t. a \ b)" unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast lemma convex_le_PDPlus_PDUnit_iff [simp]: "(PDPlus t u \\ PDUnit a) = (t \\ PDUnit a \ u \\ PDUnit a)" unfolding convex_le_PDUnit_lemma2 Rep_PDPlus by fast lemma convex_le_PDPlus_lemma: assumes z: "PDPlus t u \\ z" shows "\v w. z = PDPlus v w \ t \\ v \ u \\ w" proof (intro exI conjI) let ?A = "{b\Rep_pd_basis z. \a\Rep_pd_basis t. a \ b}" let ?B = "{b\Rep_pd_basis z. \a\Rep_pd_basis u. a \ b}" let ?v = "Abs_pd_basis ?A" let ?w = "Abs_pd_basis ?B" have Rep_v: "Rep_pd_basis ?v = ?A" apply (rule Abs_pd_basis_inverse) apply (rule Rep_pd_basis_nonempty [of t, folded ex_in_conv, THEN exE]) apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify) apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE) apply (simp add: pd_basis_def) apply fast done have Rep_w: "Rep_pd_basis ?w = ?B" apply (rule Abs_pd_basis_inverse) apply (rule Rep_pd_basis_nonempty [of u, folded ex_in_conv, THEN exE]) apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify) apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE) apply (simp add: pd_basis_def) apply fast done show "z = PDPlus ?v ?w" apply (insert z) apply (simp add: convex_le_def, erule conjE) apply (simp add: Rep_pd_basis_inject [symmetric] Rep_PDPlus) apply (simp add: Rep_v Rep_w) apply (rule equalityI) apply (rule subsetI) apply (simp only: upper_le_def) apply (drule (1) bspec, erule bexE) apply (simp add: Rep_PDPlus) apply fast apply fast done show "t \\ ?v" "u \\ ?w" apply (insert z) apply (simp_all add: convex_le_def upper_le_def lower_le_def Rep_PDPlus Rep_v Rep_w) apply fast+ done qed lemma convex_le_induct [induct set: convex_le]: assumes le: "t \\ u" assumes 2: "\t u v. \P t u; P u v\ \ P t v" assumes 3: "\a b. a \ b \ P (PDUnit a) (PDUnit b)" assumes 4: "\t u v w. \P t v; P u w\ \ P (PDPlus t u) (PDPlus v w)" shows "P t u" using le apply (induct t arbitrary: u rule: pd_basis_induct) apply (erule rev_mp) apply (induct_tac u rule: pd_basis_induct1) apply (simp add: 3) apply (simp, clarify, rename_tac a b t) apply (subgoal_tac "P (PDPlus (PDUnit a) (PDUnit a)) (PDPlus (PDUnit b) t)") apply (simp add: PDPlus_absorb) apply (erule (1) 4 [OF 3]) apply (drule convex_le_PDPlus_lemma, clarify) apply (simp add: 4) done lemma pd_take_convex_chain: "pd_take n t \\ pd_take (Suc n) t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_basis.take_chain) apply (simp add: PDPlus_convex_mono) done lemma pd_take_convex_le: "pd_take i t \\ t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_basis.take_less) apply (simp add: PDPlus_convex_mono) done lemma pd_take_convex_mono: "t \\ u \ pd_take n t \\ pd_take n u" apply (erule convex_le_induct) apply (erule (1) convex_le_trans) apply (simp add: compact_basis.take_mono) apply (simp add: PDPlus_convex_mono) done subsection {* Type definition *} typedef (open) 'a convex_pd = "{S::'a pd_basis set. convex_le.ideal S}" by (fast intro: convex_le.ideal_principal) instantiation convex_pd :: (profinite) sq_ord begin definition "x \ y \ Rep_convex_pd x \ Rep_convex_pd y" instance .. end instance convex_pd :: (profinite) po by (rule convex_le.typedef_ideal_po [OF type_definition_convex_pd sq_le_convex_pd_def]) instance convex_pd :: (profinite) cpo by (rule convex_le.typedef_ideal_cpo [OF type_definition_convex_pd sq_le_convex_pd_def]) lemma Rep_convex_pd_lub: "chain Y \ Rep_convex_pd (\i. Y i) = (\i. Rep_convex_pd (Y i))" by (rule convex_le.typedef_ideal_rep_contlub [OF type_definition_convex_pd sq_le_convex_pd_def]) lemma ideal_Rep_convex_pd: "convex_le.ideal (Rep_convex_pd xs)" by (rule Rep_convex_pd [unfolded mem_Collect_eq]) definition convex_principal :: "'a pd_basis \ 'a convex_pd" where "convex_principal t = Abs_convex_pd {u. u \\ t}" lemma Rep_convex_principal: "Rep_convex_pd (convex_principal t) = {u. u \\ t}" unfolding convex_principal_def by (simp add: Abs_convex_pd_inverse convex_le.ideal_principal) -interpretation convex_pd!: +interpretation convex_pd: ideal_completion convex_le pd_take convex_principal Rep_convex_pd apply unfold_locales apply (rule pd_take_convex_le) apply (rule pd_take_idem) apply (erule pd_take_convex_mono) apply (rule pd_take_convex_chain) apply (rule finite_range_pd_take) apply (rule pd_take_covers) apply (rule ideal_Rep_convex_pd) apply (erule Rep_convex_pd_lub) apply (rule Rep_convex_principal) apply (simp only: sq_le_convex_pd_def) done text {* Convex powerdomain is pointed *} lemma convex_pd_minimal: "convex_principal (PDUnit compact_bot) \ ys" by (induct ys rule: convex_pd.principal_induct, simp, simp) instance convex_pd :: (bifinite) pcpo by intro_classes (fast intro: convex_pd_minimal) lemma inst_convex_pd_pcpo: "\ = convex_principal (PDUnit compact_bot)" by (rule convex_pd_minimal [THEN UU_I, symmetric]) text {* Convex powerdomain is profinite *} instantiation convex_pd :: (profinite) profinite begin definition approx_convex_pd_def: "approx = convex_pd.completion_approx" instance apply (intro_classes, unfold approx_convex_pd_def) apply (rule convex_pd.chain_completion_approx) apply (rule convex_pd.lub_completion_approx) apply (rule convex_pd.completion_approx_idem) apply (rule convex_pd.finite_fixes_completion_approx) done end instance convex_pd :: (bifinite) bifinite .. lemma approx_convex_principal [simp]: "approx n\(convex_principal t) = convex_principal (pd_take n t)" unfolding approx_convex_pd_def by (rule convex_pd.completion_approx_principal) lemma approx_eq_convex_principal: "\t\Rep_convex_pd xs. approx n\xs = convex_principal (pd_take n t)" unfolding approx_convex_pd_def by (rule convex_pd.completion_approx_eq_principal) subsection {* Monadic unit and plus *} definition convex_unit :: "'a \ 'a convex_pd" where "convex_unit = compact_basis.basis_fun (\a. convex_principal (PDUnit a))" definition convex_plus :: "'a convex_pd \ 'a convex_pd \ 'a convex_pd" where "convex_plus = convex_pd.basis_fun (\t. convex_pd.basis_fun (\u. convex_principal (PDPlus t u)))" abbreviation convex_add :: "'a convex_pd \ 'a convex_pd \ 'a convex_pd" (infixl "+\" 65) where "xs +\ ys == convex_plus\xs\ys" syntax "_convex_pd" :: "args \ 'a convex_pd" ("{_}\") translations "{x,xs}\" == "{x}\ +\ {xs}\" "{x}\" == "CONST convex_unit\x" lemma convex_unit_Rep_compact_basis [simp]: "{Rep_compact_basis a}\ = convex_principal (PDUnit a)" unfolding convex_unit_def by (simp add: compact_basis.basis_fun_principal PDUnit_convex_mono) lemma convex_plus_principal [simp]: "convex_principal t +\ convex_principal u = convex_principal (PDPlus t u)" unfolding convex_plus_def by (simp add: convex_pd.basis_fun_principal convex_pd.basis_fun_mono PDPlus_convex_mono) lemma approx_convex_unit [simp]: "approx n\{x}\ = {approx n\x}\" apply (induct x rule: compact_basis.principal_induct, simp) apply (simp add: approx_Rep_compact_basis) done lemma approx_convex_plus [simp]: "approx n\(xs +\ ys) = approx n\xs +\ approx n\ys" by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) lemma convex_plus_assoc: "(xs +\ ys) +\ zs = xs +\ (ys +\ zs)" apply (induct xs ys arbitrary: zs rule: convex_pd.principal_induct2, simp, simp) apply (rule_tac x=zs in convex_pd.principal_induct, simp) apply (simp add: PDPlus_assoc) done lemma convex_plus_commute: "xs +\ ys = ys +\ xs" apply (induct xs ys rule: convex_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_commute) done lemma convex_plus_absorb [simp]: "xs +\ xs = xs" apply (induct xs rule: convex_pd.principal_induct, simp) apply (simp add: PDPlus_absorb) done lemma convex_plus_left_commute: "xs +\ (ys +\ zs) = ys +\ (xs +\ zs)" by (rule mk_left_commute [of "op +\", OF convex_plus_assoc convex_plus_commute]) lemma convex_plus_left_absorb [simp]: "xs +\ (xs +\ ys) = xs +\ ys" by (simp only: convex_plus_assoc [symmetric] convex_plus_absorb) text {* Useful for @{text "simp add: convex_plus_ac"} *} lemmas convex_plus_ac = convex_plus_assoc convex_plus_commute convex_plus_left_commute text {* Useful for @{text "simp only: convex_plus_aci"} *} lemmas convex_plus_aci = convex_plus_ac convex_plus_absorb convex_plus_left_absorb lemma convex_unit_less_plus_iff [simp]: "{x}\ \ ys +\ zs \ {x}\ \ ys \ {x}\ \ zs" apply (rule iffI) apply (subgoal_tac "adm (\f. f\{x}\ \ f\ys \ f\{x}\ \ f\zs)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) apply (cut_tac x="approx i\ys" in convex_pd.compact_imp_principal, simp) apply (cut_tac x="approx i\zs" in convex_pd.compact_imp_principal, simp) apply (clarify, simp) apply simp apply simp apply (erule conjE) apply (subst convex_plus_absorb [of "{x}\", symmetric]) apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) done lemma convex_plus_less_unit_iff [simp]: "xs +\ ys \ {z}\ \ xs \ {z}\ \ ys \ {z}\" apply (rule iffI) apply (subgoal_tac "adm (\f. f\xs \ f\{z}\ \ f\ys \ f\{z}\)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) apply (cut_tac x="approx i\xs" in convex_pd.compact_imp_principal, simp) apply (cut_tac x="approx i\ys" in convex_pd.compact_imp_principal, simp) apply (cut_tac x="approx i\z" in compact_basis.compact_imp_principal, simp) apply (clarify, simp) apply simp apply simp apply (erule conjE) apply (subst convex_plus_absorb [of "{z}\", symmetric]) apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) done lemma convex_unit_less_iff [simp]: "{x}\ \ {y}\ \ x \ y" apply (rule iffI) apply (rule profinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg, simp) apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) apply (cut_tac x="approx i\y" in compact_basis.compact_imp_principal, simp) apply clarsimp apply (erule monofun_cfun_arg) done lemma convex_unit_eq_iff [simp]: "{x}\ = {y}\ \ x = y" unfolding po_eq_conv by simp lemma convex_unit_strict [simp]: "{\}\ = \" unfolding inst_convex_pd_pcpo Rep_compact_bot [symmetric] by simp lemma convex_unit_strict_iff [simp]: "{x}\ = \ \ x = \" unfolding convex_unit_strict [symmetric] by (rule convex_unit_eq_iff) lemma compact_convex_unit_iff [simp]: "compact {x}\ \ compact x" unfolding profinite_compact_iff by simp lemma compact_convex_plus [simp]: "\compact xs; compact ys\ \ compact (xs +\ ys)" by (auto dest!: convex_pd.compact_imp_principal) subsection {* Induction rules *} lemma convex_pd_induct1: assumes P: "adm P" assumes unit: "\x. P {x}\" assumes insert: "\x ys. \P {x}\; P ys\ \ P ({x}\ +\ ys)" shows "P (xs::'a convex_pd)" apply (induct xs rule: convex_pd.principal_induct, rule P) apply (induct_tac a rule: pd_basis_induct1) apply (simp only: convex_unit_Rep_compact_basis [symmetric]) apply (rule unit) apply (simp only: convex_unit_Rep_compact_basis [symmetric] convex_plus_principal [symmetric]) apply (erule insert [OF unit]) done lemma convex_pd_induct: assumes P: "adm P" assumes unit: "\x. P {x}\" assumes plus: "\xs ys. \P xs; P ys\ \ P (xs +\ ys)" shows "P (xs::'a convex_pd)" apply (induct xs rule: convex_pd.principal_induct, rule P) apply (induct_tac a rule: pd_basis_induct) apply (simp only: convex_unit_Rep_compact_basis [symmetric] unit) apply (simp only: convex_plus_principal [symmetric] plus) done subsection {* Monadic bind *} definition convex_bind_basis :: "'a pd_basis \ ('a \ 'b convex_pd) \ 'b convex_pd" where "convex_bind_basis = fold_pd (\a. \ f. f\(Rep_compact_basis a)) (\x y. \ f. x\f +\ y\f)" lemma ACI_convex_bind: "ab_semigroup_idem_mult (\x y. \ f. x\f +\ y\f)" apply unfold_locales apply (simp add: convex_plus_assoc) apply (simp add: convex_plus_commute) apply (simp add: eta_cfun) done lemma convex_bind_basis_simps [simp]: "convex_bind_basis (PDUnit a) = (\ f. f\(Rep_compact_basis a))" "convex_bind_basis (PDPlus t u) = (\ f. convex_bind_basis t\f +\ convex_bind_basis u\f)" unfolding convex_bind_basis_def apply - apply (rule fold_pd_PDUnit [OF ACI_convex_bind]) apply (rule fold_pd_PDPlus [OF ACI_convex_bind]) done lemma monofun_LAM: "\cont f; cont g; \x. f x \ g x\ \ (\ x. f x) \ (\ x. g x)" by (simp add: expand_cfun_less) lemma convex_bind_basis_mono: "t \\ u \ convex_bind_basis t \ convex_bind_basis u" apply (erule convex_le_induct) apply (erule (1) trans_less) apply (simp add: monofun_LAM monofun_cfun) apply (simp add: monofun_LAM monofun_cfun) done definition convex_bind :: "'a convex_pd \ ('a \ 'b convex_pd) \ 'b convex_pd" where "convex_bind = convex_pd.basis_fun convex_bind_basis" lemma convex_bind_principal [simp]: "convex_bind\(convex_principal t) = convex_bind_basis t" unfolding convex_bind_def apply (rule convex_pd.basis_fun_principal) apply (erule convex_bind_basis_mono) done lemma convex_bind_unit [simp]: "convex_bind\{x}\\f = f\x" by (induct x rule: compact_basis.principal_induct, simp, simp) lemma convex_bind_plus [simp]: "convex_bind\(xs +\ ys)\f = convex_bind\xs\f +\ convex_bind\ys\f" by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) lemma convex_bind_strict [simp]: "convex_bind\\\f = f\\" unfolding convex_unit_strict [symmetric] by (rule convex_bind_unit) subsection {* Map and join *} definition convex_map :: "('a \ 'b) \ 'a convex_pd \ 'b convex_pd" where "convex_map = (\ f xs. convex_bind\xs\(\ x. {f\x}\))" definition convex_join :: "'a convex_pd convex_pd \ 'a convex_pd" where "convex_join = (\ xss. convex_bind\xss\(\ xs. xs))" lemma convex_map_unit [simp]: "convex_map\f\(convex_unit\x) = convex_unit\(f\x)" unfolding convex_map_def by simp lemma convex_map_plus [simp]: "convex_map\f\(xs +\ ys) = convex_map\f\xs +\ convex_map\f\ys" unfolding convex_map_def by simp lemma convex_join_unit [simp]: "convex_join\{xs}\ = xs" unfolding convex_join_def by simp lemma convex_join_plus [simp]: "convex_join\(xss +\ yss) = convex_join\xss +\ convex_join\yss" unfolding convex_join_def by simp lemma convex_map_ident: "convex_map\(\ x. x)\xs = xs" by (induct xs rule: convex_pd_induct, simp_all) lemma convex_map_map: "convex_map\f\(convex_map\g\xs) = convex_map\(\ x. f\(g\x))\xs" by (induct xs rule: convex_pd_induct, simp_all) lemma convex_join_map_unit: "convex_join\(convex_map\convex_unit\xs) = xs" by (induct xs rule: convex_pd_induct, simp_all) lemma convex_join_map_join: "convex_join\(convex_map\convex_join\xsss) = convex_join\(convex_join\xsss)" by (induct xsss rule: convex_pd_induct, simp_all) lemma convex_join_map_map: "convex_join\(convex_map\(convex_map\f)\xss) = convex_map\f\(convex_join\xss)" by (induct xss rule: convex_pd_induct, simp_all) lemma convex_map_approx: "convex_map\(approx n)\xs = approx n\xs" by (induct xs rule: convex_pd_induct, simp_all) subsection {* Conversions to other powerdomains *} text {* Convex to upper *} lemma convex_le_imp_upper_le: "t \\ u \ t \\ u" unfolding convex_le_def by simp definition convex_to_upper :: "'a convex_pd \ 'a upper_pd" where "convex_to_upper = convex_pd.basis_fun upper_principal" lemma convex_to_upper_principal [simp]: "convex_to_upper\(convex_principal t) = upper_principal t" unfolding convex_to_upper_def apply (rule convex_pd.basis_fun_principal) apply (rule upper_pd.principal_mono) apply (erule convex_le_imp_upper_le) done lemma convex_to_upper_unit [simp]: "convex_to_upper\{x}\ = {x}\" by (induct x rule: compact_basis.principal_induct, simp, simp) lemma convex_to_upper_plus [simp]: "convex_to_upper\(xs +\ ys) = convex_to_upper\xs +\ convex_to_upper\ys" by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) lemma approx_convex_to_upper: "approx i\(convex_to_upper\xs) = convex_to_upper\(approx i\xs)" by (induct xs rule: convex_pd_induct, simp, simp, simp) lemma convex_to_upper_bind [simp]: "convex_to_upper\(convex_bind\xs\f) = upper_bind\(convex_to_upper\xs)\(convex_to_upper oo f)" by (induct xs rule: convex_pd_induct, simp, simp, simp) lemma convex_to_upper_map [simp]: "convex_to_upper\(convex_map\f\xs) = upper_map\f\(convex_to_upper\xs)" by (simp add: convex_map_def upper_map_def cfcomp_LAM) lemma convex_to_upper_join [simp]: "convex_to_upper\(convex_join\xss) = upper_bind\(convex_to_upper\xss)\convex_to_upper" by (simp add: convex_join_def upper_join_def cfcomp_LAM eta_cfun) text {* Convex to lower *} lemma convex_le_imp_lower_le: "t \\ u \ t \\ u" unfolding convex_le_def by simp definition convex_to_lower :: "'a convex_pd \ 'a lower_pd" where "convex_to_lower = convex_pd.basis_fun lower_principal" lemma convex_to_lower_principal [simp]: "convex_to_lower\(convex_principal t) = lower_principal t" unfolding convex_to_lower_def apply (rule convex_pd.basis_fun_principal) apply (rule lower_pd.principal_mono) apply (erule convex_le_imp_lower_le) done lemma convex_to_lower_unit [simp]: "convex_to_lower\{x}\ = {x}\" by (induct x rule: compact_basis.principal_induct, simp, simp) lemma convex_to_lower_plus [simp]: "convex_to_lower\(xs +\ ys) = convex_to_lower\xs +\ convex_to_lower\ys" by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp) lemma approx_convex_to_lower: "approx i\(convex_to_lower\xs) = convex_to_lower\(approx i\xs)" by (induct xs rule: convex_pd_induct, simp, simp, simp) lemma convex_to_lower_bind [simp]: "convex_to_lower\(convex_bind\xs\f) = lower_bind\(convex_to_lower\xs)\(convex_to_lower oo f)" by (induct xs rule: convex_pd_induct, simp, simp, simp) lemma convex_to_lower_map [simp]: "convex_to_lower\(convex_map\f\xs) = lower_map\f\(convex_to_lower\xs)" by (simp add: convex_map_def lower_map_def cfcomp_LAM) lemma convex_to_lower_join [simp]: "convex_to_lower\(convex_join\xss) = lower_bind\(convex_to_lower\xss)\convex_to_lower" by (simp add: convex_join_def lower_join_def cfcomp_LAM eta_cfun) text {* Ordering property *} lemma convex_pd_less_iff: "(xs \ ys) = (convex_to_upper\xs \ convex_to_upper\ys \ convex_to_lower\xs \ convex_to_lower\ys)" apply (safe elim!: monofun_cfun_arg) apply (rule profinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg) apply (drule_tac f="approx i" in monofun_cfun_arg) apply (cut_tac x="approx i\xs" in convex_pd.compact_imp_principal, simp) apply (cut_tac x="approx i\ys" in convex_pd.compact_imp_principal, simp) apply clarify apply (simp add: approx_convex_to_upper approx_convex_to_lower convex_le_def) done lemmas convex_plus_less_plus_iff = convex_pd_less_iff [where xs="xs +\ ys" and ys="zs +\ ws", standard] lemmas convex_pd_less_simps = convex_unit_less_plus_iff convex_plus_less_unit_iff convex_plus_less_plus_iff convex_unit_less_iff convex_to_upper_unit convex_to_upper_plus convex_to_lower_unit convex_to_lower_plus upper_pd_less_simps lower_pd_less_simps end diff --git a/src/HOLCF/LowerPD.thy b/src/HOLCF/LowerPD.thy --- a/src/HOLCF/LowerPD.thy +++ b/src/HOLCF/LowerPD.thy @@ -1,494 +1,494 @@ (* Title: HOLCF/LowerPD.thy Author: Brian Huffman *) header {* Lower powerdomain *} theory LowerPD imports CompactBasis begin subsection {* Basis preorder *} definition lower_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix "\\" 50) where "lower_le = (\u v. \x\Rep_pd_basis u. \y\Rep_pd_basis v. x \ y)" lemma lower_le_refl [simp]: "t \\ t" unfolding lower_le_def by fast lemma lower_le_trans: "\t \\ u; u \\ v\ \ t \\ v" unfolding lower_le_def apply (rule ballI) apply (drule (1) bspec, erule bexE) apply (drule (1) bspec, erule bexE) apply (erule rev_bexI) apply (erule (1) trans_less) done -interpretation lower_le!: preorder lower_le +interpretation lower_le: preorder lower_le by (rule preorder.intro, rule lower_le_refl, rule lower_le_trans) lemma lower_le_minimal [simp]: "PDUnit compact_bot \\ t" unfolding lower_le_def Rep_PDUnit by (simp, rule Rep_pd_basis_nonempty [folded ex_in_conv]) lemma PDUnit_lower_mono: "x \ y \ PDUnit x \\ PDUnit y" unfolding lower_le_def Rep_PDUnit by fast lemma PDPlus_lower_mono: "\s \\ t; u \\ v\ \ PDPlus s u \\ PDPlus t v" unfolding lower_le_def Rep_PDPlus by fast lemma PDPlus_lower_less: "t \\ PDPlus t u" unfolding lower_le_def Rep_PDPlus by fast lemma lower_le_PDUnit_PDUnit_iff [simp]: "(PDUnit a \\ PDUnit b) = a \ b" unfolding lower_le_def Rep_PDUnit by fast lemma lower_le_PDUnit_PDPlus_iff: "(PDUnit a \\ PDPlus t u) = (PDUnit a \\ t \ PDUnit a \\ u)" unfolding lower_le_def Rep_PDPlus Rep_PDUnit by fast lemma lower_le_PDPlus_iff: "(PDPlus t u \\ v) = (t \\ v \ u \\ v)" unfolding lower_le_def Rep_PDPlus by fast lemma lower_le_induct [induct set: lower_le]: assumes le: "t \\ u" assumes 1: "\a b. a \ b \ P (PDUnit a) (PDUnit b)" assumes 2: "\t u a. P (PDUnit a) t \ P (PDUnit a) (PDPlus t u)" assumes 3: "\t u v. \P t v; P u v\ \ P (PDPlus t u) v" shows "P t u" using le apply (induct t arbitrary: u rule: pd_basis_induct) apply (erule rev_mp) apply (induct_tac u rule: pd_basis_induct) apply (simp add: 1) apply (simp add: lower_le_PDUnit_PDPlus_iff) apply (simp add: 2) apply (subst PDPlus_commute) apply (simp add: 2) apply (simp add: lower_le_PDPlus_iff 3) done lemma pd_take_lower_chain: "pd_take n t \\ pd_take (Suc n) t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_basis.take_chain) apply (simp add: PDPlus_lower_mono) done lemma pd_take_lower_le: "pd_take i t \\ t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_basis.take_less) apply (simp add: PDPlus_lower_mono) done lemma pd_take_lower_mono: "t \\ u \ pd_take n t \\ pd_take n u" apply (erule lower_le_induct) apply (simp add: compact_basis.take_mono) apply (simp add: lower_le_PDUnit_PDPlus_iff) apply (simp add: lower_le_PDPlus_iff) done subsection {* Type definition *} typedef (open) 'a lower_pd = "{S::'a pd_basis set. lower_le.ideal S}" by (fast intro: lower_le.ideal_principal) instantiation lower_pd :: (profinite) sq_ord begin definition "x \ y \ Rep_lower_pd x \ Rep_lower_pd y" instance .. end instance lower_pd :: (profinite) po by (rule lower_le.typedef_ideal_po [OF type_definition_lower_pd sq_le_lower_pd_def]) instance lower_pd :: (profinite) cpo by (rule lower_le.typedef_ideal_cpo [OF type_definition_lower_pd sq_le_lower_pd_def]) lemma Rep_lower_pd_lub: "chain Y \ Rep_lower_pd (\i. Y i) = (\i. Rep_lower_pd (Y i))" by (rule lower_le.typedef_ideal_rep_contlub [OF type_definition_lower_pd sq_le_lower_pd_def]) lemma ideal_Rep_lower_pd: "lower_le.ideal (Rep_lower_pd xs)" by (rule Rep_lower_pd [unfolded mem_Collect_eq]) definition lower_principal :: "'a pd_basis \ 'a lower_pd" where "lower_principal t = Abs_lower_pd {u. u \\ t}" lemma Rep_lower_principal: "Rep_lower_pd (lower_principal t) = {u. u \\ t}" unfolding lower_principal_def by (simp add: Abs_lower_pd_inverse lower_le.ideal_principal) -interpretation lower_pd!: +interpretation lower_pd: ideal_completion lower_le pd_take lower_principal Rep_lower_pd apply unfold_locales apply (rule pd_take_lower_le) apply (rule pd_take_idem) apply (erule pd_take_lower_mono) apply (rule pd_take_lower_chain) apply (rule finite_range_pd_take) apply (rule pd_take_covers) apply (rule ideal_Rep_lower_pd) apply (erule Rep_lower_pd_lub) apply (rule Rep_lower_principal) apply (simp only: sq_le_lower_pd_def) done text {* Lower powerdomain is pointed *} lemma lower_pd_minimal: "lower_principal (PDUnit compact_bot) \ ys" by (induct ys rule: lower_pd.principal_induct, simp, simp) instance lower_pd :: (bifinite) pcpo by intro_classes (fast intro: lower_pd_minimal) lemma inst_lower_pd_pcpo: "\ = lower_principal (PDUnit compact_bot)" by (rule lower_pd_minimal [THEN UU_I, symmetric]) text {* Lower powerdomain is profinite *} instantiation lower_pd :: (profinite) profinite begin definition approx_lower_pd_def: "approx = lower_pd.completion_approx" instance apply (intro_classes, unfold approx_lower_pd_def) apply (rule lower_pd.chain_completion_approx) apply (rule lower_pd.lub_completion_approx) apply (rule lower_pd.completion_approx_idem) apply (rule lower_pd.finite_fixes_completion_approx) done end instance lower_pd :: (bifinite) bifinite .. lemma approx_lower_principal [simp]: "approx n\(lower_principal t) = lower_principal (pd_take n t)" unfolding approx_lower_pd_def by (rule lower_pd.completion_approx_principal) lemma approx_eq_lower_principal: "\t\Rep_lower_pd xs. approx n\xs = lower_principal (pd_take n t)" unfolding approx_lower_pd_def by (rule lower_pd.completion_approx_eq_principal) subsection {* Monadic unit and plus *} definition lower_unit :: "'a \ 'a lower_pd" where "lower_unit = compact_basis.basis_fun (\a. lower_principal (PDUnit a))" definition lower_plus :: "'a lower_pd \ 'a lower_pd \ 'a lower_pd" where "lower_plus = lower_pd.basis_fun (\t. lower_pd.basis_fun (\u. lower_principal (PDPlus t u)))" abbreviation lower_add :: "'a lower_pd \ 'a lower_pd \ 'a lower_pd" (infixl "+\" 65) where "xs +\ ys == lower_plus\xs\ys" syntax "_lower_pd" :: "args \ 'a lower_pd" ("{_}\") translations "{x,xs}\" == "{x}\ +\ {xs}\" "{x}\" == "CONST lower_unit\x" lemma lower_unit_Rep_compact_basis [simp]: "{Rep_compact_basis a}\ = lower_principal (PDUnit a)" unfolding lower_unit_def by (simp add: compact_basis.basis_fun_principal PDUnit_lower_mono) lemma lower_plus_principal [simp]: "lower_principal t +\ lower_principal u = lower_principal (PDPlus t u)" unfolding lower_plus_def by (simp add: lower_pd.basis_fun_principal lower_pd.basis_fun_mono PDPlus_lower_mono) lemma approx_lower_unit [simp]: "approx n\{x}\ = {approx n\x}\" apply (induct x rule: compact_basis.principal_induct, simp) apply (simp add: approx_Rep_compact_basis) done lemma approx_lower_plus [simp]: "approx n\(xs +\ ys) = (approx n\xs) +\ (approx n\ys)" by (induct xs ys rule: lower_pd.principal_induct2, simp, simp, simp) lemma lower_plus_assoc: "(xs +\ ys) +\ zs = xs +\ (ys +\ zs)" apply (induct xs ys arbitrary: zs rule: lower_pd.principal_induct2, simp, simp) apply (rule_tac x=zs in lower_pd.principal_induct, simp) apply (simp add: PDPlus_assoc) done lemma lower_plus_commute: "xs +\ ys = ys +\ xs" apply (induct xs ys rule: lower_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_commute) done lemma lower_plus_absorb [simp]: "xs +\ xs = xs" apply (induct xs rule: lower_pd.principal_induct, simp) apply (simp add: PDPlus_absorb) done lemma lower_plus_left_commute: "xs +\ (ys +\ zs) = ys +\ (xs +\ zs)" by (rule mk_left_commute [of "op +\", OF lower_plus_assoc lower_plus_commute]) lemma lower_plus_left_absorb [simp]: "xs +\ (xs +\ ys) = xs +\ ys" by (simp only: lower_plus_assoc [symmetric] lower_plus_absorb) text {* Useful for @{text "simp add: lower_plus_ac"} *} lemmas lower_plus_ac = lower_plus_assoc lower_plus_commute lower_plus_left_commute text {* Useful for @{text "simp only: lower_plus_aci"} *} lemmas lower_plus_aci = lower_plus_ac lower_plus_absorb lower_plus_left_absorb lemma lower_plus_less1: "xs \ xs +\ ys" apply (induct xs ys rule: lower_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_lower_less) done lemma lower_plus_less2: "ys \ xs +\ ys" by (subst lower_plus_commute, rule lower_plus_less1) lemma lower_plus_least: "\xs \ zs; ys \ zs\ \ xs +\ ys \ zs" apply (subst lower_plus_absorb [of zs, symmetric]) apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) done lemma lower_plus_less_iff: "xs +\ ys \ zs \ xs \ zs \ ys \ zs" apply safe apply (erule trans_less [OF lower_plus_less1]) apply (erule trans_less [OF lower_plus_less2]) apply (erule (1) lower_plus_least) done lemma lower_unit_less_plus_iff: "{x}\ \ ys +\ zs \ {x}\ \ ys \ {x}\ \ zs" apply (rule iffI) apply (subgoal_tac "adm (\f. f\{x}\ \ f\ys \ f\{x}\ \ f\zs)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) apply (cut_tac x="approx i\ys" in lower_pd.compact_imp_principal, simp) apply (cut_tac x="approx i\zs" in lower_pd.compact_imp_principal, simp) apply (clarify, simp add: lower_le_PDUnit_PDPlus_iff) apply simp apply simp apply (erule disjE) apply (erule trans_less [OF _ lower_plus_less1]) apply (erule trans_less [OF _ lower_plus_less2]) done lemma lower_unit_less_iff [simp]: "{x}\ \ {y}\ \ x \ y" apply (rule iffI) apply (rule profinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg, simp) apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) apply (cut_tac x="approx i\y" in compact_basis.compact_imp_principal, simp) apply clarsimp apply (erule monofun_cfun_arg) done lemmas lower_pd_less_simps = lower_unit_less_iff lower_plus_less_iff lower_unit_less_plus_iff lemma lower_unit_eq_iff [simp]: "{x}\ = {y}\ \ x = y" by (simp add: po_eq_conv) lemma lower_unit_strict [simp]: "{\}\ = \" unfolding inst_lower_pd_pcpo Rep_compact_bot [symmetric] by simp lemma lower_unit_strict_iff [simp]: "{x}\ = \ \ x = \" unfolding lower_unit_strict [symmetric] by (rule lower_unit_eq_iff) lemma lower_plus_strict_iff [simp]: "xs +\ ys = \ \ xs = \ \ ys = \" apply safe apply (rule UU_I, erule subst, rule lower_plus_less1) apply (rule UU_I, erule subst, rule lower_plus_less2) apply (rule lower_plus_absorb) done lemma lower_plus_strict1 [simp]: "\ +\ ys = ys" apply (rule antisym_less [OF _ lower_plus_less2]) apply (simp add: lower_plus_least) done lemma lower_plus_strict2 [simp]: "xs +\ \ = xs" apply (rule antisym_less [OF _ lower_plus_less1]) apply (simp add: lower_plus_least) done lemma compact_lower_unit_iff [simp]: "compact {x}\ \ compact x" unfolding profinite_compact_iff by simp lemma compact_lower_plus [simp]: "\compact xs; compact ys\ \ compact (xs +\ ys)" by (auto dest!: lower_pd.compact_imp_principal) subsection {* Induction rules *} lemma lower_pd_induct1: assumes P: "adm P" assumes unit: "\x. P {x}\" assumes insert: "\x ys. \P {x}\; P ys\ \ P ({x}\ +\ ys)" shows "P (xs::'a lower_pd)" apply (induct xs rule: lower_pd.principal_induct, rule P) apply (induct_tac a rule: pd_basis_induct1) apply (simp only: lower_unit_Rep_compact_basis [symmetric]) apply (rule unit) apply (simp only: lower_unit_Rep_compact_basis [symmetric] lower_plus_principal [symmetric]) apply (erule insert [OF unit]) done lemma lower_pd_induct: assumes P: "adm P" assumes unit: "\x. P {x}\" assumes plus: "\xs ys. \P xs; P ys\ \ P (xs +\ ys)" shows "P (xs::'a lower_pd)" apply (induct xs rule: lower_pd.principal_induct, rule P) apply (induct_tac a rule: pd_basis_induct) apply (simp only: lower_unit_Rep_compact_basis [symmetric] unit) apply (simp only: lower_plus_principal [symmetric] plus) done subsection {* Monadic bind *} definition lower_bind_basis :: "'a pd_basis \ ('a \ 'b lower_pd) \ 'b lower_pd" where "lower_bind_basis = fold_pd (\a. \ f. f\(Rep_compact_basis a)) (\x y. \ f. x\f +\ y\f)" lemma ACI_lower_bind: "ab_semigroup_idem_mult (\x y. \ f. x\f +\ y\f)" apply unfold_locales apply (simp add: lower_plus_assoc) apply (simp add: lower_plus_commute) apply (simp add: eta_cfun) done lemma lower_bind_basis_simps [simp]: "lower_bind_basis (PDUnit a) = (\ f. f\(Rep_compact_basis a))" "lower_bind_basis (PDPlus t u) = (\ f. lower_bind_basis t\f +\ lower_bind_basis u\f)" unfolding lower_bind_basis_def apply - apply (rule fold_pd_PDUnit [OF ACI_lower_bind]) apply (rule fold_pd_PDPlus [OF ACI_lower_bind]) done lemma lower_bind_basis_mono: "t \\ u \ lower_bind_basis t \ lower_bind_basis u" unfolding expand_cfun_less apply (erule lower_le_induct, safe) apply (simp add: monofun_cfun) apply (simp add: rev_trans_less [OF lower_plus_less1]) apply (simp add: lower_plus_less_iff) done definition lower_bind :: "'a lower_pd \ ('a \ 'b lower_pd) \ 'b lower_pd" where "lower_bind = lower_pd.basis_fun lower_bind_basis" lemma lower_bind_principal [simp]: "lower_bind\(lower_principal t) = lower_bind_basis t" unfolding lower_bind_def apply (rule lower_pd.basis_fun_principal) apply (erule lower_bind_basis_mono) done lemma lower_bind_unit [simp]: "lower_bind\{x}\\f = f\x" by (induct x rule: compact_basis.principal_induct, simp, simp) lemma lower_bind_plus [simp]: "lower_bind\(xs +\ ys)\f = lower_bind\xs\f +\ lower_bind\ys\f" by (induct xs ys rule: lower_pd.principal_induct2, simp, simp, simp) lemma lower_bind_strict [simp]: "lower_bind\\\f = f\\" unfolding lower_unit_strict [symmetric] by (rule lower_bind_unit) subsection {* Map and join *} definition lower_map :: "('a \ 'b) \ 'a lower_pd \ 'b lower_pd" where "lower_map = (\ f xs. lower_bind\xs\(\ x. {f\x}\))" definition lower_join :: "'a lower_pd lower_pd \ 'a lower_pd" where "lower_join = (\ xss. lower_bind\xss\(\ xs. xs))" lemma lower_map_unit [simp]: "lower_map\f\{x}\ = {f\x}\" unfolding lower_map_def by simp lemma lower_map_plus [simp]: "lower_map\f\(xs +\ ys) = lower_map\f\xs +\ lower_map\f\ys" unfolding lower_map_def by simp lemma lower_join_unit [simp]: "lower_join\{xs}\ = xs" unfolding lower_join_def by simp lemma lower_join_plus [simp]: "lower_join\(xss +\ yss) = lower_join\xss +\ lower_join\yss" unfolding lower_join_def by simp lemma lower_map_ident: "lower_map\(\ x. x)\xs = xs" by (induct xs rule: lower_pd_induct, simp_all) lemma lower_map_map: "lower_map\f\(lower_map\g\xs) = lower_map\(\ x. f\(g\x))\xs" by (induct xs rule: lower_pd_induct, simp_all) lemma lower_join_map_unit: "lower_join\(lower_map\lower_unit\xs) = xs" by (induct xs rule: lower_pd_induct, simp_all) lemma lower_join_map_join: "lower_join\(lower_map\lower_join\xsss) = lower_join\(lower_join\xsss)" by (induct xsss rule: lower_pd_induct, simp_all) lemma lower_join_map_map: "lower_join\(lower_map\(lower_map\f)\xss) = lower_map\f\(lower_join\xss)" by (induct xss rule: lower_pd_induct, simp_all) lemma lower_map_approx: "lower_map\(approx n)\xs = approx n\xs" by (induct xs rule: lower_pd_induct, simp_all) end diff --git a/src/HOLCF/Universal.thy b/src/HOLCF/Universal.thy --- a/src/HOLCF/Universal.thy +++ b/src/HOLCF/Universal.thy @@ -1,860 +1,860 @@ (* Title: HOLCF/Universal.thy Author: Brian Huffman *) theory Universal imports CompactBasis NatIso begin subsection {* Basis datatype *} types ubasis = nat definition node :: "nat \ ubasis \ ubasis set \ ubasis" where "node i a S = Suc (prod2nat (i, prod2nat (a, set2nat S)))" lemma node_not_0 [simp]: "node i a S \ 0" unfolding node_def by simp lemma node_gt_0 [simp]: "0 < node i a S" unfolding node_def by simp lemma node_inject [simp]: "\finite S; finite T\ \ node i a S = node j b T \ i = j \ a = b \ S = T" unfolding node_def by simp lemma node_gt0: "i < node i a S" unfolding node_def less_Suc_eq_le by (rule le_prod2nat_1) lemma node_gt1: "a < node i a S" unfolding node_def less_Suc_eq_le by (rule order_trans [OF le_prod2nat_1 le_prod2nat_2]) lemma nat_less_power2: "n < 2^n" by (induct n) simp_all lemma node_gt2: "\finite S; b \ S\ \ b < node i a S" unfolding node_def less_Suc_eq_le set2nat_def apply (rule order_trans [OF _ le_prod2nat_2]) apply (rule order_trans [OF _ le_prod2nat_2]) apply (rule order_trans [where y="setsum (op ^ 2) {b}"]) apply (simp add: nat_less_power2 [THEN order_less_imp_le]) apply (erule setsum_mono2, simp, simp) done lemma eq_prod2nat_pairI: "\fst (nat2prod x) = a; snd (nat2prod x) = b\ \ x = prod2nat (a, b)" by (erule subst, erule subst, simp) lemma node_cases: assumes 1: "x = 0 \ P" assumes 2: "\i a S. \finite S; x = node i a S\ \ P" shows "P" apply (cases x) apply (erule 1) apply (rule 2) apply (rule finite_nat2set) apply (simp add: node_def) apply (rule eq_prod2nat_pairI [OF refl]) apply (rule eq_prod2nat_pairI [OF refl refl]) done lemma node_induct: assumes 1: "P 0" assumes 2: "\i a S. \P a; finite S; \b\S. P b\ \ P (node i a S)" shows "P x" apply (induct x rule: nat_less_induct) apply (case_tac n rule: node_cases) apply (simp add: 1) apply (simp add: 2 node_gt1 node_gt2) done subsection {* Basis ordering *} inductive ubasis_le :: "nat \ nat \ bool" where ubasis_le_refl: "ubasis_le a a" | ubasis_le_trans: "\ubasis_le a b; ubasis_le b c\ \ ubasis_le a c" | ubasis_le_lower: "finite S \ ubasis_le a (node i a S)" | ubasis_le_upper: "\finite S; b \ S; ubasis_le a b\ \ ubasis_le (node i a S) b" lemma ubasis_le_minimal: "ubasis_le 0 x" apply (induct x rule: node_induct) apply (rule ubasis_le_refl) apply (erule ubasis_le_trans) apply (erule ubasis_le_lower) done subsubsection {* Generic take function *} function ubasis_until :: "(ubasis \ bool) \ ubasis \ ubasis" where "ubasis_until P 0 = 0" | "finite S \ ubasis_until P (node i a S) = (if P (node i a S) then node i a S else ubasis_until P a)" apply clarify apply (rule_tac x=b in node_cases) apply simp apply simp apply fast apply simp apply simp apply simp done termination ubasis_until apply (relation "measure snd") apply (rule wf_measure) apply (simp add: node_gt1) done lemma ubasis_until: "P 0 \ P (ubasis_until P x)" by (induct x rule: node_induct) simp_all lemma ubasis_until': "0 < ubasis_until P x \ P (ubasis_until P x)" by (induct x rule: node_induct) auto lemma ubasis_until_same: "P x \ ubasis_until P x = x" by (induct x rule: node_induct) simp_all lemma ubasis_until_idem: "P 0 \ ubasis_until P (ubasis_until P x) = ubasis_until P x" by (rule ubasis_until_same [OF ubasis_until]) lemma ubasis_until_0: "\x. x \ 0 \ \ P x \ ubasis_until P x = 0" by (induct x rule: node_induct) simp_all lemma ubasis_until_less: "ubasis_le (ubasis_until P x) x" apply (induct x rule: node_induct) apply (simp add: ubasis_le_refl) apply (simp add: ubasis_le_refl) apply (rule impI) apply (erule ubasis_le_trans) apply (erule ubasis_le_lower) done lemma ubasis_until_chain: assumes PQ: "\x. P x \ Q x" shows "ubasis_le (ubasis_until P x) (ubasis_until Q x)" apply (induct x rule: node_induct) apply (simp add: ubasis_le_refl) apply (simp add: ubasis_le_refl) apply (simp add: PQ) apply clarify apply (rule ubasis_le_trans) apply (rule ubasis_until_less) apply (erule ubasis_le_lower) done lemma ubasis_until_mono: assumes "\i a S b. \finite S; P (node i a S); b \ S; ubasis_le a b\ \ P b" shows "ubasis_le a b \ ubasis_le (ubasis_until P a) (ubasis_until P b)" proof (induct set: ubasis_le) case (ubasis_le_refl a) show ?case by (rule ubasis_le.ubasis_le_refl) next case (ubasis_le_trans a b c) thus ?case by - (rule ubasis_le.ubasis_le_trans) next case (ubasis_le_lower S a i) thus ?case apply (clarsimp simp add: ubasis_le_refl) apply (rule ubasis_le_trans [OF ubasis_until_less]) apply (erule ubasis_le.ubasis_le_lower) done next case (ubasis_le_upper S b a i) thus ?case apply clarsimp apply (subst ubasis_until_same) apply (erule (3) prems) apply (erule (2) ubasis_le.ubasis_le_upper) done qed lemma finite_range_ubasis_until: "finite {x. P x} \ finite (range (ubasis_until P))" apply (rule finite_subset [where B="insert 0 {x. P x}"]) apply (clarsimp simp add: ubasis_until') apply simp done subsubsection {* Take function for @{typ ubasis} *} definition ubasis_take :: "nat \ ubasis \ ubasis" where "ubasis_take n = ubasis_until (\x. x \ n)" lemma ubasis_take_le: "ubasis_take n x \ n" unfolding ubasis_take_def by (rule ubasis_until, rule le0) lemma ubasis_take_same: "x \ n \ ubasis_take n x = x" unfolding ubasis_take_def by (rule ubasis_until_same) lemma ubasis_take_idem: "ubasis_take n (ubasis_take n x) = ubasis_take n x" by (rule ubasis_take_same [OF ubasis_take_le]) lemma ubasis_take_0 [simp]: "ubasis_take 0 x = 0" unfolding ubasis_take_def by (simp add: ubasis_until_0) lemma ubasis_take_less: "ubasis_le (ubasis_take n x) x" unfolding ubasis_take_def by (rule ubasis_until_less) lemma ubasis_take_chain: "ubasis_le (ubasis_take n x) (ubasis_take (Suc n) x)" unfolding ubasis_take_def by (rule ubasis_until_chain) simp lemma ubasis_take_mono: assumes "ubasis_le x y" shows "ubasis_le (ubasis_take n x) (ubasis_take n y)" unfolding ubasis_take_def apply (rule ubasis_until_mono [OF _ prems]) apply (frule (2) order_less_le_trans [OF node_gt2]) apply (erule order_less_imp_le) done lemma finite_range_ubasis_take: "finite (range (ubasis_take n))" apply (rule finite_subset [where B="{..n}"]) apply (simp add: subset_eq ubasis_take_le) apply simp done lemma ubasis_take_covers: "\n. ubasis_take n x = x" apply (rule exI [where x=x]) apply (simp add: ubasis_take_same) done -interpretation udom!: preorder ubasis_le +interpretation udom: preorder ubasis_le apply default apply (rule ubasis_le_refl) apply (erule (1) ubasis_le_trans) done -interpretation udom!: basis_take ubasis_le ubasis_take +interpretation udom: basis_take ubasis_le ubasis_take apply default apply (rule ubasis_take_less) apply (rule ubasis_take_idem) apply (erule ubasis_take_mono) apply (rule ubasis_take_chain) apply (rule finite_range_ubasis_take) apply (rule ubasis_take_covers) done subsection {* Defining the universal domain by ideal completion *} typedef (open) udom = "{S. udom.ideal S}" by (fast intro: udom.ideal_principal) instantiation udom :: sq_ord begin definition "x \ y \ Rep_udom x \ Rep_udom y" instance .. end instance udom :: po by (rule udom.typedef_ideal_po [OF type_definition_udom sq_le_udom_def]) instance udom :: cpo by (rule udom.typedef_ideal_cpo [OF type_definition_udom sq_le_udom_def]) lemma Rep_udom_lub: "chain Y \ Rep_udom (\i. Y i) = (\i. Rep_udom (Y i))" by (rule udom.typedef_ideal_rep_contlub [OF type_definition_udom sq_le_udom_def]) lemma ideal_Rep_udom: "udom.ideal (Rep_udom xs)" by (rule Rep_udom [unfolded mem_Collect_eq]) definition udom_principal :: "nat \ udom" where "udom_principal t = Abs_udom {u. ubasis_le u t}" lemma Rep_udom_principal: "Rep_udom (udom_principal t) = {u. ubasis_le u t}" unfolding udom_principal_def by (simp add: Abs_udom_inverse udom.ideal_principal) -interpretation udom!: +interpretation udom: ideal_completion ubasis_le ubasis_take udom_principal Rep_udom apply unfold_locales apply (rule ideal_Rep_udom) apply (erule Rep_udom_lub) apply (rule Rep_udom_principal) apply (simp only: sq_le_udom_def) done text {* Universal domain is pointed *} lemma udom_minimal: "udom_principal 0 \ x" apply (induct x rule: udom.principal_induct) apply (simp, simp add: ubasis_le_minimal) done instance udom :: pcpo by intro_classes (fast intro: udom_minimal) lemma inst_udom_pcpo: "\ = udom_principal 0" by (rule udom_minimal [THEN UU_I, symmetric]) text {* Universal domain is bifinite *} instantiation udom :: bifinite begin definition approx_udom_def: "approx = udom.completion_approx" instance apply (intro_classes, unfold approx_udom_def) apply (rule udom.chain_completion_approx) apply (rule udom.lub_completion_approx) apply (rule udom.completion_approx_idem) apply (rule udom.finite_fixes_completion_approx) done end lemma approx_udom_principal [simp]: "approx n\(udom_principal x) = udom_principal (ubasis_take n x)" unfolding approx_udom_def by (rule udom.completion_approx_principal) lemma approx_eq_udom_principal: "\a\Rep_udom x. approx n\x = udom_principal (ubasis_take n a)" unfolding approx_udom_def by (rule udom.completion_approx_eq_principal) subsection {* Universality of @{typ udom} *} defaultsort bifinite subsubsection {* Choosing a maximal element from a finite set *} lemma finite_has_maximal: fixes A :: "'a::po set" shows "\finite A; A \ {}\ \ \x\A. \y\A. x \ y \ x = y" proof (induct rule: finite_ne_induct) case (singleton x) show ?case by simp next case (insert a A) from `\x\A. \y\A. x \ y \ x = y` obtain x where x: "x \ A" and x_eq: "\y. \y \ A; x \ y\ \ x = y" by fast show ?case proof (intro bexI ballI impI) fix y assume "y \ insert a A" and "(if x \ a then a else x) \ y" thus "(if x \ a then a else x) = y" apply auto apply (frule (1) trans_less) apply (frule (1) x_eq) apply (rule antisym_less, assumption) apply simp apply (erule (1) x_eq) done next show "(if x \ a then a else x) \ insert a A" by (simp add: x) qed qed definition choose :: "'a compact_basis set \ 'a compact_basis" where "choose A = (SOME x. x \ {x\A. \y\A. x \ y \ x = y})" lemma choose_lemma: "\finite A; A \ {}\ \ choose A \ {x\A. \y\A. x \ y \ x = y}" unfolding choose_def apply (rule someI_ex) apply (frule (1) finite_has_maximal, fast) done lemma maximal_choose: "\finite A; y \ A; choose A \ y\ \ choose A = y" apply (cases "A = {}", simp) apply (frule (1) choose_lemma, simp) done lemma choose_in: "\finite A; A \ {}\ \ choose A \ A" by (frule (1) choose_lemma, simp) function choose_pos :: "'a compact_basis set \ 'a compact_basis \ nat" where "choose_pos A x = (if finite A \ x \ A \ x \ choose A then Suc (choose_pos (A - {choose A}) x) else 0)" by auto termination choose_pos apply (relation "measure (card \ fst)", simp) apply clarsimp apply (rule card_Diff1_less) apply assumption apply (erule choose_in) apply clarsimp done declare choose_pos.simps [simp del] lemma choose_pos_choose: "finite A \ choose_pos A (choose A) = 0" by (simp add: choose_pos.simps) lemma inj_on_choose_pos [OF refl]: "\card A = n; finite A\ \ inj_on (choose_pos A) A" apply (induct n arbitrary: A) apply simp apply (case_tac "A = {}", simp) apply (frule (1) choose_in) apply (rule inj_onI) apply (drule_tac x="A - {choose A}" in meta_spec, simp) apply (simp add: choose_pos.simps) apply (simp split: split_if_asm) apply (erule (1) inj_onD, simp, simp) done lemma choose_pos_bounded [OF refl]: "\card A = n; finite A; x \ A\ \ choose_pos A x < n" apply (induct n arbitrary: A) apply simp apply (case_tac "A = {}", simp) apply (frule (1) choose_in) apply (subst choose_pos.simps) apply simp done lemma choose_pos_lessD: "\choose_pos A x < choose_pos A y; finite A; x \ A; y \ A\ \ \ x \ y" apply (induct A x arbitrary: y rule: choose_pos.induct) apply simp apply (case_tac "x = choose A") apply simp apply (rule notI) apply (frule (2) maximal_choose) apply simp apply (case_tac "y = choose A") apply (simp add: choose_pos_choose) apply (drule_tac x=y in meta_spec) apply simp apply (erule meta_mp) apply (simp add: choose_pos.simps) done subsubsection {* Rank of basis elements *} primrec cb_take :: "nat \ 'a compact_basis \ 'a compact_basis" where "cb_take 0 = (\x. compact_bot)" | "cb_take (Suc n) = compact_take n" lemma cb_take_covers: "\n. cb_take n x = x" apply (rule exE [OF compact_basis.take_covers [where a=x]]) apply (rename_tac n, rule_tac x="Suc n" in exI, simp) done lemma cb_take_less: "cb_take n x \ x" by (cases n, simp, simp add: compact_basis.take_less) lemma cb_take_idem: "cb_take n (cb_take n x) = cb_take n x" by (cases n, simp, simp add: compact_basis.take_take) lemma cb_take_mono: "x \ y \ cb_take n x \ cb_take n y" by (cases n, simp, simp add: compact_basis.take_mono) lemma cb_take_chain_le: "m \ n \ cb_take m x \ cb_take n x" apply (cases m, simp) apply (cases n, simp) apply (simp add: compact_basis.take_chain_le) done lemma range_const: "range (\x. c) = {c}" by auto lemma finite_range_cb_take: "finite (range (cb_take n))" apply (cases n) apply (simp add: range_const) apply (simp add: compact_basis.finite_range_take) done definition rank :: "'a compact_basis \ nat" where "rank x = (LEAST n. cb_take n x = x)" lemma compact_approx_rank: "cb_take (rank x) x = x" unfolding rank_def apply (rule LeastI_ex) apply (rule cb_take_covers) done lemma rank_leD: "rank x \ n \ cb_take n x = x" apply (rule antisym_less [OF cb_take_less]) apply (subst compact_approx_rank [symmetric]) apply (erule cb_take_chain_le) done lemma rank_leI: "cb_take n x = x \ rank x \ n" unfolding rank_def by (rule Least_le) lemma rank_le_iff: "rank x \ n \ cb_take n x = x" by (rule iffI [OF rank_leD rank_leI]) lemma rank_compact_bot [simp]: "rank compact_bot = 0" using rank_leI [of 0 compact_bot] by simp lemma rank_eq_0_iff [simp]: "rank x = 0 \ x = compact_bot" using rank_le_iff [of x 0] by auto definition rank_le :: "'a compact_basis \ 'a compact_basis set" where "rank_le x = {y. rank y \ rank x}" definition rank_lt :: "'a compact_basis \ 'a compact_basis set" where "rank_lt x = {y. rank y < rank x}" definition rank_eq :: "'a compact_basis \ 'a compact_basis set" where "rank_eq x = {y. rank y = rank x}" lemma rank_eq_cong: "rank x = rank y \ rank_eq x = rank_eq y" unfolding rank_eq_def by simp lemma rank_lt_cong: "rank x = rank y \ rank_lt x = rank_lt y" unfolding rank_lt_def by simp lemma rank_eq_subset: "rank_eq x \ rank_le x" unfolding rank_eq_def rank_le_def by auto lemma rank_lt_subset: "rank_lt x \ rank_le x" unfolding rank_lt_def rank_le_def by auto lemma finite_rank_le: "finite (rank_le x)" unfolding rank_le_def apply (rule finite_subset [where B="range (cb_take (rank x))"]) apply clarify apply (rule range_eqI) apply (erule rank_leD [symmetric]) apply (rule finite_range_cb_take) done lemma finite_rank_eq: "finite (rank_eq x)" by (rule finite_subset [OF rank_eq_subset finite_rank_le]) lemma finite_rank_lt: "finite (rank_lt x)" by (rule finite_subset [OF rank_lt_subset finite_rank_le]) lemma rank_lt_Int_rank_eq: "rank_lt x \ rank_eq x = {}" unfolding rank_lt_def rank_eq_def rank_le_def by auto lemma rank_lt_Un_rank_eq: "rank_lt x \ rank_eq x = rank_le x" unfolding rank_lt_def rank_eq_def rank_le_def by auto subsubsection {* Sequencing basis elements *} definition place :: "'a compact_basis \ nat" where "place x = card (rank_lt x) + choose_pos (rank_eq x) x" lemma place_bounded: "place x < card (rank_le x)" unfolding place_def apply (rule ord_less_eq_trans) apply (rule add_strict_left_mono) apply (rule choose_pos_bounded) apply (rule finite_rank_eq) apply (simp add: rank_eq_def) apply (subst card_Un_disjoint [symmetric]) apply (rule finite_rank_lt) apply (rule finite_rank_eq) apply (rule rank_lt_Int_rank_eq) apply (simp add: rank_lt_Un_rank_eq) done lemma place_ge: "card (rank_lt x) \ place x" unfolding place_def by simp lemma place_rank_mono: fixes x y :: "'a compact_basis" shows "rank x < rank y \ place x < place y" apply (rule less_le_trans [OF place_bounded]) apply (rule order_trans [OF _ place_ge]) apply (rule card_mono) apply (rule finite_rank_lt) apply (simp add: rank_le_def rank_lt_def subset_eq) done lemma place_eqD: "place x = place y \ x = y" apply (rule linorder_cases [where x="rank x" and y="rank y"]) apply (drule place_rank_mono, simp) apply (simp add: place_def) apply (rule inj_on_choose_pos [where A="rank_eq x", THEN inj_onD]) apply (rule finite_rank_eq) apply (simp cong: rank_lt_cong rank_eq_cong) apply (simp add: rank_eq_def) apply (simp add: rank_eq_def) apply (drule place_rank_mono, simp) done lemma inj_place: "inj place" by (rule inj_onI, erule place_eqD) subsubsection {* Embedding and projection on basis elements *} definition sub :: "'a compact_basis \ 'a compact_basis" where "sub x = (case rank x of 0 \ compact_bot | Suc k \ cb_take k x)" lemma rank_sub_less: "x \ compact_bot \ rank (sub x) < rank x" unfolding sub_def apply (cases "rank x", simp) apply (simp add: less_Suc_eq_le) apply (rule rank_leI) apply (rule cb_take_idem) done lemma place_sub_less: "x \ compact_bot \ place (sub x) < place x" apply (rule place_rank_mono) apply (erule rank_sub_less) done lemma sub_below: "sub x \ x" unfolding sub_def by (cases "rank x", simp_all add: cb_take_less) lemma rank_less_imp_below_sub: "\x \ y; rank x < rank y\ \ x \ sub y" unfolding sub_def apply (cases "rank y", simp) apply (simp add: less_Suc_eq_le) apply (subgoal_tac "cb_take nat x \ cb_take nat y") apply (simp add: rank_leD) apply (erule cb_take_mono) done function basis_emb :: "'a compact_basis \ ubasis" where "basis_emb x = (if x = compact_bot then 0 else node (place x) (basis_emb (sub x)) (basis_emb ` {y. place y < place x \ x \ y}))" by auto termination basis_emb apply (relation "measure place", simp) apply (simp add: place_sub_less) apply simp done declare basis_emb.simps [simp del] lemma basis_emb_compact_bot [simp]: "basis_emb compact_bot = 0" by (simp add: basis_emb.simps) lemma fin1: "finite {y. place y < place x \ x \ y}" apply (subst Collect_conj_eq) apply (rule finite_Int) apply (rule disjI1) apply (subgoal_tac "finite (place -` {n. n < place x})", simp) apply (rule finite_vimageI [OF _ inj_place]) apply (simp add: lessThan_def [symmetric]) done lemma fin2: "finite (basis_emb ` {y. place y < place x \ x \ y})" by (rule finite_imageI [OF fin1]) lemma rank_place_mono: "\place x < place y; x \ y\ \ rank x < rank y" apply (rule linorder_cases, assumption) apply (simp add: place_def cong: rank_lt_cong rank_eq_cong) apply (drule choose_pos_lessD) apply (rule finite_rank_eq) apply (simp add: rank_eq_def) apply (simp add: rank_eq_def) apply simp apply (drule place_rank_mono, simp) done lemma basis_emb_mono: "x \ y \ ubasis_le (basis_emb x) (basis_emb y)" proof (induct n \ "max (place x) (place y)" arbitrary: x y rule: less_induct) case (less n) hence IH: "\(a::'a compact_basis) b. \max (place a) (place b) < max (place x) (place y); a \ b\ \ ubasis_le (basis_emb a) (basis_emb b)" by simp show ?case proof (rule linorder_cases) assume "place x < place y" then have "rank x < rank y" using `x \ y` by (rule rank_place_mono) with `place x < place y` show ?case apply (case_tac "y = compact_bot", simp) apply (simp add: basis_emb.simps [of y]) apply (rule ubasis_le_trans [OF _ ubasis_le_lower [OF fin2]]) apply (rule IH) apply (simp add: less_max_iff_disj) apply (erule place_sub_less) apply (erule rank_less_imp_below_sub [OF `x \ y`]) done next assume "place x = place y" hence "x = y" by (rule place_eqD) thus ?case by (simp add: ubasis_le_refl) next assume "place x > place y" with `x \ y` show ?case apply (case_tac "x = compact_bot", simp add: ubasis_le_minimal) apply (simp add: basis_emb.simps [of x]) apply (rule ubasis_le_upper [OF fin2], simp) apply (rule IH) apply (simp add: less_max_iff_disj) apply (erule place_sub_less) apply (erule rev_trans_less) apply (rule sub_below) done qed qed lemma inj_basis_emb: "inj basis_emb" apply (rule inj_onI) apply (case_tac "x = compact_bot") apply (case_tac [!] "y = compact_bot") apply simp apply (simp add: basis_emb.simps) apply (simp add: basis_emb.simps) apply (simp add: basis_emb.simps) apply (simp add: fin2 inj_eq [OF inj_place]) done definition basis_prj :: "ubasis \ 'a compact_basis" where "basis_prj x = inv basis_emb (ubasis_until (\x. x \ range (basis_emb :: 'a compact_basis \ ubasis)) x)" lemma basis_prj_basis_emb: "\x. basis_prj (basis_emb x) = x" unfolding basis_prj_def apply (subst ubasis_until_same) apply (rule rangeI) apply (rule inv_f_f) apply (rule inj_basis_emb) done lemma basis_prj_node: "\finite S; node i a S \ range (basis_emb :: 'a compact_basis \ nat)\ \ basis_prj (node i a S) = (basis_prj a :: 'a compact_basis)" unfolding basis_prj_def by simp lemma basis_prj_0: "basis_prj 0 = compact_bot" apply (subst basis_emb_compact_bot [symmetric]) apply (rule basis_prj_basis_emb) done lemma node_eq_basis_emb_iff: "finite S \ node i a S = basis_emb x \ x \ compact_bot \ i = place x \ a = basis_emb (sub x) \ S = basis_emb ` {y. place y < place x \ x \ y}" apply (cases "x = compact_bot", simp) apply (simp add: basis_emb.simps [of x]) apply (simp add: fin2) done lemma basis_prj_mono: "ubasis_le a b \ basis_prj a \ basis_prj b" proof (induct a b rule: ubasis_le.induct) case (ubasis_le_refl a) show ?case by (rule refl_less) next case (ubasis_le_trans a b c) thus ?case by - (rule trans_less) next case (ubasis_le_lower S a i) thus ?case apply (cases "node i a S \ range (basis_emb :: 'a compact_basis \ nat)") apply (erule rangeE, rename_tac x) apply (simp add: basis_prj_basis_emb) apply (simp add: node_eq_basis_emb_iff) apply (simp add: basis_prj_basis_emb) apply (rule sub_below) apply (simp add: basis_prj_node) done next case (ubasis_le_upper S b a i) thus ?case apply (cases "node i a S \ range (basis_emb :: 'a compact_basis \ nat)") apply (erule rangeE, rename_tac x) apply (simp add: basis_prj_basis_emb) apply (clarsimp simp add: node_eq_basis_emb_iff) apply (simp add: basis_prj_basis_emb) apply (simp add: basis_prj_node) done qed lemma basis_emb_prj_less: "ubasis_le (basis_emb (basis_prj x)) x" unfolding basis_prj_def apply (subst f_inv_f [where f=basis_emb]) apply (rule ubasis_until) apply (rule range_eqI [where x=compact_bot]) apply simp apply (rule ubasis_until_less) done hide (open) const node choose choose_pos place sub subsubsection {* EP-pair from any bifinite domain into @{typ udom} *} definition udom_emb :: "'a::bifinite \ udom" where "udom_emb = compact_basis.basis_fun (\x. udom_principal (basis_emb x))" definition udom_prj :: "udom \ 'a::bifinite" where "udom_prj = udom.basis_fun (\x. Rep_compact_basis (basis_prj x))" lemma udom_emb_principal: "udom_emb\(Rep_compact_basis x) = udom_principal (basis_emb x)" unfolding udom_emb_def apply (rule compact_basis.basis_fun_principal) apply (rule udom.principal_mono) apply (erule basis_emb_mono) done lemma udom_prj_principal: "udom_prj\(udom_principal x) = Rep_compact_basis (basis_prj x)" unfolding udom_prj_def apply (rule udom.basis_fun_principal) apply (rule compact_basis.principal_mono) apply (erule basis_prj_mono) done lemma ep_pair_udom: "ep_pair udom_emb udom_prj" apply default apply (rule compact_basis.principal_induct, simp) apply (simp add: udom_emb_principal udom_prj_principal) apply (simp add: basis_prj_basis_emb) apply (rule udom.principal_induct, simp) apply (simp add: udom_emb_principal udom_prj_principal) apply (rule basis_emb_prj_less) done end diff --git a/src/HOLCF/UpperPD.thy b/src/HOLCF/UpperPD.thy --- a/src/HOLCF/UpperPD.thy +++ b/src/HOLCF/UpperPD.thy @@ -1,489 +1,489 @@ (* Title: HOLCF/UpperPD.thy Author: Brian Huffman *) header {* Upper powerdomain *} theory UpperPD imports CompactBasis begin subsection {* Basis preorder *} definition upper_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix "\\" 50) where "upper_le = (\u v. \y\Rep_pd_basis v. \x\Rep_pd_basis u. x \ y)" lemma upper_le_refl [simp]: "t \\ t" unfolding upper_le_def by fast lemma upper_le_trans: "\t \\ u; u \\ v\ \ t \\ v" unfolding upper_le_def apply (rule ballI) apply (drule (1) bspec, erule bexE) apply (drule (1) bspec, erule bexE) apply (erule rev_bexI) apply (erule (1) trans_less) done -interpretation upper_le!: preorder upper_le +interpretation upper_le: preorder upper_le by (rule preorder.intro, rule upper_le_refl, rule upper_le_trans) lemma upper_le_minimal [simp]: "PDUnit compact_bot \\ t" unfolding upper_le_def Rep_PDUnit by simp lemma PDUnit_upper_mono: "x \ y \ PDUnit x \\ PDUnit y" unfolding upper_le_def Rep_PDUnit by simp lemma PDPlus_upper_mono: "\s \\ t; u \\ v\ \ PDPlus s u \\ PDPlus t v" unfolding upper_le_def Rep_PDPlus by fast lemma PDPlus_upper_less: "PDPlus t u \\ t" unfolding upper_le_def Rep_PDPlus by fast lemma upper_le_PDUnit_PDUnit_iff [simp]: "(PDUnit a \\ PDUnit b) = a \ b" unfolding upper_le_def Rep_PDUnit by fast lemma upper_le_PDPlus_PDUnit_iff: "(PDPlus t u \\ PDUnit a) = (t \\ PDUnit a \ u \\ PDUnit a)" unfolding upper_le_def Rep_PDPlus Rep_PDUnit by fast lemma upper_le_PDPlus_iff: "(t \\ PDPlus u v) = (t \\ u \ t \\ v)" unfolding upper_le_def Rep_PDPlus by fast lemma upper_le_induct [induct set: upper_le]: assumes le: "t \\ u" assumes 1: "\a b. a \ b \ P (PDUnit a) (PDUnit b)" assumes 2: "\t u a. P t (PDUnit a) \ P (PDPlus t u) (PDUnit a)" assumes 3: "\t u v. \P t u; P t v\ \ P t (PDPlus u v)" shows "P t u" using le apply (induct u arbitrary: t rule: pd_basis_induct) apply (erule rev_mp) apply (induct_tac t rule: pd_basis_induct) apply (simp add: 1) apply (simp add: upper_le_PDPlus_PDUnit_iff) apply (simp add: 2) apply (subst PDPlus_commute) apply (simp add: 2) apply (simp add: upper_le_PDPlus_iff 3) done lemma pd_take_upper_chain: "pd_take n t \\ pd_take (Suc n) t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_basis.take_chain) apply (simp add: PDPlus_upper_mono) done lemma pd_take_upper_le: "pd_take i t \\ t" apply (induct t rule: pd_basis_induct) apply (simp add: compact_basis.take_less) apply (simp add: PDPlus_upper_mono) done lemma pd_take_upper_mono: "t \\ u \ pd_take n t \\ pd_take n u" apply (erule upper_le_induct) apply (simp add: compact_basis.take_mono) apply (simp add: upper_le_PDPlus_PDUnit_iff) apply (simp add: upper_le_PDPlus_iff) done subsection {* Type definition *} typedef (open) 'a upper_pd = "{S::'a pd_basis set. upper_le.ideal S}" by (fast intro: upper_le.ideal_principal) instantiation upper_pd :: (profinite) sq_ord begin definition "x \ y \ Rep_upper_pd x \ Rep_upper_pd y" instance .. end instance upper_pd :: (profinite) po by (rule upper_le.typedef_ideal_po [OF type_definition_upper_pd sq_le_upper_pd_def]) instance upper_pd :: (profinite) cpo by (rule upper_le.typedef_ideal_cpo [OF type_definition_upper_pd sq_le_upper_pd_def]) lemma Rep_upper_pd_lub: "chain Y \ Rep_upper_pd (\i. Y i) = (\i. Rep_upper_pd (Y i))" by (rule upper_le.typedef_ideal_rep_contlub [OF type_definition_upper_pd sq_le_upper_pd_def]) lemma ideal_Rep_upper_pd: "upper_le.ideal (Rep_upper_pd xs)" by (rule Rep_upper_pd [unfolded mem_Collect_eq]) definition upper_principal :: "'a pd_basis \ 'a upper_pd" where "upper_principal t = Abs_upper_pd {u. u \\ t}" lemma Rep_upper_principal: "Rep_upper_pd (upper_principal t) = {u. u \\ t}" unfolding upper_principal_def by (simp add: Abs_upper_pd_inverse upper_le.ideal_principal) -interpretation upper_pd!: +interpretation upper_pd: ideal_completion upper_le pd_take upper_principal Rep_upper_pd apply unfold_locales apply (rule pd_take_upper_le) apply (rule pd_take_idem) apply (erule pd_take_upper_mono) apply (rule pd_take_upper_chain) apply (rule finite_range_pd_take) apply (rule pd_take_covers) apply (rule ideal_Rep_upper_pd) apply (erule Rep_upper_pd_lub) apply (rule Rep_upper_principal) apply (simp only: sq_le_upper_pd_def) done text {* Upper powerdomain is pointed *} lemma upper_pd_minimal: "upper_principal (PDUnit compact_bot) \ ys" by (induct ys rule: upper_pd.principal_induct, simp, simp) instance upper_pd :: (bifinite) pcpo by intro_classes (fast intro: upper_pd_minimal) lemma inst_upper_pd_pcpo: "\ = upper_principal (PDUnit compact_bot)" by (rule upper_pd_minimal [THEN UU_I, symmetric]) text {* Upper powerdomain is profinite *} instantiation upper_pd :: (profinite) profinite begin definition approx_upper_pd_def: "approx = upper_pd.completion_approx" instance apply (intro_classes, unfold approx_upper_pd_def) apply (rule upper_pd.chain_completion_approx) apply (rule upper_pd.lub_completion_approx) apply (rule upper_pd.completion_approx_idem) apply (rule upper_pd.finite_fixes_completion_approx) done end instance upper_pd :: (bifinite) bifinite .. lemma approx_upper_principal [simp]: "approx n\(upper_principal t) = upper_principal (pd_take n t)" unfolding approx_upper_pd_def by (rule upper_pd.completion_approx_principal) lemma approx_eq_upper_principal: "\t\Rep_upper_pd xs. approx n\xs = upper_principal (pd_take n t)" unfolding approx_upper_pd_def by (rule upper_pd.completion_approx_eq_principal) subsection {* Monadic unit and plus *} definition upper_unit :: "'a \ 'a upper_pd" where "upper_unit = compact_basis.basis_fun (\a. upper_principal (PDUnit a))" definition upper_plus :: "'a upper_pd \ 'a upper_pd \ 'a upper_pd" where "upper_plus = upper_pd.basis_fun (\t. upper_pd.basis_fun (\u. upper_principal (PDPlus t u)))" abbreviation upper_add :: "'a upper_pd \ 'a upper_pd \ 'a upper_pd" (infixl "+\" 65) where "xs +\ ys == upper_plus\xs\ys" syntax "_upper_pd" :: "args \ 'a upper_pd" ("{_}\") translations "{x,xs}\" == "{x}\ +\ {xs}\" "{x}\" == "CONST upper_unit\x" lemma upper_unit_Rep_compact_basis [simp]: "{Rep_compact_basis a}\ = upper_principal (PDUnit a)" unfolding upper_unit_def by (simp add: compact_basis.basis_fun_principal PDUnit_upper_mono) lemma upper_plus_principal [simp]: "upper_principal t +\ upper_principal u = upper_principal (PDPlus t u)" unfolding upper_plus_def by (simp add: upper_pd.basis_fun_principal upper_pd.basis_fun_mono PDPlus_upper_mono) lemma approx_upper_unit [simp]: "approx n\{x}\ = {approx n\x}\" apply (induct x rule: compact_basis.principal_induct, simp) apply (simp add: approx_Rep_compact_basis) done lemma approx_upper_plus [simp]: "approx n\(xs +\ ys) = (approx n\xs) +\ (approx n\ys)" by (induct xs ys rule: upper_pd.principal_induct2, simp, simp, simp) lemma upper_plus_assoc: "(xs +\ ys) +\ zs = xs +\ (ys +\ zs)" apply (induct xs ys arbitrary: zs rule: upper_pd.principal_induct2, simp, simp) apply (rule_tac x=zs in upper_pd.principal_induct, simp) apply (simp add: PDPlus_assoc) done lemma upper_plus_commute: "xs +\ ys = ys +\ xs" apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_commute) done lemma upper_plus_absorb [simp]: "xs +\ xs = xs" apply (induct xs rule: upper_pd.principal_induct, simp) apply (simp add: PDPlus_absorb) done lemma upper_plus_left_commute: "xs +\ (ys +\ zs) = ys +\ (xs +\ zs)" by (rule mk_left_commute [of "op +\", OF upper_plus_assoc upper_plus_commute]) lemma upper_plus_left_absorb [simp]: "xs +\ (xs +\ ys) = xs +\ ys" by (simp only: upper_plus_assoc [symmetric] upper_plus_absorb) text {* Useful for @{text "simp add: upper_plus_ac"} *} lemmas upper_plus_ac = upper_plus_assoc upper_plus_commute upper_plus_left_commute text {* Useful for @{text "simp only: upper_plus_aci"} *} lemmas upper_plus_aci = upper_plus_ac upper_plus_absorb upper_plus_left_absorb lemma upper_plus_less1: "xs +\ ys \ xs" apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp) apply (simp add: PDPlus_upper_less) done lemma upper_plus_less2: "xs +\ ys \ ys" by (subst upper_plus_commute, rule upper_plus_less1) lemma upper_plus_greatest: "\xs \ ys; xs \ zs\ \ xs \ ys +\ zs" apply (subst upper_plus_absorb [of xs, symmetric]) apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) done lemma upper_less_plus_iff: "xs \ ys +\ zs \ xs \ ys \ xs \ zs" apply safe apply (erule trans_less [OF _ upper_plus_less1]) apply (erule trans_less [OF _ upper_plus_less2]) apply (erule (1) upper_plus_greatest) done lemma upper_plus_less_unit_iff: "xs +\ ys \ {z}\ \ xs \ {z}\ \ ys \ {z}\" apply (rule iffI) apply (subgoal_tac "adm (\f. f\xs \ f\{z}\ \ f\ys \ f\{z}\)") apply (drule admD, rule chain_approx) apply (drule_tac f="approx i" in monofun_cfun_arg) apply (cut_tac x="approx i\xs" in upper_pd.compact_imp_principal, simp) apply (cut_tac x="approx i\ys" in upper_pd.compact_imp_principal, simp) apply (cut_tac x="approx i\z" in compact_basis.compact_imp_principal, simp) apply (clarify, simp add: upper_le_PDPlus_PDUnit_iff) apply simp apply simp apply (erule disjE) apply (erule trans_less [OF upper_plus_less1]) apply (erule trans_less [OF upper_plus_less2]) done lemma upper_unit_less_iff [simp]: "{x}\ \ {y}\ \ x \ y" apply (rule iffI) apply (rule profinite_less_ext) apply (drule_tac f="approx i" in monofun_cfun_arg, simp) apply (cut_tac x="approx i\x" in compact_basis.compact_imp_principal, simp) apply (cut_tac x="approx i\y" in compact_basis.compact_imp_principal, simp) apply clarsimp apply (erule monofun_cfun_arg) done lemmas upper_pd_less_simps = upper_unit_less_iff upper_less_plus_iff upper_plus_less_unit_iff lemma upper_unit_eq_iff [simp]: "{x}\ = {y}\ \ x = y" unfolding po_eq_conv by simp lemma upper_unit_strict [simp]: "{\}\ = \" unfolding inst_upper_pd_pcpo Rep_compact_bot [symmetric] by simp lemma upper_plus_strict1 [simp]: "\ +\ ys = \" by (rule UU_I, rule upper_plus_less1) lemma upper_plus_strict2 [simp]: "xs +\ \ = \" by (rule UU_I, rule upper_plus_less2) lemma upper_unit_strict_iff [simp]: "{x}\ = \ \ x = \" unfolding upper_unit_strict [symmetric] by (rule upper_unit_eq_iff) lemma upper_plus_strict_iff [simp]: "xs +\ ys = \ \ xs = \ \ ys = \" apply (rule iffI) apply (erule rev_mp) apply (rule upper_pd.principal_induct2 [where x=xs and y=ys], simp, simp) apply (simp add: inst_upper_pd_pcpo upper_pd.principal_eq_iff upper_le_PDPlus_PDUnit_iff) apply auto done lemma compact_upper_unit_iff [simp]: "compact {x}\ \ compact x" unfolding profinite_compact_iff by simp lemma compact_upper_plus [simp]: "\compact xs; compact ys\ \ compact (xs +\ ys)" by (auto dest!: upper_pd.compact_imp_principal) subsection {* Induction rules *} lemma upper_pd_induct1: assumes P: "adm P" assumes unit: "\x. P {x}\" assumes insert: "\x ys. \P {x}\; P ys\ \ P ({x}\ +\ ys)" shows "P (xs::'a upper_pd)" apply (induct xs rule: upper_pd.principal_induct, rule P) apply (induct_tac a rule: pd_basis_induct1) apply (simp only: upper_unit_Rep_compact_basis [symmetric]) apply (rule unit) apply (simp only: upper_unit_Rep_compact_basis [symmetric] upper_plus_principal [symmetric]) apply (erule insert [OF unit]) done lemma upper_pd_induct: assumes P: "adm P" assumes unit: "\x. P {x}\" assumes plus: "\xs ys. \P xs; P ys\ \ P (xs +\ ys)" shows "P (xs::'a upper_pd)" apply (induct xs rule: upper_pd.principal_induct, rule P) apply (induct_tac a rule: pd_basis_induct) apply (simp only: upper_unit_Rep_compact_basis [symmetric] unit) apply (simp only: upper_plus_principal [symmetric] plus) done subsection {* Monadic bind *} definition upper_bind_basis :: "'a pd_basis \ ('a \ 'b upper_pd) \ 'b upper_pd" where "upper_bind_basis = fold_pd (\a. \ f. f\(Rep_compact_basis a)) (\x y. \ f. x\f +\ y\f)" lemma ACI_upper_bind: "ab_semigroup_idem_mult (\x y. \ f. x\f +\ y\f)" apply unfold_locales apply (simp add: upper_plus_assoc) apply (simp add: upper_plus_commute) apply (simp add: eta_cfun) done lemma upper_bind_basis_simps [simp]: "upper_bind_basis (PDUnit a) = (\ f. f\(Rep_compact_basis a))" "upper_bind_basis (PDPlus t u) = (\ f. upper_bind_basis t\f +\ upper_bind_basis u\f)" unfolding upper_bind_basis_def apply - apply (rule fold_pd_PDUnit [OF ACI_upper_bind]) apply (rule fold_pd_PDPlus [OF ACI_upper_bind]) done lemma upper_bind_basis_mono: "t \\ u \ upper_bind_basis t \ upper_bind_basis u" unfolding expand_cfun_less apply (erule upper_le_induct, safe) apply (simp add: monofun_cfun) apply (simp add: trans_less [OF upper_plus_less1]) apply (simp add: upper_less_plus_iff) done definition upper_bind :: "'a upper_pd \ ('a \ 'b upper_pd) \ 'b upper_pd" where "upper_bind = upper_pd.basis_fun upper_bind_basis" lemma upper_bind_principal [simp]: "upper_bind\(upper_principal t) = upper_bind_basis t" unfolding upper_bind_def apply (rule upper_pd.basis_fun_principal) apply (erule upper_bind_basis_mono) done lemma upper_bind_unit [simp]: "upper_bind\{x}\\f = f\x" by (induct x rule: compact_basis.principal_induct, simp, simp) lemma upper_bind_plus [simp]: "upper_bind\(xs +\ ys)\f = upper_bind\xs\f +\ upper_bind\ys\f" by (induct xs ys rule: upper_pd.principal_induct2, simp, simp, simp) lemma upper_bind_strict [simp]: "upper_bind\\\f = f\\" unfolding upper_unit_strict [symmetric] by (rule upper_bind_unit) subsection {* Map and join *} definition upper_map :: "('a \ 'b) \ 'a upper_pd \ 'b upper_pd" where "upper_map = (\ f xs. upper_bind\xs\(\ x. {f\x}\))" definition upper_join :: "'a upper_pd upper_pd \ 'a upper_pd" where "upper_join = (\ xss. upper_bind\xss\(\ xs. xs))" lemma upper_map_unit [simp]: "upper_map\f\{x}\ = {f\x}\" unfolding upper_map_def by simp lemma upper_map_plus [simp]: "upper_map\f\(xs +\ ys) = upper_map\f\xs +\ upper_map\f\ys" unfolding upper_map_def by simp lemma upper_join_unit [simp]: "upper_join\{xs}\ = xs" unfolding upper_join_def by simp lemma upper_join_plus [simp]: "upper_join\(xss +\ yss) = upper_join\xss +\ upper_join\yss" unfolding upper_join_def by simp lemma upper_map_ident: "upper_map\(\ x. x)\xs = xs" by (induct xs rule: upper_pd_induct, simp_all) lemma upper_map_map: "upper_map\f\(upper_map\g\xs) = upper_map\(\ x. f\(g\x))\xs" by (induct xs rule: upper_pd_induct, simp_all) lemma upper_join_map_unit: "upper_join\(upper_map\upper_unit\xs) = xs" by (induct xs rule: upper_pd_induct, simp_all) lemma upper_join_map_join: "upper_join\(upper_map\upper_join\xsss) = upper_join\(upper_join\xsss)" by (induct xsss rule: upper_pd_induct, simp_all) lemma upper_join_map_map: "upper_join\(upper_map\(upper_map\f)\xss) = upper_map\f\(upper_join\xss)" by (induct xss rule: upper_pd_induct, simp_all) lemma upper_map_approx: "upper_map\(approx n)\xs = approx n\xs" by (induct xs rule: upper_pd_induct, simp_all) end diff --git a/src/ZF/Constructible/L_axioms.thy b/src/ZF/Constructible/L_axioms.thy --- a/src/ZF/Constructible/L_axioms.thy +++ b/src/ZF/Constructible/L_axioms.thy @@ -1,1392 +1,1392 @@ (* Title: ZF/Constructible/L_axioms.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) header {* The ZF Axioms (Except Separation) in L *} theory L_axioms imports Formula Relative Reflection MetaExists begin text {* The class L satisfies the premises of locale @{text M_trivial} *} lemma transL: "[| y\x; L(x) |] ==> L(y)" apply (insert Transset_Lset) apply (simp add: Transset_def L_def, blast) done lemma nonempty: "L(0)" apply (simp add: L_def) apply (blast intro: zero_in_Lset) done theorem upair_ax: "upair_ax(L)" apply (simp add: upair_ax_def upair_def, clarify) apply (rule_tac x="{x,y}" in rexI) apply (simp_all add: doubleton_in_L) done theorem Union_ax: "Union_ax(L)" apply (simp add: Union_ax_def big_union_def, clarify) apply (rule_tac x="Union(x)" in rexI) apply (simp_all add: Union_in_L, auto) apply (blast intro: transL) done theorem power_ax: "power_ax(L)" apply (simp add: power_ax_def powerset_def Relative.subset_def, clarify) apply (rule_tac x="{y \ Pow(x). L(y)}" in rexI) apply (simp_all add: LPow_in_L, auto) apply (blast intro: transL) done text{*We don't actually need @{term L} to satisfy the foundation axiom.*} theorem foundation_ax: "foundation_ax(L)" apply (simp add: foundation_ax_def) apply (rule rallI) apply (cut_tac A=x in foundation) apply (blast intro: transL) done subsection{*For L to satisfy Replacement *} (*Can't move these to Formula unless the definition of univalent is moved there too!*) lemma LReplace_in_Lset: "[|X \ Lset(i); univalent(L,X,Q); Ord(i)|] ==> \j. Ord(j) & Replace(X, %x y. Q(x,y) & L(y)) \ Lset(j)" apply (rule_tac x="\y \ Replace(X, %x y. Q(x,y) & L(y)). succ(lrank(y))" in exI) apply simp apply clarify apply (rule_tac a=x in UN_I) apply (simp_all add: Replace_iff univalent_def) apply (blast dest: transL L_I) done lemma LReplace_in_L: "[|L(X); univalent(L,X,Q)|] ==> \Y. L(Y) & Replace(X, %x y. Q(x,y) & L(y)) \ Y" apply (drule L_D, clarify) apply (drule LReplace_in_Lset, assumption+) apply (blast intro: L_I Lset_in_Lset_succ) done theorem replacement: "replacement(L,P)" apply (simp add: replacement_def, clarify) apply (frule LReplace_in_L, assumption+, clarify) apply (rule_tac x=Y in rexI) apply (simp_all add: Replace_iff univalent_def, blast) done subsection{*Instantiating the locale @{text M_trivial}*} text{*No instances of Separation yet.*} lemma Lset_mono_le: "mono_le_subset(Lset)" by (simp add: mono_le_subset_def le_imp_subset Lset_mono) lemma Lset_cont: "cont_Ord(Lset)" by (simp add: cont_Ord_def Limit_Lset_eq OUnion_def Limit_is_Ord) lemmas L_nat = Ord_in_L [OF Ord_nat] theorem M_trivial_L: "PROP M_trivial(L)" apply (rule M_trivial.intro) apply (erule (1) transL) apply (rule upair_ax) apply (rule Union_ax) apply (rule power_ax) apply (rule replacement) apply (rule L_nat) done -interpretation L: M_trivial L by (rule M_trivial_L) +interpretation L?: M_trivial L by (rule M_trivial_L) (* Replaces the following declarations... lemmas rall_abs = M_trivial.rall_abs [OF M_trivial_L] and rex_abs = M_trivial.rex_abs [OF M_trivial_L] ... declare rall_abs [simp] declare rex_abs [simp] ...and dozens of similar ones. *) subsection{*Instantiation of the locale @{text reflection}*} text{*instances of locale constants*} definition L_F0 :: "[i=>o,i] => i" where "L_F0(P,y) == \ b. (\z. L(z) \ P()) --> (\z\Lset(b). P())" definition L_FF :: "[i=>o,i] => i" where "L_FF(P) == \a. \y\Lset(a). L_F0(P,y)" definition L_ClEx :: "[i=>o,i] => o" where "L_ClEx(P) == \a. Limit(a) \ normalize(L_FF(P),a) = a" text{*We must use the meta-existential quantifier; otherwise the reflection terms become enormous!*} definition L_Reflects :: "[i=>o,[i,i]=>o] => prop" ("(3REFLECTS/ [_,/ _])") where "REFLECTS[P,Q] == (??Cl. Closed_Unbounded(Cl) & (\a. Cl(a) --> (\x \ Lset(a). P(x) <-> Q(a,x))))" theorem Triv_reflection: "REFLECTS[P, \a x. P(x)]" apply (simp add: L_Reflects_def) apply (rule meta_exI) apply (rule Closed_Unbounded_Ord) done theorem Not_reflection: "REFLECTS[P,Q] ==> REFLECTS[\x. ~P(x), \a x. ~Q(a,x)]" apply (unfold L_Reflects_def) apply (erule meta_exE) apply (rule_tac x=Cl in meta_exI, simp) done theorem And_reflection: "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] ==> REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" apply (unfold L_Reflects_def) apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) apply (simp add: Closed_Unbounded_Int, blast) done theorem Or_reflection: "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] ==> REFLECTS[\x. P(x) \ P'(x), \a x. Q(a,x) \ Q'(a,x)]" apply (unfold L_Reflects_def) apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) apply (simp add: Closed_Unbounded_Int, blast) done theorem Imp_reflection: "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] ==> REFLECTS[\x. P(x) --> P'(x), \a x. Q(a,x) --> Q'(a,x)]" apply (unfold L_Reflects_def) apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) apply (simp add: Closed_Unbounded_Int, blast) done theorem Iff_reflection: "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] ==> REFLECTS[\x. P(x) <-> P'(x), \a x. Q(a,x) <-> Q'(a,x)]" apply (unfold L_Reflects_def) apply (elim meta_exE) apply (rule_tac x="\a. Cl(a) \ Cla(a)" in meta_exI) apply (simp add: Closed_Unbounded_Int, blast) done lemma reflection_Lset: "reflection(Lset)" by (blast intro: reflection.intro Lset_mono_le Lset_cont Formula.Pair_in_LLimit)+ theorem Ex_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] ==> REFLECTS[\x. \z. L(z) \ P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) apply (elim meta_exE) apply (rule meta_exI) apply (erule reflection.Ex_reflection [OF reflection_Lset]) done theorem All_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] ==> REFLECTS[\x. \z. L(z) --> P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) apply (elim meta_exE) apply (rule meta_exI) apply (erule reflection.All_reflection [OF reflection_Lset]) done theorem Rex_reflection: "REFLECTS[ \x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] ==> REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" apply (unfold rex_def) apply (intro And_reflection Ex_reflection, assumption) done theorem Rall_reflection: "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] ==> REFLECTS[\x. \z[L]. P(x,z), \a x. \z\Lset(a). Q(a,x,z)]" apply (unfold rall_def) apply (intro Imp_reflection All_reflection, assumption) done text{*This version handles an alternative form of the bounded quantifier in the second argument of @{text REFLECTS}.*} theorem Rex_reflection': "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] ==> REFLECTS[\x. \z[L]. P(x,z), \a x. \z[##Lset(a)]. Q(a,x,z)]" apply (unfold setclass_def rex_def) apply (erule Rex_reflection [unfolded rex_def Bex_def]) done text{*As above.*} theorem Rall_reflection': "REFLECTS[\x. P(fst(x),snd(x)), \a x. Q(a,fst(x),snd(x))] ==> REFLECTS[\x. \z[L]. P(x,z), \a x. \z[##Lset(a)]. Q(a,x,z)]" apply (unfold setclass_def rall_def) apply (erule Rall_reflection [unfolded rall_def Ball_def]) done lemmas FOL_reflections = Triv_reflection Not_reflection And_reflection Or_reflection Imp_reflection Iff_reflection Ex_reflection All_reflection Rex_reflection Rall_reflection Rex_reflection' Rall_reflection' lemma ReflectsD: "[|REFLECTS[P,Q]; Ord(i)|] ==> \j. ix \ Lset(j). P(x) <-> Q(j,x))" apply (unfold L_Reflects_def Closed_Unbounded_def) apply (elim meta_exE, clarify) apply (blast dest!: UnboundedD) done lemma ReflectsE: "[| REFLECTS[P,Q]; Ord(i); !!j. [|ix \ Lset(j). P(x) <-> Q(j,x)|] ==> R |] ==> R" by (drule ReflectsD, assumption, blast) lemma Collect_mem_eq: "{x\A. x\B} = A \ B" by blast subsection{*Internalized Formulas for some Set-Theoretic Concepts*} subsubsection{*Some numbers to help write de Bruijn indices*} abbreviation digit3 :: i ("3") where "3 == succ(2)" abbreviation digit4 :: i ("4") where "4 == succ(3)" abbreviation digit5 :: i ("5") where "5 == succ(4)" abbreviation digit6 :: i ("6") where "6 == succ(5)" abbreviation digit7 :: i ("7") where "7 == succ(6)" abbreviation digit8 :: i ("8") where "8 == succ(7)" abbreviation digit9 :: i ("9") where "9 == succ(8)" subsubsection{*The Empty Set, Internalized*} definition empty_fm :: "i=>i" where "empty_fm(x) == Forall(Neg(Member(0,succ(x))))" lemma empty_type [TC]: "x \ nat ==> empty_fm(x) \ formula" by (simp add: empty_fm_def) lemma sats_empty_fm [simp]: "[| x \ nat; env \ list(A)|] ==> sats(A, empty_fm(x), env) <-> empty(##A, nth(x,env))" by (simp add: empty_fm_def empty_def) lemma empty_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> empty(##A, x) <-> sats(A, empty_fm(i), env)" by simp theorem empty_reflection: "REFLECTS[\x. empty(L,f(x)), \i x. empty(##Lset(i),f(x))]" apply (simp only: empty_def) apply (intro FOL_reflections) done text{*Not used. But maybe useful?*} lemma Transset_sats_empty_fm_eq_0: "[| n \ nat; env \ list(A); Transset(A)|] ==> sats(A, empty_fm(n), env) <-> nth(n,env) = 0" apply (simp add: empty_fm_def empty_def Transset_def, auto) apply (case_tac "n < length(env)") apply (frule nth_type, assumption+, blast) apply (simp_all add: not_lt_iff_le nth_eq_0) done subsubsection{*Unordered Pairs, Internalized*} definition upair_fm :: "[i,i,i]=>i" where "upair_fm(x,y,z) == And(Member(x,z), And(Member(y,z), Forall(Implies(Member(0,succ(z)), Or(Equal(0,succ(x)), Equal(0,succ(y)))))))" lemma upair_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> upair_fm(x,y,z) \ formula" by (simp add: upair_fm_def) lemma sats_upair_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, upair_fm(x,y,z), env) <-> upair(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: upair_fm_def upair_def) lemma upair_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> upair(##A, x, y, z) <-> sats(A, upair_fm(i,j,k), env)" by (simp add: sats_upair_fm) text{*Useful? At least it refers to "real" unordered pairs*} lemma sats_upair_fm2 [simp]: "[| x \ nat; y \ nat; z < length(env); env \ list(A); Transset(A)|] ==> sats(A, upair_fm(x,y,z), env) <-> nth(z,env) = {nth(x,env), nth(y,env)}" apply (frule lt_length_in_nat, assumption) apply (simp add: upair_fm_def Transset_def, auto) apply (blast intro: nth_type) done theorem upair_reflection: "REFLECTS[\x. upair(L,f(x),g(x),h(x)), \i x. upair(##Lset(i),f(x),g(x),h(x))]" apply (simp add: upair_def) apply (intro FOL_reflections) done subsubsection{*Ordered pairs, Internalized*} definition pair_fm :: "[i,i,i]=>i" where "pair_fm(x,y,z) == Exists(And(upair_fm(succ(x),succ(x),0), Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0), upair_fm(1,0,succ(succ(z)))))))" lemma pair_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> pair_fm(x,y,z) \ formula" by (simp add: pair_fm_def) lemma sats_pair_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, pair_fm(x,y,z), env) <-> pair(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: pair_fm_def pair_def) lemma pair_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> pair(##A, x, y, z) <-> sats(A, pair_fm(i,j,k), env)" by (simp add: sats_pair_fm) theorem pair_reflection: "REFLECTS[\x. pair(L,f(x),g(x),h(x)), \i x. pair(##Lset(i),f(x),g(x),h(x))]" apply (simp only: pair_def) apply (intro FOL_reflections upair_reflection) done subsubsection{*Binary Unions, Internalized*} definition union_fm :: "[i,i,i]=>i" where "union_fm(x,y,z) == Forall(Iff(Member(0,succ(z)), Or(Member(0,succ(x)),Member(0,succ(y)))))" lemma union_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> union_fm(x,y,z) \ formula" by (simp add: union_fm_def) lemma sats_union_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, union_fm(x,y,z), env) <-> union(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: union_fm_def union_def) lemma union_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> union(##A, x, y, z) <-> sats(A, union_fm(i,j,k), env)" by (simp add: sats_union_fm) theorem union_reflection: "REFLECTS[\x. union(L,f(x),g(x),h(x)), \i x. union(##Lset(i),f(x),g(x),h(x))]" apply (simp only: union_def) apply (intro FOL_reflections) done subsubsection{*Set ``Cons,'' Internalized*} definition cons_fm :: "[i,i,i]=>i" where "cons_fm(x,y,z) == Exists(And(upair_fm(succ(x),succ(x),0), union_fm(0,succ(y),succ(z))))" lemma cons_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> cons_fm(x,y,z) \ formula" by (simp add: cons_fm_def) lemma sats_cons_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, cons_fm(x,y,z), env) <-> is_cons(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: cons_fm_def is_cons_def) lemma cons_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> is_cons(##A, x, y, z) <-> sats(A, cons_fm(i,j,k), env)" by simp theorem cons_reflection: "REFLECTS[\x. is_cons(L,f(x),g(x),h(x)), \i x. is_cons(##Lset(i),f(x),g(x),h(x))]" apply (simp only: is_cons_def) apply (intro FOL_reflections upair_reflection union_reflection) done subsubsection{*Successor Function, Internalized*} definition succ_fm :: "[i,i]=>i" where "succ_fm(x,y) == cons_fm(x,x,y)" lemma succ_type [TC]: "[| x \ nat; y \ nat |] ==> succ_fm(x,y) \ formula" by (simp add: succ_fm_def) lemma sats_succ_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] ==> sats(A, succ_fm(x,y), env) <-> successor(##A, nth(x,env), nth(y,env))" by (simp add: succ_fm_def successor_def) lemma successor_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> successor(##A, x, y) <-> sats(A, succ_fm(i,j), env)" by simp theorem successor_reflection: "REFLECTS[\x. successor(L,f(x),g(x)), \i x. successor(##Lset(i),f(x),g(x))]" apply (simp only: successor_def) apply (intro cons_reflection) done subsubsection{*The Number 1, Internalized*} (* "number1(M,a) == (\x[M]. empty(M,x) & successor(M,x,a))" *) definition number1_fm :: "i=>i" where "number1_fm(a) == Exists(And(empty_fm(0), succ_fm(0,succ(a))))" lemma number1_type [TC]: "x \ nat ==> number1_fm(x) \ formula" by (simp add: number1_fm_def) lemma sats_number1_fm [simp]: "[| x \ nat; env \ list(A)|] ==> sats(A, number1_fm(x), env) <-> number1(##A, nth(x,env))" by (simp add: number1_fm_def number1_def) lemma number1_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> number1(##A, x) <-> sats(A, number1_fm(i), env)" by simp theorem number1_reflection: "REFLECTS[\x. number1(L,f(x)), \i x. number1(##Lset(i),f(x))]" apply (simp only: number1_def) apply (intro FOL_reflections empty_reflection successor_reflection) done subsubsection{*Big Union, Internalized*} (* "big_union(M,A,z) == \x[M]. x \ z <-> (\y[M]. y\A & x \ y)" *) definition big_union_fm :: "[i,i]=>i" where "big_union_fm(A,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(A))), Member(1,0)))))" lemma big_union_type [TC]: "[| x \ nat; y \ nat |] ==> big_union_fm(x,y) \ formula" by (simp add: big_union_fm_def) lemma sats_big_union_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] ==> sats(A, big_union_fm(x,y), env) <-> big_union(##A, nth(x,env), nth(y,env))" by (simp add: big_union_fm_def big_union_def) lemma big_union_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> big_union(##A, x, y) <-> sats(A, big_union_fm(i,j), env)" by simp theorem big_union_reflection: "REFLECTS[\x. big_union(L,f(x),g(x)), \i x. big_union(##Lset(i),f(x),g(x))]" apply (simp only: big_union_def) apply (intro FOL_reflections) done subsubsection{*Variants of Satisfaction Definitions for Ordinals, etc.*} text{*The @{text sats} theorems below are standard versions of the ones proved in theory @{text Formula}. They relate elements of type @{term formula} to relativized concepts such as @{term subset} or @{term ordinal} rather than to real concepts such as @{term Ord}. Now that we have instantiated the locale @{text M_trivial}, we no longer require the earlier versions.*} lemma sats_subset_fm': "[|x \ nat; y \ nat; env \ list(A)|] ==> sats(A, subset_fm(x,y), env) <-> subset(##A, nth(x,env), nth(y,env))" by (simp add: subset_fm_def Relative.subset_def) theorem subset_reflection: "REFLECTS[\x. subset(L,f(x),g(x)), \i x. subset(##Lset(i),f(x),g(x))]" apply (simp only: Relative.subset_def) apply (intro FOL_reflections) done lemma sats_transset_fm': "[|x \ nat; env \ list(A)|] ==> sats(A, transset_fm(x), env) <-> transitive_set(##A, nth(x,env))" by (simp add: sats_subset_fm' transset_fm_def transitive_set_def) theorem transitive_set_reflection: "REFLECTS[\x. transitive_set(L,f(x)), \i x. transitive_set(##Lset(i),f(x))]" apply (simp only: transitive_set_def) apply (intro FOL_reflections subset_reflection) done lemma sats_ordinal_fm': "[|x \ nat; env \ list(A)|] ==> sats(A, ordinal_fm(x), env) <-> ordinal(##A,nth(x,env))" by (simp add: sats_transset_fm' ordinal_fm_def ordinal_def) lemma ordinal_iff_sats: "[| nth(i,env) = x; i \ nat; env \ list(A)|] ==> ordinal(##A, x) <-> sats(A, ordinal_fm(i), env)" by (simp add: sats_ordinal_fm') theorem ordinal_reflection: "REFLECTS[\x. ordinal(L,f(x)), \i x. ordinal(##Lset(i),f(x))]" apply (simp only: ordinal_def) apply (intro FOL_reflections transitive_set_reflection) done subsubsection{*Membership Relation, Internalized*} definition Memrel_fm :: "[i,i]=>i" where "Memrel_fm(A,r) == Forall(Iff(Member(0,succ(r)), Exists(And(Member(0,succ(succ(A))), Exists(And(Member(0,succ(succ(succ(A)))), And(Member(1,0), pair_fm(1,0,2))))))))" lemma Memrel_type [TC]: "[| x \ nat; y \ nat |] ==> Memrel_fm(x,y) \ formula" by (simp add: Memrel_fm_def) lemma sats_Memrel_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] ==> sats(A, Memrel_fm(x,y), env) <-> membership(##A, nth(x,env), nth(y,env))" by (simp add: Memrel_fm_def membership_def) lemma Memrel_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> membership(##A, x, y) <-> sats(A, Memrel_fm(i,j), env)" by simp theorem membership_reflection: "REFLECTS[\x. membership(L,f(x),g(x)), \i x. membership(##Lset(i),f(x),g(x))]" apply (simp only: membership_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Predecessor Set, Internalized*} definition pred_set_fm :: "[i,i,i,i]=>i" where "pred_set_fm(A,x,r,B) == Forall(Iff(Member(0,succ(B)), Exists(And(Member(0,succ(succ(r))), And(Member(1,succ(succ(A))), pair_fm(1,succ(succ(x)),0))))))" lemma pred_set_type [TC]: "[| A \ nat; x \ nat; r \ nat; B \ nat |] ==> pred_set_fm(A,x,r,B) \ formula" by (simp add: pred_set_fm_def) lemma sats_pred_set_fm [simp]: "[| U \ nat; x \ nat; r \ nat; B \ nat; env \ list(A)|] ==> sats(A, pred_set_fm(U,x,r,B), env) <-> pred_set(##A, nth(U,env), nth(x,env), nth(r,env), nth(B,env))" by (simp add: pred_set_fm_def pred_set_def) lemma pred_set_iff_sats: "[| nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B; i \ nat; j \ nat; k \ nat; l \ nat; env \ list(A)|] ==> pred_set(##A,U,x,r,B) <-> sats(A, pred_set_fm(i,j,k,l), env)" by (simp add: sats_pred_set_fm) theorem pred_set_reflection: "REFLECTS[\x. pred_set(L,f(x),g(x),h(x),b(x)), \i x. pred_set(##Lset(i),f(x),g(x),h(x),b(x))]" apply (simp only: pred_set_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Domain of a Relation, Internalized*} (* "is_domain(M,r,z) == \x[M]. (x \ z <-> (\w[M]. w\r & (\y[M]. pair(M,x,y,w))))" *) definition domain_fm :: "[i,i]=>i" where "domain_fm(r,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(pair_fm(2,0,1))))))" lemma domain_type [TC]: "[| x \ nat; y \ nat |] ==> domain_fm(x,y) \ formula" by (simp add: domain_fm_def) lemma sats_domain_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] ==> sats(A, domain_fm(x,y), env) <-> is_domain(##A, nth(x,env), nth(y,env))" by (simp add: domain_fm_def is_domain_def) lemma domain_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> is_domain(##A, x, y) <-> sats(A, domain_fm(i,j), env)" by simp theorem domain_reflection: "REFLECTS[\x. is_domain(L,f(x),g(x)), \i x. is_domain(##Lset(i),f(x),g(x))]" apply (simp only: is_domain_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Range of a Relation, Internalized*} (* "is_range(M,r,z) == \y[M]. (y \ z <-> (\w[M]. w\r & (\x[M]. pair(M,x,y,w))))" *) definition range_fm :: "[i,i]=>i" where "range_fm(r,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(pair_fm(0,2,1))))))" lemma range_type [TC]: "[| x \ nat; y \ nat |] ==> range_fm(x,y) \ formula" by (simp add: range_fm_def) lemma sats_range_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] ==> sats(A, range_fm(x,y), env) <-> is_range(##A, nth(x,env), nth(y,env))" by (simp add: range_fm_def is_range_def) lemma range_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> is_range(##A, x, y) <-> sats(A, range_fm(i,j), env)" by simp theorem range_reflection: "REFLECTS[\x. is_range(L,f(x),g(x)), \i x. is_range(##Lset(i),f(x),g(x))]" apply (simp only: is_range_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Field of a Relation, Internalized*} (* "is_field(M,r,z) == \dr[M]. is_domain(M,r,dr) & (\rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" *) definition field_fm :: "[i,i]=>i" where "field_fm(r,z) == Exists(And(domain_fm(succ(r),0), Exists(And(range_fm(succ(succ(r)),0), union_fm(1,0,succ(succ(z)))))))" lemma field_type [TC]: "[| x \ nat; y \ nat |] ==> field_fm(x,y) \ formula" by (simp add: field_fm_def) lemma sats_field_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] ==> sats(A, field_fm(x,y), env) <-> is_field(##A, nth(x,env), nth(y,env))" by (simp add: field_fm_def is_field_def) lemma field_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> is_field(##A, x, y) <-> sats(A, field_fm(i,j), env)" by simp theorem field_reflection: "REFLECTS[\x. is_field(L,f(x),g(x)), \i x. is_field(##Lset(i),f(x),g(x))]" apply (simp only: is_field_def) apply (intro FOL_reflections domain_reflection range_reflection union_reflection) done subsubsection{*Image under a Relation, Internalized*} (* "image(M,r,A,z) == \y[M]. (y \ z <-> (\w[M]. w\r & (\x[M]. x\A & pair(M,x,y,w))))" *) definition image_fm :: "[i,i,i]=>i" where "image_fm(r,A,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(And(Member(0,succ(succ(succ(A)))), pair_fm(0,2,1)))))))" lemma image_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> image_fm(x,y,z) \ formula" by (simp add: image_fm_def) lemma sats_image_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, image_fm(x,y,z), env) <-> image(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: image_fm_def Relative.image_def) lemma image_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> image(##A, x, y, z) <-> sats(A, image_fm(i,j,k), env)" by (simp add: sats_image_fm) theorem image_reflection: "REFLECTS[\x. image(L,f(x),g(x),h(x)), \i x. image(##Lset(i),f(x),g(x),h(x))]" apply (simp only: Relative.image_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Pre-Image under a Relation, Internalized*} (* "pre_image(M,r,A,z) == \x[M]. x \ z <-> (\w[M]. w\r & (\y[M]. y\A & pair(M,x,y,w)))" *) definition pre_image_fm :: "[i,i,i]=>i" where "pre_image_fm(r,A,z) == Forall(Iff(Member(0,succ(z)), Exists(And(Member(0,succ(succ(r))), Exists(And(Member(0,succ(succ(succ(A)))), pair_fm(2,0,1)))))))" lemma pre_image_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> pre_image_fm(x,y,z) \ formula" by (simp add: pre_image_fm_def) lemma sats_pre_image_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, pre_image_fm(x,y,z), env) <-> pre_image(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: pre_image_fm_def Relative.pre_image_def) lemma pre_image_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> pre_image(##A, x, y, z) <-> sats(A, pre_image_fm(i,j,k), env)" by (simp add: sats_pre_image_fm) theorem pre_image_reflection: "REFLECTS[\x. pre_image(L,f(x),g(x),h(x)), \i x. pre_image(##Lset(i),f(x),g(x),h(x))]" apply (simp only: Relative.pre_image_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Function Application, Internalized*} (* "fun_apply(M,f,x,y) == (\xs[M]. \fxs[M]. upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))" *) definition fun_apply_fm :: "[i,i,i]=>i" where "fun_apply_fm(f,x,y) == Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1), And(image_fm(succ(succ(f)), 1, 0), big_union_fm(0,succ(succ(y)))))))" lemma fun_apply_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> fun_apply_fm(x,y,z) \ formula" by (simp add: fun_apply_fm_def) lemma sats_fun_apply_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, fun_apply_fm(x,y,z), env) <-> fun_apply(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: fun_apply_fm_def fun_apply_def) lemma fun_apply_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> fun_apply(##A, x, y, z) <-> sats(A, fun_apply_fm(i,j,k), env)" by simp theorem fun_apply_reflection: "REFLECTS[\x. fun_apply(L,f(x),g(x),h(x)), \i x. fun_apply(##Lset(i),f(x),g(x),h(x))]" apply (simp only: fun_apply_def) apply (intro FOL_reflections upair_reflection image_reflection big_union_reflection) done subsubsection{*The Concept of Relation, Internalized*} (* "is_relation(M,r) == (\z[M]. z\r --> (\x[M]. \y[M]. pair(M,x,y,z)))" *) definition relation_fm :: "i=>i" where "relation_fm(r) == Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))" lemma relation_type [TC]: "[| x \ nat |] ==> relation_fm(x) \ formula" by (simp add: relation_fm_def) lemma sats_relation_fm [simp]: "[| x \ nat; env \ list(A)|] ==> sats(A, relation_fm(x), env) <-> is_relation(##A, nth(x,env))" by (simp add: relation_fm_def is_relation_def) lemma relation_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> is_relation(##A, x) <-> sats(A, relation_fm(i), env)" by simp theorem is_relation_reflection: "REFLECTS[\x. is_relation(L,f(x)), \i x. is_relation(##Lset(i),f(x))]" apply (simp only: is_relation_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*The Concept of Function, Internalized*} (* "is_function(M,r) == \x[M]. \y[M]. \y'[M]. \p[M]. \p'[M]. pair(M,x,y,p) --> pair(M,x,y',p') --> p\r --> p'\r --> y=y'" *) definition function_fm :: "i=>i" where "function_fm(r) == Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,3,1), Implies(pair_fm(4,2,0), Implies(Member(1,r#+5), Implies(Member(0,r#+5), Equal(3,2))))))))))" lemma function_type [TC]: "[| x \ nat |] ==> function_fm(x) \ formula" by (simp add: function_fm_def) lemma sats_function_fm [simp]: "[| x \ nat; env \ list(A)|] ==> sats(A, function_fm(x), env) <-> is_function(##A, nth(x,env))" by (simp add: function_fm_def is_function_def) lemma is_function_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> is_function(##A, x) <-> sats(A, function_fm(i), env)" by simp theorem is_function_reflection: "REFLECTS[\x. is_function(L,f(x)), \i x. is_function(##Lset(i),f(x))]" apply (simp only: is_function_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Typed Functions, Internalized*} (* "typed_function(M,A,B,r) == is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) & (\u[M]. u\r --> (\x[M]. \y[M]. pair(M,x,y,u) --> y\B))" *) definition typed_function_fm :: "[i,i,i]=>i" where "typed_function_fm(A,B,r) == And(function_fm(r), And(relation_fm(r), And(domain_fm(r,A), Forall(Implies(Member(0,succ(r)), Forall(Forall(Implies(pair_fm(1,0,2),Member(0,B#+3)))))))))" lemma typed_function_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> typed_function_fm(x,y,z) \ formula" by (simp add: typed_function_fm_def) lemma sats_typed_function_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, typed_function_fm(x,y,z), env) <-> typed_function(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: typed_function_fm_def typed_function_def) lemma typed_function_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> typed_function(##A, x, y, z) <-> sats(A, typed_function_fm(i,j,k), env)" by simp lemmas function_reflections = empty_reflection number1_reflection upair_reflection pair_reflection union_reflection big_union_reflection cons_reflection successor_reflection fun_apply_reflection subset_reflection transitive_set_reflection membership_reflection pred_set_reflection domain_reflection range_reflection field_reflection image_reflection pre_image_reflection is_relation_reflection is_function_reflection lemmas function_iff_sats = empty_iff_sats number1_iff_sats upair_iff_sats pair_iff_sats union_iff_sats big_union_iff_sats cons_iff_sats successor_iff_sats fun_apply_iff_sats Memrel_iff_sats pred_set_iff_sats domain_iff_sats range_iff_sats field_iff_sats image_iff_sats pre_image_iff_sats relation_iff_sats is_function_iff_sats theorem typed_function_reflection: "REFLECTS[\x. typed_function(L,f(x),g(x),h(x)), \i x. typed_function(##Lset(i),f(x),g(x),h(x))]" apply (simp only: typed_function_def) apply (intro FOL_reflections function_reflections) done subsubsection{*Composition of Relations, Internalized*} (* "composition(M,r,s,t) == \p[M]. p \ t <-> (\x[M]. \y[M]. \z[M]. \xy[M]. \yz[M]. pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) & xy \ s & yz \ r)" *) definition composition_fm :: "[i,i,i]=>i" where "composition_fm(r,s,t) == Forall(Iff(Member(0,succ(t)), Exists(Exists(Exists(Exists(Exists( And(pair_fm(4,2,5), And(pair_fm(4,3,1), And(pair_fm(3,2,0), And(Member(1,s#+6), Member(0,r#+6))))))))))))" lemma composition_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> composition_fm(x,y,z) \ formula" by (simp add: composition_fm_def) lemma sats_composition_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, composition_fm(x,y,z), env) <-> composition(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: composition_fm_def composition_def) lemma composition_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> composition(##A, x, y, z) <-> sats(A, composition_fm(i,j,k), env)" by simp theorem composition_reflection: "REFLECTS[\x. composition(L,f(x),g(x),h(x)), \i x. composition(##Lset(i),f(x),g(x),h(x))]" apply (simp only: composition_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Injections, Internalized*} (* "injection(M,A,B,f) == typed_function(M,A,B,f) & (\x[M]. \x'[M]. \y[M]. \p[M]. \p'[M]. pair(M,x,y,p) --> pair(M,x',y,p') --> p\f --> p'\f --> x=x')" *) definition injection_fm :: "[i,i,i]=>i" where "injection_fm(A,B,f) == And(typed_function_fm(A,B,f), Forall(Forall(Forall(Forall(Forall( Implies(pair_fm(4,2,1), Implies(pair_fm(3,2,0), Implies(Member(1,f#+5), Implies(Member(0,f#+5), Equal(4,3)))))))))))" lemma injection_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> injection_fm(x,y,z) \ formula" by (simp add: injection_fm_def) lemma sats_injection_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, injection_fm(x,y,z), env) <-> injection(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: injection_fm_def injection_def) lemma injection_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> injection(##A, x, y, z) <-> sats(A, injection_fm(i,j,k), env)" by simp theorem injection_reflection: "REFLECTS[\x. injection(L,f(x),g(x),h(x)), \i x. injection(##Lset(i),f(x),g(x),h(x))]" apply (simp only: injection_def) apply (intro FOL_reflections function_reflections typed_function_reflection) done subsubsection{*Surjections, Internalized*} (* surjection :: "[i=>o,i,i,i] => o" "surjection(M,A,B,f) == typed_function(M,A,B,f) & (\y[M]. y\B --> (\x[M]. x\A & fun_apply(M,f,x,y)))" *) definition surjection_fm :: "[i,i,i]=>i" where "surjection_fm(A,B,f) == And(typed_function_fm(A,B,f), Forall(Implies(Member(0,succ(B)), Exists(And(Member(0,succ(succ(A))), fun_apply_fm(succ(succ(f)),0,1))))))" lemma surjection_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> surjection_fm(x,y,z) \ formula" by (simp add: surjection_fm_def) lemma sats_surjection_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, surjection_fm(x,y,z), env) <-> surjection(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: surjection_fm_def surjection_def) lemma surjection_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> surjection(##A, x, y, z) <-> sats(A, surjection_fm(i,j,k), env)" by simp theorem surjection_reflection: "REFLECTS[\x. surjection(L,f(x),g(x),h(x)), \i x. surjection(##Lset(i),f(x),g(x),h(x))]" apply (simp only: surjection_def) apply (intro FOL_reflections function_reflections typed_function_reflection) done subsubsection{*Bijections, Internalized*} (* bijection :: "[i=>o,i,i,i] => o" "bijection(M,A,B,f) == injection(M,A,B,f) & surjection(M,A,B,f)" *) definition bijection_fm :: "[i,i,i]=>i" where "bijection_fm(A,B,f) == And(injection_fm(A,B,f), surjection_fm(A,B,f))" lemma bijection_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> bijection_fm(x,y,z) \ formula" by (simp add: bijection_fm_def) lemma sats_bijection_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, bijection_fm(x,y,z), env) <-> bijection(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: bijection_fm_def bijection_def) lemma bijection_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> bijection(##A, x, y, z) <-> sats(A, bijection_fm(i,j,k), env)" by simp theorem bijection_reflection: "REFLECTS[\x. bijection(L,f(x),g(x),h(x)), \i x. bijection(##Lset(i),f(x),g(x),h(x))]" apply (simp only: bijection_def) apply (intro And_reflection injection_reflection surjection_reflection) done subsubsection{*Restriction of a Relation, Internalized*} (* "restriction(M,r,A,z) == \x[M]. x \ z <-> (x \ r & (\u[M]. u\A & (\v[M]. pair(M,u,v,x))))" *) definition restriction_fm :: "[i,i,i]=>i" where "restriction_fm(r,A,z) == Forall(Iff(Member(0,succ(z)), And(Member(0,succ(r)), Exists(And(Member(0,succ(succ(A))), Exists(pair_fm(1,0,2)))))))" lemma restriction_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> restriction_fm(x,y,z) \ formula" by (simp add: restriction_fm_def) lemma sats_restriction_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, restriction_fm(x,y,z), env) <-> restriction(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: restriction_fm_def restriction_def) lemma restriction_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> restriction(##A, x, y, z) <-> sats(A, restriction_fm(i,j,k), env)" by simp theorem restriction_reflection: "REFLECTS[\x. restriction(L,f(x),g(x),h(x)), \i x. restriction(##Lset(i),f(x),g(x),h(x))]" apply (simp only: restriction_def) apply (intro FOL_reflections pair_reflection) done subsubsection{*Order-Isomorphisms, Internalized*} (* order_isomorphism :: "[i=>o,i,i,i,i,i] => o" "order_isomorphism(M,A,r,B,s,f) == bijection(M,A,B,f) & (\x[M]. x\A --> (\y[M]. y\A --> (\p[M]. \fx[M]. \fy[M]. \q[M]. pair(M,x,y,p) --> fun_apply(M,f,x,fx) --> fun_apply(M,f,y,fy) --> pair(M,fx,fy,q) --> (p\r <-> q\s))))" *) definition order_isomorphism_fm :: "[i,i,i,i,i]=>i" where "order_isomorphism_fm(A,r,B,s,f) == And(bijection_fm(A,B,f), Forall(Implies(Member(0,succ(A)), Forall(Implies(Member(0,succ(succ(A))), Forall(Forall(Forall(Forall( Implies(pair_fm(5,4,3), Implies(fun_apply_fm(f#+6,5,2), Implies(fun_apply_fm(f#+6,4,1), Implies(pair_fm(2,1,0), Iff(Member(3,r#+6), Member(0,s#+6)))))))))))))))" lemma order_isomorphism_type [TC]: "[| A \ nat; r \ nat; B \ nat; s \ nat; f \ nat |] ==> order_isomorphism_fm(A,r,B,s,f) \ formula" by (simp add: order_isomorphism_fm_def) lemma sats_order_isomorphism_fm [simp]: "[| U \ nat; r \ nat; B \ nat; s \ nat; f \ nat; env \ list(A)|] ==> sats(A, order_isomorphism_fm(U,r,B,s,f), env) <-> order_isomorphism(##A, nth(U,env), nth(r,env), nth(B,env), nth(s,env), nth(f,env))" by (simp add: order_isomorphism_fm_def order_isomorphism_def) lemma order_isomorphism_iff_sats: "[| nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s; nth(k',env) = f; i \ nat; j \ nat; k \ nat; j' \ nat; k' \ nat; env \ list(A)|] ==> order_isomorphism(##A,U,r,B,s,f) <-> sats(A, order_isomorphism_fm(i,j,k,j',k'), env)" by simp theorem order_isomorphism_reflection: "REFLECTS[\x. order_isomorphism(L,f(x),g(x),h(x),g'(x),h'(x)), \i x. order_isomorphism(##Lset(i),f(x),g(x),h(x),g'(x),h'(x))]" apply (simp only: order_isomorphism_def) apply (intro FOL_reflections function_reflections bijection_reflection) done subsubsection{*Limit Ordinals, Internalized*} text{*A limit ordinal is a non-empty, successor-closed ordinal*} (* "limit_ordinal(M,a) == ordinal(M,a) & ~ empty(M,a) & (\x[M]. x\a --> (\y[M]. y\a & successor(M,x,y)))" *) definition limit_ordinal_fm :: "i=>i" where "limit_ordinal_fm(x) == And(ordinal_fm(x), And(Neg(empty_fm(x)), Forall(Implies(Member(0,succ(x)), Exists(And(Member(0,succ(succ(x))), succ_fm(1,0)))))))" lemma limit_ordinal_type [TC]: "x \ nat ==> limit_ordinal_fm(x) \ formula" by (simp add: limit_ordinal_fm_def) lemma sats_limit_ordinal_fm [simp]: "[| x \ nat; env \ list(A)|] ==> sats(A, limit_ordinal_fm(x), env) <-> limit_ordinal(##A, nth(x,env))" by (simp add: limit_ordinal_fm_def limit_ordinal_def sats_ordinal_fm') lemma limit_ordinal_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> limit_ordinal(##A, x) <-> sats(A, limit_ordinal_fm(i), env)" by simp theorem limit_ordinal_reflection: "REFLECTS[\x. limit_ordinal(L,f(x)), \i x. limit_ordinal(##Lset(i),f(x))]" apply (simp only: limit_ordinal_def) apply (intro FOL_reflections ordinal_reflection empty_reflection successor_reflection) done subsubsection{*Finite Ordinals: The Predicate ``Is A Natural Number''*} (* "finite_ordinal(M,a) == ordinal(M,a) & ~ limit_ordinal(M,a) & (\x[M]. x\a --> ~ limit_ordinal(M,x))" *) definition finite_ordinal_fm :: "i=>i" where "finite_ordinal_fm(x) == And(ordinal_fm(x), And(Neg(limit_ordinal_fm(x)), Forall(Implies(Member(0,succ(x)), Neg(limit_ordinal_fm(0))))))" lemma finite_ordinal_type [TC]: "x \ nat ==> finite_ordinal_fm(x) \ formula" by (simp add: finite_ordinal_fm_def) lemma sats_finite_ordinal_fm [simp]: "[| x \ nat; env \ list(A)|] ==> sats(A, finite_ordinal_fm(x), env) <-> finite_ordinal(##A, nth(x,env))" by (simp add: finite_ordinal_fm_def sats_ordinal_fm' finite_ordinal_def) lemma finite_ordinal_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> finite_ordinal(##A, x) <-> sats(A, finite_ordinal_fm(i), env)" by simp theorem finite_ordinal_reflection: "REFLECTS[\x. finite_ordinal(L,f(x)), \i x. finite_ordinal(##Lset(i),f(x))]" apply (simp only: finite_ordinal_def) apply (intro FOL_reflections ordinal_reflection limit_ordinal_reflection) done subsubsection{*Omega: The Set of Natural Numbers*} (* omega(M,a) == limit_ordinal(M,a) & (\x[M]. x\a --> ~ limit_ordinal(M,x)) *) definition omega_fm :: "i=>i" where "omega_fm(x) == And(limit_ordinal_fm(x), Forall(Implies(Member(0,succ(x)), Neg(limit_ordinal_fm(0)))))" lemma omega_type [TC]: "x \ nat ==> omega_fm(x) \ formula" by (simp add: omega_fm_def) lemma sats_omega_fm [simp]: "[| x \ nat; env \ list(A)|] ==> sats(A, omega_fm(x), env) <-> omega(##A, nth(x,env))" by (simp add: omega_fm_def omega_def) lemma omega_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; env \ list(A)|] ==> omega(##A, x) <-> sats(A, omega_fm(i), env)" by simp theorem omega_reflection: "REFLECTS[\x. omega(L,f(x)), \i x. omega(##Lset(i),f(x))]" apply (simp only: omega_def) apply (intro FOL_reflections limit_ordinal_reflection) done lemmas fun_plus_reflections = typed_function_reflection composition_reflection injection_reflection surjection_reflection bijection_reflection restriction_reflection order_isomorphism_reflection finite_ordinal_reflection ordinal_reflection limit_ordinal_reflection omega_reflection lemmas fun_plus_iff_sats = typed_function_iff_sats composition_iff_sats injection_iff_sats surjection_iff_sats bijection_iff_sats restriction_iff_sats order_isomorphism_iff_sats finite_ordinal_iff_sats ordinal_iff_sats limit_ordinal_iff_sats omega_iff_sats end diff --git a/src/ZF/Constructible/Rec_Separation.thy b/src/ZF/Constructible/Rec_Separation.thy --- a/src/ZF/Constructible/Rec_Separation.thy +++ b/src/ZF/Constructible/Rec_Separation.thy @@ -1,441 +1,441 @@ (* Title: ZF/Constructible/Rec_Separation.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) header {*Separation for Facts About Recursion*} theory Rec_Separation imports Separation Internalize begin text{*This theory proves all instances needed for locales @{text "M_trancl"} and @{text "M_datatypes"}*} lemma eq_succ_imp_lt: "[|i = succ(j); Ord(i)|] ==> jnnat[M]. \n[M]. \n'[M]. omega(M,nnat) & n\nnat & successor(M,n,n') & (\f[M]. typed_function(M,n',A,f) & (\x[M]. \y[M]. \zero[M]. pair(M,x,y,p) & empty(M,zero) & fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & (\j[M]. j\n --> (\fj[M]. \sj[M]. \fsj[M]. \ffp[M]. fun_apply(M,f,j,fj) & successor(M,j,sj) & fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \ r)))"*) definition rtran_closure_mem_fm :: "[i,i,i]=>i" where "rtran_closure_mem_fm(A,r,p) == Exists(Exists(Exists( And(omega_fm(2), And(Member(1,2), And(succ_fm(1,0), Exists(And(typed_function_fm(1, A#+4, 0), And(Exists(Exists(Exists( And(pair_fm(2,1,p#+7), And(empty_fm(0), And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))), Forall(Implies(Member(0,3), Exists(Exists(Exists(Exists( And(fun_apply_fm(5,4,3), And(succ_fm(4,2), And(fun_apply_fm(5,2,1), And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))" lemma rtran_closure_mem_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> rtran_closure_mem_fm(x,y,z) \ formula" by (simp add: rtran_closure_mem_fm_def) lemma sats_rtran_closure_mem_fm [simp]: "[| x \ nat; y \ nat; z \ nat; env \ list(A)|] ==> sats(A, rtran_closure_mem_fm(x,y,z), env) <-> rtran_closure_mem(##A, nth(x,env), nth(y,env), nth(z,env))" by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def) lemma rtran_closure_mem_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i \ nat; j \ nat; k \ nat; env \ list(A)|] ==> rtran_closure_mem(##A, x, y, z) <-> sats(A, rtran_closure_mem_fm(i,j,k), env)" by (simp add: sats_rtran_closure_mem_fm) lemma rtran_closure_mem_reflection: "REFLECTS[\x. rtran_closure_mem(L,f(x),g(x),h(x)), \i x. rtran_closure_mem(##Lset(i),f(x),g(x),h(x))]" apply (simp only: rtran_closure_mem_def) apply (intro FOL_reflections function_reflections fun_plus_reflections) done text{*Separation for @{term "rtrancl(r)"}.*} lemma rtrancl_separation: "[| L(r); L(A) |] ==> separation (L, rtran_closure_mem(L,A,r))" apply (rule gen_separation_multi [OF rtran_closure_mem_reflection, of "{r,A}"], auto) apply (rule_tac env="[r,A]" in DPow_LsetI) apply (rule rtran_closure_mem_iff_sats sep_rules | simp)+ done subsubsection{*Reflexive/Transitive Closure, Internalized*} (* "rtran_closure(M,r,s) == \A[M]. is_field(M,r,A) --> (\p[M]. p \ s <-> rtran_closure_mem(M,A,r,p))" *) definition rtran_closure_fm :: "[i,i]=>i" where "rtran_closure_fm(r,s) == Forall(Implies(field_fm(succ(r),0), Forall(Iff(Member(0,succ(succ(s))), rtran_closure_mem_fm(1,succ(succ(r)),0)))))" lemma rtran_closure_type [TC]: "[| x \ nat; y \ nat |] ==> rtran_closure_fm(x,y) \ formula" by (simp add: rtran_closure_fm_def) lemma sats_rtran_closure_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] ==> sats(A, rtran_closure_fm(x,y), env) <-> rtran_closure(##A, nth(x,env), nth(y,env))" by (simp add: rtran_closure_fm_def rtran_closure_def) lemma rtran_closure_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> rtran_closure(##A, x, y) <-> sats(A, rtran_closure_fm(i,j), env)" by simp theorem rtran_closure_reflection: "REFLECTS[\x. rtran_closure(L,f(x),g(x)), \i x. rtran_closure(##Lset(i),f(x),g(x))]" apply (simp only: rtran_closure_def) apply (intro FOL_reflections function_reflections rtran_closure_mem_reflection) done subsubsection{*Transitive Closure of a Relation, Internalized*} (* "tran_closure(M,r,t) == \s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *) definition tran_closure_fm :: "[i,i]=>i" where "tran_closure_fm(r,s) == Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))" lemma tran_closure_type [TC]: "[| x \ nat; y \ nat |] ==> tran_closure_fm(x,y) \ formula" by (simp add: tran_closure_fm_def) lemma sats_tran_closure_fm [simp]: "[| x \ nat; y \ nat; env \ list(A)|] ==> sats(A, tran_closure_fm(x,y), env) <-> tran_closure(##A, nth(x,env), nth(y,env))" by (simp add: tran_closure_fm_def tran_closure_def) lemma tran_closure_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; i \ nat; j \ nat; env \ list(A)|] ==> tran_closure(##A, x, y) <-> sats(A, tran_closure_fm(i,j), env)" by simp theorem tran_closure_reflection: "REFLECTS[\x. tran_closure(L,f(x),g(x)), \i x. tran_closure(##Lset(i),f(x),g(x))]" apply (simp only: tran_closure_def) apply (intro FOL_reflections function_reflections rtran_closure_reflection composition_reflection) done subsubsection{*Separation for the Proof of @{text "wellfounded_on_trancl"}*} lemma wellfounded_trancl_reflects: "REFLECTS[\x. \w[L]. \wx[L]. \rp[L]. w \ Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \ rp, \i x. \w \ Lset(i). \wx \ Lset(i). \rp \ Lset(i). w \ Z & pair(##Lset(i),w,x,wx) & tran_closure(##Lset(i),r,rp) & wx \ rp]" by (intro FOL_reflections function_reflections fun_plus_reflections tran_closure_reflection) lemma wellfounded_trancl_separation: "[| L(r); L(Z) |] ==> separation (L, \x. \w[L]. \wx[L]. \rp[L]. w \ Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \ rp)" apply (rule gen_separation_multi [OF wellfounded_trancl_reflects, of "{r,Z}"], auto) apply (rule_tac env="[r,Z]" in DPow_LsetI) apply (rule sep_rules tran_closure_iff_sats | simp)+ done subsubsection{*Instantiating the locale @{text M_trancl}*} lemma M_trancl_axioms_L: "M_trancl_axioms(L)" apply (rule M_trancl_axioms.intro) apply (assumption | rule rtrancl_separation wellfounded_trancl_separation)+ done theorem M_trancl_L: "PROP M_trancl(L)" by (rule M_trancl.intro [OF M_basic_L M_trancl_axioms_L]) -interpretation L: M_trancl L by (rule M_trancl_L) +interpretation L?: M_trancl L by (rule M_trancl_L) subsection{*@{term L} is Closed Under the Operator @{term list}*} subsubsection{*Instances of Replacement for Lists*} lemma list_replacement1_Reflects: "REFLECTS [\x. \u[L]. u \ B \ (\y[L]. pair(L,u,y,x) \ is_wfrec(L, iterates_MH(L, is_list_functor(L,A), 0), memsn, u, y)), \i x. \u \ Lset(i). u \ B \ (\y \ Lset(i). pair(##Lset(i), u, y, x) \ is_wfrec(##Lset(i), iterates_MH(##Lset(i), is_list_functor(##Lset(i), A), 0), memsn, u, y))]" by (intro FOL_reflections function_reflections is_wfrec_reflection iterates_MH_reflection list_functor_reflection) lemma list_replacement1: "L(A) ==> iterates_replacement(L, is_list_functor(L,A), 0)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) apply (rule strong_replacementI) apply (rule_tac u="{B,A,n,0,Memrel(succ(n))}" in gen_separation_multi [OF list_replacement1_Reflects], auto simp add: nonempty) apply (rule_tac env="[B,A,n,0,Memrel(succ(n))]" in DPow_LsetI) apply (rule sep_rules is_nat_case_iff_sats list_functor_iff_sats is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+ done lemma list_replacement2_Reflects: "REFLECTS [\x. \u[L]. u \ B & u \ nat & is_iterates(L, is_list_functor(L, A), 0, u, x), \i x. \u \ Lset(i). u \ B & u \ nat & is_iterates(##Lset(i), is_list_functor(##Lset(i), A), 0, u, x)]" by (intro FOL_reflections is_iterates_reflection list_functor_reflection) lemma list_replacement2: "L(A) ==> strong_replacement(L, \n y. n\nat & is_iterates(L, is_list_functor(L,A), 0, n, y))" apply (rule strong_replacementI) apply (rule_tac u="{A,B,0,nat}" in gen_separation_multi [OF list_replacement2_Reflects], auto simp add: L_nat nonempty) apply (rule_tac env="[A,B,0,nat]" in DPow_LsetI) apply (rule sep_rules list_functor_iff_sats is_iterates_iff_sats | simp)+ done subsection{*@{term L} is Closed Under the Operator @{term formula}*} subsubsection{*Instances of Replacement for Formulas*} (*FIXME: could prove a lemma iterates_replacementI to eliminate the need to expand iterates_replacement and wfrec_replacement*) lemma formula_replacement1_Reflects: "REFLECTS [\x. \u[L]. u \ B & (\y[L]. pair(L,u,y,x) & is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)), \i x. \u \ Lset(i). u \ B & (\y \ Lset(i). pair(##Lset(i), u, y, x) & is_wfrec(##Lset(i), iterates_MH(##Lset(i), is_formula_functor(##Lset(i)), 0), memsn, u, y))]" by (intro FOL_reflections function_reflections is_wfrec_reflection iterates_MH_reflection formula_functor_reflection) lemma formula_replacement1: "iterates_replacement(L, is_formula_functor(L), 0)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) apply (rule strong_replacementI) apply (rule_tac u="{B,n,0,Memrel(succ(n))}" in gen_separation_multi [OF formula_replacement1_Reflects], auto simp add: nonempty) apply (rule_tac env="[n,B,0,Memrel(succ(n))]" in DPow_LsetI) apply (rule sep_rules is_nat_case_iff_sats formula_functor_iff_sats is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+ done lemma formula_replacement2_Reflects: "REFLECTS [\x. \u[L]. u \ B & u \ nat & is_iterates(L, is_formula_functor(L), 0, u, x), \i x. \u \ Lset(i). u \ B & u \ nat & is_iterates(##Lset(i), is_formula_functor(##Lset(i)), 0, u, x)]" by (intro FOL_reflections is_iterates_reflection formula_functor_reflection) lemma formula_replacement2: "strong_replacement(L, \n y. n\nat & is_iterates(L, is_formula_functor(L), 0, n, y))" apply (rule strong_replacementI) apply (rule_tac u="{B,0,nat}" in gen_separation_multi [OF formula_replacement2_Reflects], auto simp add: nonempty L_nat) apply (rule_tac env="[B,0,nat]" in DPow_LsetI) apply (rule sep_rules formula_functor_iff_sats is_iterates_iff_sats | simp)+ done text{*NB The proofs for type @{term formula} are virtually identical to those for @{term "list(A)"}. It was a cut-and-paste job! *} subsubsection{*The Formula @{term is_nth}, Internalized*} (* "is_nth(M,n,l,Z) == \X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" *) definition nth_fm :: "[i,i,i]=>i" where "nth_fm(n,l,Z) == Exists(And(is_iterates_fm(tl_fm(1,0), succ(l), succ(n), 0), hd_fm(0,succ(Z))))" lemma nth_fm_type [TC]: "[| x \ nat; y \ nat; z \ nat |] ==> nth_fm(x,y,z) \ formula" by (simp add: nth_fm_def) lemma sats_nth_fm [simp]: "[| x < length(env); y \ nat; z \ nat; env \ list(A)|] ==> sats(A, nth_fm(x,y,z), env) <-> is_nth(##A, nth(x,env), nth(y,env), nth(z,env))" apply (frule lt_length_in_nat, assumption) apply (simp add: nth_fm_def is_nth_def sats_is_iterates_fm) done lemma nth_iff_sats: "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; i < length(env); j \ nat; k \ nat; env \ list(A)|] ==> is_nth(##A, x, y, z) <-> sats(A, nth_fm(i,j,k), env)" by (simp add: sats_nth_fm) theorem nth_reflection: "REFLECTS[\x. is_nth(L, f(x), g(x), h(x)), \i x. is_nth(##Lset(i), f(x), g(x), h(x))]" apply (simp only: is_nth_def) apply (intro FOL_reflections is_iterates_reflection hd_reflection tl_reflection) done subsubsection{*An Instance of Replacement for @{term nth}*} (*FIXME: could prove a lemma iterates_replacementI to eliminate the need to expand iterates_replacement and wfrec_replacement*) lemma nth_replacement_Reflects: "REFLECTS [\x. \u[L]. u \ B & (\y[L]. pair(L,u,y,x) & is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)), \i x. \u \ Lset(i). u \ B & (\y \ Lset(i). pair(##Lset(i), u, y, x) & is_wfrec(##Lset(i), iterates_MH(##Lset(i), is_tl(##Lset(i)), z), memsn, u, y))]" by (intro FOL_reflections function_reflections is_wfrec_reflection iterates_MH_reflection tl_reflection) lemma nth_replacement: "L(w) ==> iterates_replacement(L, is_tl(L), w)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) apply (rule strong_replacementI) apply (rule_tac u="{B,w,Memrel(succ(n))}" in gen_separation_multi [OF nth_replacement_Reflects], auto) apply (rule_tac env="[B,w,Memrel(succ(n))]" in DPow_LsetI) apply (rule sep_rules is_nat_case_iff_sats tl_iff_sats is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+ done subsubsection{*Instantiating the locale @{text M_datatypes}*} lemma M_datatypes_axioms_L: "M_datatypes_axioms(L)" apply (rule M_datatypes_axioms.intro) apply (assumption | rule list_replacement1 list_replacement2 formula_replacement1 formula_replacement2 nth_replacement)+ done theorem M_datatypes_L: "PROP M_datatypes(L)" apply (rule M_datatypes.intro) apply (rule M_trancl_L) apply (rule M_datatypes_axioms_L) done -interpretation L: M_datatypes L by (rule M_datatypes_L) +interpretation L?: M_datatypes L by (rule M_datatypes_L) subsection{*@{term L} is Closed Under the Operator @{term eclose}*} subsubsection{*Instances of Replacement for @{term eclose}*} lemma eclose_replacement1_Reflects: "REFLECTS [\x. \u[L]. u \ B & (\y[L]. pair(L,u,y,x) & is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)), \i x. \u \ Lset(i). u \ B & (\y \ Lset(i). pair(##Lset(i), u, y, x) & is_wfrec(##Lset(i), iterates_MH(##Lset(i), big_union(##Lset(i)), A), memsn, u, y))]" by (intro FOL_reflections function_reflections is_wfrec_reflection iterates_MH_reflection) lemma eclose_replacement1: "L(A) ==> iterates_replacement(L, big_union(L), A)" apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) apply (rule strong_replacementI) apply (rule_tac u="{B,A,n,Memrel(succ(n))}" in gen_separation_multi [OF eclose_replacement1_Reflects], auto) apply (rule_tac env="[B,A,n,Memrel(succ(n))]" in DPow_LsetI) apply (rule sep_rules iterates_MH_iff_sats is_nat_case_iff_sats is_wfrec_iff_sats big_union_iff_sats quasinat_iff_sats | simp)+ done lemma eclose_replacement2_Reflects: "REFLECTS [\x. \u[L]. u \ B & u \ nat & is_iterates(L, big_union(L), A, u, x), \i x. \u \ Lset(i). u \ B & u \ nat & is_iterates(##Lset(i), big_union(##Lset(i)), A, u, x)]" by (intro FOL_reflections function_reflections is_iterates_reflection) lemma eclose_replacement2: "L(A) ==> strong_replacement(L, \n y. n\nat & is_iterates(L, big_union(L), A, n, y))" apply (rule strong_replacementI) apply (rule_tac u="{A,B,nat}" in gen_separation_multi [OF eclose_replacement2_Reflects], auto simp add: L_nat) apply (rule_tac env="[A,B,nat]" in DPow_LsetI) apply (rule sep_rules is_iterates_iff_sats big_union_iff_sats | simp)+ done subsubsection{*Instantiating the locale @{text M_eclose}*} lemma M_eclose_axioms_L: "M_eclose_axioms(L)" apply (rule M_eclose_axioms.intro) apply (assumption | rule eclose_replacement1 eclose_replacement2)+ done theorem M_eclose_L: "PROP M_eclose(L)" apply (rule M_eclose.intro) apply (rule M_datatypes_L) apply (rule M_eclose_axioms_L) done -interpretation L: M_eclose L by (rule M_eclose_L) +interpretation L?: M_eclose L by (rule M_eclose_L) end diff --git a/src/ZF/Constructible/Separation.thy b/src/ZF/Constructible/Separation.thy --- a/src/ZF/Constructible/Separation.thy +++ b/src/ZF/Constructible/Separation.thy @@ -1,311 +1,311 @@ (* Title: ZF/Constructible/Separation.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory *) header{*Early Instances of Separation and Strong Replacement*} theory Separation imports L_axioms WF_absolute begin text{*This theory proves all instances needed for locale @{text "M_basic"}*} text{*Helps us solve for de Bruijn indices!*} lemma nth_ConsI: "[|nth(n,l) = x; n \ nat|] ==> nth(succ(n), Cons(a,l)) = x" by simp lemmas nth_rules = nth_0 nth_ConsI nat_0I nat_succI lemmas sep_rules = nth_0 nth_ConsI FOL_iff_sats function_iff_sats fun_plus_iff_sats lemma Collect_conj_in_DPow: "[| {x\A. P(x)} \ DPow(A); {x\A. Q(x)} \ DPow(A) |] ==> {x\A. P(x) & Q(x)} \ DPow(A)" by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric]) lemma Collect_conj_in_DPow_Lset: "[|z \ Lset(j); {x \ Lset(j). P(x)} \ DPow(Lset(j))|] ==> {x \ Lset(j). x \ z & P(x)} \ DPow(Lset(j))" apply (frule mem_Lset_imp_subset_Lset) apply (simp add: Collect_conj_in_DPow Collect_mem_eq subset_Int_iff2 elem_subset_in_DPow) done lemma separation_CollectI: "(\z. L(z) ==> L({x \ z . P(x)})) ==> separation(L, \x. P(x))" apply (unfold separation_def, clarify) apply (rule_tac x="{x\z. P(x)}" in rexI) apply simp_all done text{*Reduces the original comprehension to the reflected one*} lemma reflection_imp_L_separation: "[| \x\Lset(j). P(x) <-> Q(x); {x \ Lset(j) . Q(x)} \ DPow(Lset(j)); Ord(j); z \ Lset(j)|] ==> L({x \ z . P(x)})" apply (rule_tac i = "succ(j)" in L_I) prefer 2 apply simp apply (subgoal_tac "{x \ z. P(x)} = {x \ Lset(j). x \ z & (Q(x))}") prefer 2 apply (blast dest: mem_Lset_imp_subset_Lset) apply (simp add: Lset_succ Collect_conj_in_DPow_Lset) done text{*Encapsulates the standard proof script for proving instances of Separation.*} lemma gen_separation: assumes reflection: "REFLECTS [P,Q]" and Lu: "L(u)" and collI: "!!j. u \ Lset(j) \ Collect(Lset(j), Q(j)) \ DPow(Lset(j))" shows "separation(L,P)" apply (rule separation_CollectI) apply (rule_tac A="{u,z}" in subset_LsetE, blast intro: Lu) apply (rule ReflectsE [OF reflection], assumption) apply (drule subset_Lset_ltD, assumption) apply (erule reflection_imp_L_separation) apply (simp_all add: lt_Ord2, clarify) apply (rule collI, assumption) done text{*As above, but typically @{term u} is a finite enumeration such as @{term "{a,b}"}; thus the new subgoal gets the assumption @{term "{a,b} \ Lset(i)"}, which is logically equivalent to @{term "a \ Lset(i)"} and @{term "b \ Lset(i)"}.*} lemma gen_separation_multi: assumes reflection: "REFLECTS [P,Q]" and Lu: "L(u)" and collI: "!!j. u \ Lset(j) \ Collect(Lset(j), Q(j)) \ DPow(Lset(j))" shows "separation(L,P)" apply (rule gen_separation [OF reflection Lu]) apply (drule mem_Lset_imp_subset_Lset) apply (erule collI) done subsection{*Separation for Intersection*} lemma Inter_Reflects: "REFLECTS[\x. \y[L]. y\A --> x \ y, \i x. \y\Lset(i). y\A --> x \ y]" by (intro FOL_reflections) lemma Inter_separation: "L(A) ==> separation(L, \x. \y[L]. y\A --> x\y)" apply (rule gen_separation [OF Inter_Reflects], simp) apply (rule DPow_LsetI) txt{*I leave this one example of a manual proof. The tedium of manually instantiating @{term i}, @{term j} and @{term env} is obvious. *} apply (rule ball_iff_sats) apply (rule imp_iff_sats) apply (rule_tac [2] i=1 and j=0 and env="[y,x,A]" in mem_iff_sats) apply (rule_tac i=0 and j=2 in mem_iff_sats) apply (simp_all add: succ_Un_distrib [symmetric]) done subsection{*Separation for Set Difference*} lemma Diff_Reflects: "REFLECTS[\x. x \ B, \i x. x \ B]" by (intro FOL_reflections) lemma Diff_separation: "L(B) ==> separation(L, \x. x \ B)" apply (rule gen_separation [OF Diff_Reflects], simp) apply (rule_tac env="[B]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Cartesian Product*} lemma cartprod_Reflects: "REFLECTS[\z. \x[L]. x\A & (\y[L]. y\B & pair(L,x,y,z)), \i z. \x\Lset(i). x\A & (\y\Lset(i). y\B & pair(##Lset(i),x,y,z))]" by (intro FOL_reflections function_reflections) lemma cartprod_separation: "[| L(A); L(B) |] ==> separation(L, \z. \x[L]. x\A & (\y[L]. y\B & pair(L,x,y,z)))" apply (rule gen_separation_multi [OF cartprod_Reflects, of "{A,B}"], auto) apply (rule_tac env="[A,B]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Image*} lemma image_Reflects: "REFLECTS[\y. \p[L]. p\r & (\x[L]. x\A & pair(L,x,y,p)), \i y. \p\Lset(i). p\r & (\x\Lset(i). x\A & pair(##Lset(i),x,y,p))]" by (intro FOL_reflections function_reflections) lemma image_separation: "[| L(A); L(r) |] ==> separation(L, \y. \p[L]. p\r & (\x[L]. x\A & pair(L,x,y,p)))" apply (rule gen_separation_multi [OF image_Reflects, of "{A,r}"], auto) apply (rule_tac env="[A,r]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Converse*} lemma converse_Reflects: "REFLECTS[\z. \p[L]. p\r & (\x[L]. \y[L]. pair(L,x,y,p) & pair(L,y,x,z)), \i z. \p\Lset(i). p\r & (\x\Lset(i). \y\Lset(i). pair(##Lset(i),x,y,p) & pair(##Lset(i),y,x,z))]" by (intro FOL_reflections function_reflections) lemma converse_separation: "L(r) ==> separation(L, \z. \p[L]. p\r & (\x[L]. \y[L]. pair(L,x,y,p) & pair(L,y,x,z)))" apply (rule gen_separation [OF converse_Reflects], simp) apply (rule_tac env="[r]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Restriction*} lemma restrict_Reflects: "REFLECTS[\z. \x[L]. x\A & (\y[L]. pair(L,x,y,z)), \i z. \x\Lset(i). x\A & (\y\Lset(i). pair(##Lset(i),x,y,z))]" by (intro FOL_reflections function_reflections) lemma restrict_separation: "L(A) ==> separation(L, \z. \x[L]. x\A & (\y[L]. pair(L,x,y,z)))" apply (rule gen_separation [OF restrict_Reflects], simp) apply (rule_tac env="[A]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for Composition*} lemma comp_Reflects: "REFLECTS[\xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & xy\s & yz\r, \i xz. \x\Lset(i). \y\Lset(i). \z\Lset(i). \xy\Lset(i). \yz\Lset(i). pair(##Lset(i),x,z,xz) & pair(##Lset(i),x,y,xy) & pair(##Lset(i),y,z,yz) & xy\s & yz\r]" by (intro FOL_reflections function_reflections) lemma comp_separation: "[| L(r); L(s) |] ==> separation(L, \xz. \x[L]. \y[L]. \z[L]. \xy[L]. \yz[L]. pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & xy\s & yz\r)" apply (rule gen_separation_multi [OF comp_Reflects, of "{r,s}"], auto) txt{*Subgoals after applying general ``separation'' rule: @{subgoals[display,indent=0,margin=65]}*} apply (rule_tac env="[r,s]" in DPow_LsetI) txt{*Subgoals ready for automatic synthesis of a formula: @{subgoals[display,indent=0,margin=65]}*} apply (rule sep_rules | simp)+ done subsection{*Separation for Predecessors in an Order*} lemma pred_Reflects: "REFLECTS[\y. \p[L]. p\r & pair(L,y,x,p), \i y. \p \ Lset(i). p\r & pair(##Lset(i),y,x,p)]" by (intro FOL_reflections function_reflections) lemma pred_separation: "[| L(r); L(x) |] ==> separation(L, \y. \p[L]. p\r & pair(L,y,x,p))" apply (rule gen_separation_multi [OF pred_Reflects, of "{r,x}"], auto) apply (rule_tac env="[r,x]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for the Membership Relation*} lemma Memrel_Reflects: "REFLECTS[\z. \x[L]. \y[L]. pair(L,x,y,z) & x \ y, \i z. \x \ Lset(i). \y \ Lset(i). pair(##Lset(i),x,y,z) & x \ y]" by (intro FOL_reflections function_reflections) lemma Memrel_separation: "separation(L, \z. \x[L]. \y[L]. pair(L,x,y,z) & x \ y)" apply (rule gen_separation [OF Memrel_Reflects nonempty]) apply (rule_tac env="[]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Replacement for FunSpace*} lemma funspace_succ_Reflects: "REFLECTS[\z. \p[L]. p\A & (\f[L]. \b[L]. \nb[L]. \cnbf[L]. pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & upair(L,cnbf,cnbf,z)), \i z. \p \ Lset(i). p\A & (\f \ Lset(i). \b \ Lset(i). \nb \ Lset(i). \cnbf \ Lset(i). pair(##Lset(i),f,b,p) & pair(##Lset(i),n,b,nb) & is_cons(##Lset(i),nb,f,cnbf) & upair(##Lset(i),cnbf,cnbf,z))]" by (intro FOL_reflections function_reflections) lemma funspace_succ_replacement: "L(n) ==> strong_replacement(L, \p z. \f[L]. \b[L]. \nb[L]. \cnbf[L]. pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) & upair(L,cnbf,cnbf,z))" apply (rule strong_replacementI) apply (rule_tac u="{n,B}" in gen_separation_multi [OF funspace_succ_Reflects], auto) apply (rule_tac env="[n,B]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Separation for a Theorem about @{term "is_recfun"}*} lemma is_recfun_reflects: "REFLECTS[\x. \xa[L]. \xb[L]. pair(L,x,a,xa) & xa \ r & pair(L,x,b,xb) & xb \ r & (\fx[L]. \gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & fx \ gx), \i x. \xa \ Lset(i). \xb \ Lset(i). pair(##Lset(i),x,a,xa) & xa \ r & pair(##Lset(i),x,b,xb) & xb \ r & (\fx \ Lset(i). \gx \ Lset(i). fun_apply(##Lset(i),f,x,fx) & fun_apply(##Lset(i),g,x,gx) & fx \ gx)]" by (intro FOL_reflections function_reflections fun_plus_reflections) lemma is_recfun_separation: --{*for well-founded recursion*} "[| L(r); L(f); L(g); L(a); L(b) |] ==> separation(L, \x. \xa[L]. \xb[L]. pair(L,x,a,xa) & xa \ r & pair(L,x,b,xb) & xb \ r & (\fx[L]. \gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & fx \ gx))" apply (rule gen_separation_multi [OF is_recfun_reflects, of "{r,f,g,a,b}"], auto) apply (rule_tac env="[r,f,g,a,b]" in DPow_LsetI) apply (rule sep_rules | simp)+ done subsection{*Instantiating the locale @{text M_basic}*} text{*Separation (and Strong Replacement) for basic set-theoretic constructions such as intersection, Cartesian Product and image.*} lemma M_basic_axioms_L: "M_basic_axioms(L)" apply (rule M_basic_axioms.intro) apply (assumption | rule Inter_separation Diff_separation cartprod_separation image_separation converse_separation restrict_separation comp_separation pred_separation Memrel_separation funspace_succ_replacement is_recfun_separation)+ done theorem M_basic_L: "PROP M_basic(L)" by (rule M_basic.intro [OF M_trivial_L M_basic_axioms_L]) -interpretation L: M_basic L by (rule M_basic_L) +interpretation L?: M_basic L by (rule M_basic_L) end